| author | hoelzl | 
| Wed, 01 Dec 2010 21:03:02 +0100 | |
| changeset 40871 | 688f6ff859e1 | 
| parent 40653 | d921c97bdbd8 | 
| child 40906 | b5a319668955 | 
| permissions | -rw-r--r-- | 
| 37744 | 1 | (* Title: HOL/Mutabelle/mutabelle_extra.ML | 
| 34965 | 2 | Author: Stefan Berghofer, Jasmin Blanchette, Lukas Bulwahn, TU Muenchen | 
| 3 | ||
| 4 | Invokation of Counterexample generators | |
| 5 | *) | |
| 6 | signature MUTABELLE_EXTRA = | |
| 7 | sig | |
| 8 | ||
| 9 | val take_random : int -> 'a list -> 'a list | |
| 10 | ||
| 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 | 11 | datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error | Solved | Unsolved | 
| 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 | 12 | type timing = (string * int) list | 
| 34965 | 13 | |
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 14 | type mtd = string * (theory -> term -> outcome * (timing * (int * Quickcheck.report list) list option)) | 
| 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 | 15 | |
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 16 | type mutant_subentry = term * (string * (outcome * (timing * Quickcheck.report option))) list | 
| 34965 | 17 | type detailed_entry = string * bool * term * mutant_subentry list | 
| 18 | ||
| 19 | type subentry = string * int * int * int * int * int * int | |
| 20 | type entry = string * bool * subentry list | |
| 21 | type report = entry list | |
| 22 | ||
| 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 | 23 | 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 | 24 | |
| 
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 | val solve_direct_mtd : 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 | 26 | val try_mtd : mtd | 
| 34965 | 27 | (* | 
| 28 | val refute_mtd : mtd | |
| 29 | val nitpick_mtd : mtd | |
| 30 | *) | |
| 31 | ||
| 32 | val freezeT : term -> term | |
| 33 | val thms_of : bool -> theory -> thm list | |
| 34 | ||
| 35 | val string_for_report : report -> string | |
| 36 | val write_report : string -> report -> unit | |
| 37 | val mutate_theorems_and_write_report : | |
| 38 | theory -> mtd list -> thm list -> string -> unit | |
| 39 | ||
| 40 | val random_seed : real Unsynchronized.ref | |
| 41 | end; | |
| 42 | ||
| 43 | structure MutabelleExtra : MUTABELLE_EXTRA = | |
| 44 | struct | |
| 45 | ||
| 46 | (* Own seed; can't rely on the Isabelle one to stay the same *) | |
| 47 | val random_seed = Unsynchronized.ref 1.0; | |
| 48 | ||
| 49 | ||
| 50 | (* mutation 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 | 51 | (*val max_mutants = 4 | 
| 
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 | val num_mutations = 1*) | 
| 34965 | 53 | (* soundness check: *) | 
| 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 | 54 | val max_mutants = 10 | 
| 
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 | 55 | val num_mutations = 1 | 
| 34965 | 56 | |
| 57 | (* quickcheck options *) | |
| 58 | (*val quickcheck_generator = "SML"*) | |
| 59 | ||
| 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 | 60 | (* 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 | 61 | |
| 34965 | 62 | exception RANDOM; | 
| 63 | ||
| 64 | fun rmod x y = x - y * Real.realFloor (x / y); | |
| 65 | ||
| 66 | local | |
| 67 | val a = 16807.0; | |
| 68 | val m = 2147483647.0; | |
| 69 | in | |
| 70 | ||
| 71 | fun random () = CRITICAL (fn () => | |
| 72 | let val r = rmod (a * ! random_seed) m | |
| 73 | in (random_seed := r; r) end); | |
| 74 | ||
| 75 | end; | |
| 76 | ||
| 77 | fun random_range l h = | |
| 78 | if h < l orelse l < 0 then raise RANDOM | |
| 79 | else l + Real.floor (rmod (random ()) (real (h - l + 1))); | |
| 80 | ||
| 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 | 81 | 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 | 82 | | 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 | 83 | | 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 | 84 | 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 | 85 | 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 | 86 | 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 | 87 | |
| 
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 | (* 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 | 89 | |
| 
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 | 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 | 91 | |
| 
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 | 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 | 93 | | 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 | 94 | | 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 | 95 | | 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 | 96 | | 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 | 97 | | 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 | 98 | | 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 | 99 | | 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 | 100 | |
| 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 | 101 | type timing = (string * int) list | 
| 34965 | 102 | |
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 103 | type mtd = string * (theory -> term -> outcome * (timing * (int * Quickcheck.report list) list option)) | 
| 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 | 104 | |
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 105 | type mutant_subentry = term * (string * (outcome * (timing * Quickcheck.report option))) list | 
| 34965 | 106 | type detailed_entry = string * bool * term * mutant_subentry list | 
| 107 | ||
| 108 | type subentry = string * int * int * int * int * int * int | |
| 109 | type entry = string * bool * subentry list | |
| 110 | type report = entry list | |
| 111 | ||
| 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 | 112 | (* possible invocations *) | 
| 34965 | 113 | |
| 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 | 114 | (** quickcheck **) | 
| 34965 | 115 | |
| 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 | 116 | fun invoke_quickcheck change_options quickcheck_generator thy t = | 
| 39324 
05452dd66b2b
finished renaming "Auto_Counterexample" to "Auto_Tools"
 blanchet parents: 
39253diff
changeset | 117 | TimeLimit.timeLimit (Time.fromSeconds (!Auto_Tools.time_limit)) | 
| 34965 | 118 | (fn _ => | 
| 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 | 119 | case Quickcheck.test_goal_terms (change_options (ProofContext.init_global 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 | 120 | (SOME quickcheck_generator, []) [t] of | 
| 
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 | 121 | (NONE, _) => (NoCex, ([], 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 | 122 | | (SOME _, _) => (GenuineCex, ([], NONE))) () | 
| 39324 
05452dd66b2b
finished renaming "Auto_Counterexample" to "Auto_Tools"
 blanchet parents: 
39253diff
changeset | 123 |   handle TimeLimit.TimeOut => (Timeout, ([("timelimit", !Auto_Tools.time_limit)], NONE))
 | 
| 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 = | 
| 
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 | 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: 
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 | 
| 
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 | 132 | val state = Proof.theorem NONE (K I) (map (single o rpair []) [t]) (ProofContext.init_global 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 | 133 | 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 | 134 | case Solve_Direct.solve_direct false state of | 
| 
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 | 135 | (true, _) => (Solved, ([], 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 | 136 | | (false, _) => (Unsolved, ([], 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 | 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 | |
| 
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 | 141 | (** try **) | 
| 
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 | |
| 
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 | 143 | fun invoke_try 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 | 144 | 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 | 145 | val state = Proof.theorem NONE (K I) (map (single o rpair []) [t]) (ProofContext.init_global 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 | 146 | 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 | 147 | case Try.invoke_try NONE state of | 
| 
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 | 148 | true => (Solved, ([], 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 | 149 | | false => (Unsolved, ([], 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 | 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 | |
| 
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 | 152 | val try_mtd = ("try", invoke_try)
 | 
| 
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 | |
| 
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 | 154 | (* | 
| 34965 | 155 | fun invoke_refute thy t = | 
| 156 | let | |
| 157 | val res = MyRefute.refute_term thy [] t | |
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 158 |     val _ = Output.urgent_message ("Refute: " ^ res)
 | 
| 34965 | 159 | in | 
| 160 | case res of | |
| 161 | "genuine" => GenuineCex | |
| 162 | | "likely_genuine" => GenuineCex | |
| 163 | | "potential" => PotentialCex | |
| 164 | | "none" => NoCex | |
| 165 | | "unknown" => Donno | |
| 166 | | _ => Error | |
| 167 | end | |
| 168 | handle MyRefute.REFUTE (loc, details) => | |
| 169 |          (error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^
 | |
| 170 | ".")) | |
| 171 | val refute_mtd = ("refute", invoke_refute)
 | |
| 172 | *) | |
| 173 | ||
| 174 | (* | |
| 175 | open Nitpick_Util | |
| 176 | open Nitpick_Rep | |
| 177 | open Nitpick_Nut | |
| 178 | ||
| 179 | fun invoke_nitpick thy t = | |
| 180 | let | |
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36255diff
changeset | 181 | val ctxt = ProofContext.init_global thy | 
| 34965 | 182 | val state = Proof.init ctxt | 
| 183 | in | |
| 184 | let | |
| 185 | val (res, _) = Nitpick.pick_nits_in_term state (Nitpick_Isar.default_params thy []) false [] t | |
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 186 |       val _ = Output.urgent_message ("Nitpick: " ^ res)
 | 
| 34965 | 187 | in | 
| 188 | case res of | |
| 189 | "genuine" => GenuineCex | |
| 190 | | "likely_genuine" => GenuineCex | |
| 191 | | "potential" => PotentialCex | |
| 192 | | "none" => NoCex | |
| 193 | | "unknown" => Donno | |
| 194 | | _ => Error | |
| 195 | end | |
| 196 | handle ARG (loc, details) => | |
| 197 |            (error ("Bad argument(s) to " ^ quote loc ^ ": " ^ details ^ "."))
 | |
| 198 | | BAD (loc, details) => | |
| 199 |            (error ("Internal error (" ^ quote loc ^ "): " ^ details ^ "."))
 | |
| 200 | | LIMIT (_, details) => | |
| 201 |            (warning ("Limit reached: " ^ details ^ "."); Donno)
 | |
| 202 | | NOT_SUPPORTED details => | |
| 203 |            (warning ("Unsupported case: " ^ details ^ "."); Donno)
 | |
| 204 | | NUT (loc, us) => | |
| 205 |            (error ("Invalid nut" ^ plural_s_for_list us ^
 | |
| 206 |                    " (" ^ quote loc ^ "): " ^
 | |
| 207 | commas (map (string_for_nut ctxt) us) ^ ".")) | |
| 208 | | REP (loc, Rs) => | |
| 209 |            (error ("Invalid representation" ^ plural_s_for_list Rs ^
 | |
| 210 |                    " (" ^ quote loc ^ "): " ^
 | |
| 211 | commas (map string_for_rep Rs) ^ ".")) | |
| 212 | | TERM (loc, ts) => | |
| 213 |            (error ("Invalid term" ^ plural_s_for_list ts ^
 | |
| 214 |                    " (" ^ quote loc ^ "): " ^
 | |
| 215 | commas (map (Syntax.string_of_term ctxt) ts) ^ ".")) | |
| 216 | | TYPE (loc, Ts, ts) => | |
| 217 |            (error ("Invalid type" ^ plural_s_for_list Ts ^
 | |
| 218 | (if null ts then | |
| 219 | "" | |
| 220 | else | |
| 221 | " for term" ^ plural_s_for_list ts ^ " " ^ | |
| 222 | commas (map (quote o Syntax.string_of_term ctxt) ts)) ^ | |
| 223 |                    " (" ^ quote loc ^ "): " ^
 | |
| 224 | commas (map (Syntax.string_of_typ ctxt) Ts) ^ ".")) | |
| 225 | | Kodkod.SYNTAX (_, details) => | |
| 226 |            (warning ("Ill-formed Kodkodi output: " ^ details ^ "."); Error)
 | |
| 227 | | Refute.REFUTE (loc, details) => | |
| 228 |            (error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^
 | |
| 229 | ".")) | |
| 40381 
96c37a685a13
explicit indication of some remaining violations of the Isabelle/ML interrupt model;
 wenzelm parents: 
40301diff
changeset | 230 | | Exn.Interrupt => raise Exn.Interrupt (* FIXME violates Isabelle/ML exception model *) | 
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 231 |          | _ => (Output.urgent_message ("Unknown error in Nitpick"); Error)
 | 
| 34965 | 232 | end | 
| 233 | val nitpick_mtd = ("nitpick", invoke_nitpick)
 | |
| 234 | *) | |
| 235 | ||
| 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 | 236 | (* 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 | 237 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38857diff
changeset | 238 | val comms = [@{const_name HOL.eq}, @{const_name HOL.disj}, @{const_name HOL.conj}]
 | 
| 34965 | 239 | |
| 240 | val forbidden = | |
| 241 |  [(* (@{const_name "power"}, "'a"), *)
 | |
| 35325 
4123977b469d
adding ROOT.ML to HOL-Mutabelle session; uncommenting HOL.induct constants in Mutabelle session
 bulwahn parents: 
35324diff
changeset | 242 |   (*(@{const_name induct_equal}, "'a"),
 | 
| 
4123977b469d
adding ROOT.ML to HOL-Mutabelle session; uncommenting HOL.induct constants in Mutabelle session
 bulwahn parents: 
35324diff
changeset | 243 |   (@{const_name induct_implies}, "'a"),
 | 
| 
4123977b469d
adding ROOT.ML to HOL-Mutabelle session; uncommenting HOL.induct constants in Mutabelle session
 bulwahn parents: 
35324diff
changeset | 244 |   (@{const_name induct_conj}, "'a"),*)
 | 
| 34965 | 245 |   (@{const_name "undefined"}, "'a"),
 | 
| 246 |   (@{const_name "default"}, "'a"),
 | |
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 247 |   (@{const_name "dummy_pattern"}, "'a::{}"),
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 248 |   (@{const_name "HOL.simp_implies"}, "prop => prop => prop"),
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 249 |   (@{const_name "bot_fun_inst.bot_fun"}, "'a"),
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 250 |   (@{const_name "top_fun_inst.top_fun"}, "'a"),
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 251 |   (@{const_name "Pure.term"}, "'a"),
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 252 |   (@{const_name "top_class.top"}, "'a"),
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 253 |   (@{const_name "Quotient.Quot_True"}, "'a")(*,
 | 
| 34965 | 254 |   (@{const_name "uminus"}, "'a"),
 | 
| 255 |   (@{const_name "Nat.size"}, "'a"),
 | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
34974diff
changeset | 256 |   (@{const_name "Groups.abs"}, "'a") *)]
 | 
| 34965 | 257 | |
| 258 | val forbidden_thms = | |
| 259 | ["finite_intvl_succ_class", | |
| 260 | "nibble"] | |
| 261 | ||
| 262 | val forbidden_consts = | |
| 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 | 263 |  [@{const_name nibble_pair_of_char}, @{const_name "TYPE"}]
 | 
| 34965 | 264 | |
| 265 | fun is_forbidden_theorem (s, th) = | |
| 266 | let val consts = Term.add_const_names (prop_of th) [] in | |
| 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 | 267 | exists (member (op =) (space_explode "." s)) forbidden_thms orelse | 
| 
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 | 268 | exists (member (op =) forbidden_consts) consts orelse | 
| 34965 | 269 | length (space_explode "." s) <> 2 orelse | 
| 270 | String.isPrefix "type_definition" (List.last (space_explode "." s)) orelse | |
| 271 | 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 | 272 | String.isSuffix "_raw" s 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 | 273 | String.isPrefix "term_of" (List.last (space_explode "." s)) | 
| 34965 | 274 | end | 
| 275 | ||
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 276 | val forbidden_mutant_constnames = | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 277 | ["HOL.induct_equal", | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 278 | "HOL.induct_implies", | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 279 | "HOL.induct_conj", | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 280 |  @{const_name undefined},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 281 |  @{const_name default},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 282 |  @{const_name dummy_pattern},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 283 |  @{const_name "HOL.simp_implies"},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 284 |  @{const_name "bot_fun_inst.bot_fun"},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 285 |  @{const_name "top_fun_inst.top_fun"},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 286 |  @{const_name "Pure.term"},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 287 |  @{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 | 288 |  (*@{const_name "HOL.equal"},*)
 | 
| 
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 "Quotient.Quot_True"}
 | 
| 
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 "==>"}, @{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 | 291 | |
| 
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 | 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 | 293 | [ | 
| 
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.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 | 295 |    (@{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 | 296 |    (@{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 | 297 |    (@{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 | 298 |    (@{const_name "Groups.times_class.times"}, @{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 | 299 |    (@{const_name "Rings.inverse_class.divide"}, @{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 | 300 |    (@{const_name "Lattices.semilattice_inf_class.inf"}, @{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 | 301 |    (@{const_name "Lattices.semilattice_sup_class.sup"}, @{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 | 302 |    (@{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 | 303 |    (@{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 | 304 |    (@{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 | 305 |    (@{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 | 306 |    (@{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 | 307 |    (@{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 | 308 |    (@{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 | 309 |    (@{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 | 310 |    (@{const_name "Groups.one_class.one"}, @{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 | 311 |    (@{const_name "Groups.zero_class.zero"},@{typ "bool => prop"})]
 | 
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 312 | |
| 34965 | 313 | fun is_forbidden_mutant t = | 
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 314 | 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 | 315 | 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 | 316 | val consts = Term.add_consts t [] | 
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 317 | 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 | 318 | 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 | 319 | 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 | 320 | 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 | 321 | exists (member (op =) forbidden_mutant_consts) consts | 
| 34965 | 322 | end | 
| 323 | ||
| 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 | 324 | (* 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 | 325 | |
| 40248 | 326 | 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 | 327 | 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 | 328 | val ctxt = ProofContext.init_global 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 | 329 | 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 | 330 | can (TimeLimit.timeLimit (seconds 2.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 | 331 | (Quickcheck.test_goal_terms | 
| 
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 | 332 | ((Config.put Quickcheck.finite_types true #> | 
| 
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 | 333 | Config.put Quickcheck.size 1 #> Config.put Quickcheck.iterations 1) ctxt) | 
| 
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 | 334 | (SOME "random" , []))) (map (Object_Logic.atomize_term thy) (fst (Variable.import_terms true [t] ctxt))) | 
| 
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 | 335 | end | 
| 40248 | 336 | |
| 34965 | 337 | fun is_executable_thm thy th = is_executable_term thy (prop_of th) | 
| 338 | ||
| 339 | val freezeT = | |
| 340 | map_types (map_type_tvar (fn ((a, i), S) => | |
| 341 | TFree (if i = 0 then a else a ^ "_" ^ string_of_int i, S))) | |
| 342 | ||
| 343 | fun thms_of all thy = | |
| 344 | filter | |
| 345 | (fn th => (all orelse Context.theory_name (theory_of_thm th) = Context.theory_name thy) | |
| 346 | (* andalso is_executable_thm thy th *)) | |
| 347 | (map snd (filter_out is_forbidden_theorem (Mutabelle.all_unconcealed_thms_of thy))) | |
| 348 | ||
| 349 | val count = length oo filter o equal | |
| 350 | ||
| 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 | 351 | fun cpu_time description f = | 
| 
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 | 352 | let | 
| 
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 | 353 | val start = start_timing () | 
| 
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 | 354 | val result = Exn.capture f () | 
| 
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 | 355 | val time = Time.toMilliseconds (#cpu (end_timing start)) | 
| 
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 | 356 | in (Exn.release result, (description, time)) 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 | 357 | (* | 
| 
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 | 358 | 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 | 359 | 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 | 360 |     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 | 361 | 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 | 362 | 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 | 363 |     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 | 364 | 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 | 365 | *) | 
| 34965 | 366 | fun safe_invoke_mtd thy (mtd_name, invoke_mtd) t = | 
| 367 | let | |
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 368 |     val _ = Output.urgent_message ("Invoking " ^ mtd_name)
 | 
| 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 | 369 | val (res, (timing, reports)) = (*cpu_time "total time" | 
| 
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 | 370 | (fn () => *)case try (invoke_mtd thy) t of | 
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 371 | SOME (res, (timing, reports)) => (res, (timing, reports)) | 
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 372 |         | NONE => (Output.urgent_message ("**** PROBLEMS WITH " ^ 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: 
40381diff
changeset | 373 | (Error, ([], NONE))) | 
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 374 |     val _ = Output.urgent_message (" Done")
 | 
| 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 | 375 | in (res, (timing, reports)) end | 
| 34965 | 376 | |
| 377 | (* 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: 
40381diff
changeset | 378 | |
| 34965 | 379 | fun test_mutants_using_one_method thy mutants (mtd_name, invoke_mtd) = | 
| 380 | 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 | 381 | val res = map (fst o safe_invoke_mtd thy (mtd_name, invoke_mtd)) mutants | 
| 34965 | 382 | in | 
| 383 | (mtd_name, count GenuineCex res, count PotentialCex res, count NoCex res, | |
| 384 | count Donno res, count Timeout res, count Error res) | |
| 385 | end | |
| 386 | ||
| 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 | 387 | (* 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 | 388 | |
| 34965 | 389 | 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: 
36692diff
changeset | 390 | (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: 
40381diff
changeset | 391 | |
| 34965 | 392 | fun create_detailed_entry thy thm exec mutants mtds = | 
| 393 | let | |
| 394 | fun create_mutant_subentry mutant = (mutant, | |
| 395 | map (fn (mtd_name, invoke_mtd) => | |
| 396 | (mtd_name, safe_invoke_mtd thy (mtd_name, invoke_mtd) mutant)) mtds) | |
| 397 | in | |
| 36743 
ce2297415b54
prefer Thm.get_name_hint, which is closer to a user-space idea of "theorem name";
 wenzelm parents: 
36692diff
changeset | 398 | (Thm.get_name_hint thm, exec, prop_of thm, map create_mutant_subentry mutants) | 
| 34965 | 399 | end | 
| 400 | ||
| 401 | (* (theory -> thm -> bool -> term list -> mtd list -> 'a) -> theory -> mtd list -> thm -> 'a *) | |
| 402 | fun mutate_theorem create_entry thy mtds thm = | |
| 403 | let | |
| 404 | val exec = is_executable_thm thy thm | |
| 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 | 405 | val _ = Output.tracing (if exec then "EXEC" else "NOEXEC") | 
| 34965 | 406 | val mutants = | 
| 407 | (if num_mutations = 0 then | |
| 408 | [Thm.prop_of thm] | |
| 409 | else | |
| 410 | Mutabelle.mutate_mix (Thm.prop_of thm) thy comms forbidden | |
| 411 | 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 | 412 |              |> tap (fn muts => tracing ("mutants: " ^ string_of_int (length muts)))
 | 
| 34965 | 413 | |> filter_out is_forbidden_mutant | 
| 414 | val mutants = | |
| 415 | if exec then | |
| 416 | let | |
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 417 |           val _ = Output.urgent_message ("BEFORE PARTITION OF " ^
 | 
| 34965 | 418 | Int.toString (length mutants) ^ " MUTANTS") | 
| 419 | val (execs, noexecs) = List.partition (is_executable_term thy) (take_random (20 * max_mutants) mutants) | |
| 420 |           val _ = tracing ("AFTER PARTITION (" ^ Int.toString (length execs) ^
 | |
| 421 | " vs " ^ Int.toString (length noexecs) ^ ")") | |
| 422 | in | |
| 423 | execs @ take_random (Int.max (0, max_mutants - length execs)) noexecs | |
| 424 | end | |
| 425 | else | |
| 426 | mutants | |
| 427 | val mutants = mutants | |
| 428 | |> take_random max_mutants | |
| 429 | |> map Mutabelle.freeze |> map freezeT | |
| 430 | (* |> filter (not o is_forbidden_mutant) *) | |
| 431 | |> List.mapPartial (try (Sign.cert_term 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 | 432 | |> List.filter (is_some o try (cterm_of thy)) | 
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 433 |     val _ = map (fn t => Output.urgent_message ("MUTANT: " ^ Syntax.string_of_term_global thy t)) mutants
 | 
| 34965 | 434 | in | 
| 435 | create_entry thy thm exec mutants mtds | |
| 436 | end | |
| 437 | ||
| 438 | (* theory -> mtd list -> thm list -> report *) | |
| 439 | val mutate_theorems = map ooo mutate_theorem | |
| 440 | ||
| 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 | 441 | fun string_of_mutant_subentry thy thm_name (t, results) = | 
| 34965 | 442 | "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: 
35092diff
changeset | 443 | 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: 
35092diff
changeset | 444 | (map (fn (mtd_name, (outcome, timing)) => mtd_name ^ ": " ^ string_of_outcome outcome) results) ^ | 
| 34965 | 445 | "\n" | 
| 446 | ||
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 447 | (* string -> string *) | 
| 39555 
ccb223a4d49c
added XML.content_of convenience -- cover XML.body, which is the general situation;
 wenzelm parents: 
39324diff
changeset | 448 | 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 | 449 | |
| 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 | 450 | fun string_of_mutant_subentry' thy thm_name (t, results) = | 
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 451 | let | 
| 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 452 |     fun string_of_report (Quickcheck.Report {iterations = i, raised_match_errors = e,
 | 
| 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 453 | satisfied_assms = s, positive_concl_tests = p}) = | 
| 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 454 | "errors: " ^ string_of_int e ^ "; conclusion tests: " ^ string_of_int p | 
| 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 455 | fun string_of_reports NONE = "" | 
| 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 456 | | string_of_reports (SOME reports) = | 
| 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 457 | cat_lines (map (fn (size, [report]) => | 
| 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 458 | "size " ^ string_of_int size ^ ": " ^ string_of_report report) (rev reports)) | 
| 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 459 | fun string_of_mtd_result (mtd_name, (outcome, (timing, reports))) = | 
| 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 | 460 | mtd_name ^ ": " ^ string_of_outcome outcome | 
| 
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 | 461 |       (*" with time " ^ " (" ^ 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 | 462 | (*^ "\n" ^ string_of_reports reports*) | 
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 463 | in | 
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 464 | "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 | 465 | ^ 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 | 466 | 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 | 467 | |
| 34965 | 468 | fun string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) = | 
| 469 | thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ": " ^ | |
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 470 | 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 | 471 | cat_lines (map (string_of_mutant_subentry' thy thm_name) mutant_subentries) ^ "\n" | 
| 34965 | 472 | |
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 473 | 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: 
35625diff
changeset | 474 | "lemma " ^ thm_name ^ "_" ^ string_of_int (i + 1) ^ ":\n" ^ | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 475 | "\"" ^ 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: 
40381diff
changeset | 476 | "\" \nquickcheck\noops\n" | 
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 477 | |
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 478 | 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: 
35625diff
changeset | 479 |   "subsubsection {* mutants of " ^ thm_name ^ " *}\n\n" ^
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 480 | cat_lines (map_index | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 481 | (theoryfile_string_of_mutant_subentry thy thm_name) mutant_subentries) ^ "\n" | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 482 | |
| 34965 | 483 | (* subentry -> string *) | 
| 484 | fun string_for_subentry (mtd_name, genuine_cex, potential_cex, no_cex, donno, | |
| 485 | timeout, error) = | |
| 486 | " " ^ mtd_name ^ ": " ^ Int.toString genuine_cex ^ "+ " ^ | |
| 487 | Int.toString potential_cex ^ "= " ^ Int.toString no_cex ^ "- " ^ | |
| 488 | Int.toString donno ^ "? " ^ Int.toString timeout ^ "T " ^ | |
| 489 | Int.toString 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 | 490 | |
| 34965 | 491 | (* entry -> string *) | 
| 492 | fun string_for_entry (thm_name, exec, subentries) = | |
| 493 | thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ":\n" ^ | |
| 494 | 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 | 495 | |
| 34965 | 496 | (* report -> string *) | 
| 497 | fun string_for_report report = cat_lines (map string_for_entry report) | |
| 498 | ||
| 499 | (* string -> report -> unit *) | |
| 500 | fun write_report file_name = | |
| 501 | File.write (Path.explode file_name) o string_for_report | |
| 502 | ||
| 503 | (* theory -> mtd list -> thm list -> string -> unit *) | |
| 504 | fun mutate_theorems_and_write_report thy mtds thms file_name = | |
| 505 | let | |
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 506 | val _ = Output.urgent_message "Starting Mutabelle..." | 
| 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 | 507 | val ctxt = ProofContext.init_global thy | 
| 34965 | 508 | val path = Path.explode file_name | 
| 509 | (* 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 | 510 | (* | 
| 
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 | 511 | 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 | 512 | *) | 
| 
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 | 513 | (* 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 | 514 | 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 | 515 | (* 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 | 516 | (*val (gen_create_entry, gen_string_for_entry) = (create_detailed_entry, theoryfile_string_of_detailed_entry thy)*) | 
| 34965 | 517 | in | 
| 518 | File.write path ( | |
| 519 | "Mutation options = " ^ | |
| 520 | "max_mutants: " ^ string_of_int max_mutants ^ | |
| 521 | "; num_mutations: " ^ string_of_int num_mutations ^ "\n" ^ | |
| 522 | "QC options = " ^ | |
| 523 | (*"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 | 524 | "size: " ^ string_of_int (Config.get ctxt Quickcheck.size) ^ | 
| 
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 | 525 | "; iterations: " ^ string_of_int (Config.get ctxt Quickcheck.iterations) ^ "\n"); | 
| 34965 | 526 | map (File.append path o gen_string_for_entry o mutate_theorem gen_create_entry thy mtds) thms; | 
| 527 | () | |
| 528 | end | |
| 529 | ||
| 530 | end; |