src/HOL/Mutabelle/mutabelle_extra.ML
author haftmann
Thu, 08 Jul 2010 16:19:24 +0200
changeset 37743 3daaf23b9ab4
parent 36752 ce2297415b54
child 39028 848be46708dc
permissions -rw-r--r--
tuned titles
haftmann@37743
     1
(*  Title:      HOL/Mutabelle/mutabelle_extra.ML
bulwahn@34952
     2
    Author:     Stefan Berghofer, Jasmin Blanchette, Lukas Bulwahn, TU Muenchen
bulwahn@34952
     3
bulwahn@34952
     4
    Invokation of Counterexample generators
bulwahn@34952
     5
*)
bulwahn@34952
     6
signature MUTABELLE_EXTRA =
bulwahn@34952
     7
sig
bulwahn@34952
     8
bulwahn@34952
     9
val take_random : int -> 'a list -> 'a list
bulwahn@34952
    10
bulwahn@34952
    11
datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error
bulwahn@35324
    12
type timing = (string * int) list
bulwahn@34952
    13
bulwahn@35380
    14
type mtd = string * (theory -> term -> outcome * (timing * (int * Quickcheck.report list) list option))
bulwahn@35324
    15
bulwahn@35380
    16
type mutant_subentry = term * (string * (outcome * (timing * Quickcheck.report option))) list
bulwahn@34952
    17
type detailed_entry = string * bool * term * mutant_subentry list
bulwahn@34952
    18
bulwahn@34952
    19
type subentry = string * int * int * int * int * int * int
bulwahn@34952
    20
type entry = string * bool * subentry list
bulwahn@34952
    21
type report = entry list
bulwahn@34952
    22
bulwahn@34952
    23
val quickcheck_mtd : string -> mtd
bulwahn@34952
    24
(*
bulwahn@34952
    25
val refute_mtd : mtd
bulwahn@34952
    26
val nitpick_mtd : mtd
bulwahn@34952
    27
*)
bulwahn@34952
    28
bulwahn@34952
    29
val freezeT : term -> term
bulwahn@34952
    30
val thms_of : bool -> theory -> thm list
bulwahn@34952
    31
bulwahn@34952
    32
val string_for_report : report -> string
bulwahn@34952
    33
val write_report : string -> report -> unit
bulwahn@34952
    34
val mutate_theorems_and_write_report :
bulwahn@34952
    35
  theory -> mtd list -> thm list -> string -> unit
bulwahn@34952
    36
bulwahn@34952
    37
val random_seed : real Unsynchronized.ref
bulwahn@34952
    38
end;
bulwahn@34952
    39
bulwahn@34952
    40
structure MutabelleExtra : MUTABELLE_EXTRA =
bulwahn@34952
    41
struct
bulwahn@34952
    42
bulwahn@34952
    43
(* Own seed; can't rely on the Isabelle one to stay the same *)
bulwahn@34952
    44
val random_seed = Unsynchronized.ref 1.0;
bulwahn@34952
    45
bulwahn@34952
    46
bulwahn@34952
    47
(* mutation options *)
bulwahn@34952
    48
val max_mutants = 4
bulwahn@34952
    49
val num_mutations = 1
bulwahn@34952
    50
(* soundness check: *)
bulwahn@36255
    51
(*val max_mutants = 1
bulwahn@36255
    52
val num_mutations = 0*)
bulwahn@34952
    53
bulwahn@34952
    54
(* quickcheck options *)
bulwahn@34952
    55
(*val quickcheck_generator = "SML"*)
bulwahn@35523
    56
val iterations = 10
bulwahn@34952
    57
val size = 5
bulwahn@34952
    58
bulwahn@34952
    59
exception RANDOM;
bulwahn@34952
    60
bulwahn@34952
    61
fun rmod x y = x - y * Real.realFloor (x / y);
bulwahn@34952
    62
bulwahn@34952
    63
local
bulwahn@34952
    64
  val a = 16807.0;
bulwahn@34952
    65
  val m = 2147483647.0;
bulwahn@34952
    66
in
bulwahn@34952
    67
bulwahn@34952
    68
fun random () = CRITICAL (fn () =>
bulwahn@34952
    69
  let val r = rmod (a * ! random_seed) m
bulwahn@34952
    70
  in (random_seed := r; r) end);
bulwahn@34952
    71
bulwahn@34952
    72
end;
bulwahn@34952
    73
bulwahn@34952
    74
fun random_range l h =
bulwahn@34952
    75
  if h < l orelse l < 0 then raise RANDOM
bulwahn@34952
    76
  else l + Real.floor (rmod (random ()) (real (h - l + 1)));
bulwahn@34952
    77
bulwahn@34952
    78
datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error
bulwahn@35324
    79
type timing = (string * int) list
bulwahn@34952
    80
bulwahn@35380
    81
type mtd = string * (theory -> term -> outcome * (timing * (int * Quickcheck.report list) list option))
bulwahn@35324
    82
bulwahn@35380
    83
type mutant_subentry = term * (string * (outcome * (timing * Quickcheck.report option))) list
bulwahn@34952
    84
type detailed_entry = string * bool * term * mutant_subentry list
bulwahn@34952
    85
bulwahn@34952
    86
type subentry = string * int * int * int * int * int * int
bulwahn@34952
    87
type entry = string * bool * subentry list
bulwahn@34952
    88
type report = entry list
bulwahn@34952
    89
bulwahn@34952
    90
fun inst_type insts (Type (s, Ts)) = Type (s, map (inst_type insts) Ts)
bulwahn@34952
    91
  | inst_type insts T = the_default HOLogic.intT (AList.lookup op = insts T)
bulwahn@34952
    92
wenzelm@35625
    93
fun preprocess thy insts t = Object_Logic.atomize_term thy
bulwahn@34952
    94
 (map_types (inst_type insts) (Mutabelle.freeze t));
bulwahn@34952
    95
bulwahn@34952
    96
fun invoke_quickcheck quickcheck_generator thy t =
bulwahn@34952
    97
  TimeLimit.timeLimit (Time.fromSeconds (! Auto_Counterexample.time_limit))
bulwahn@34952
    98
      (fn _ =>
wenzelm@36633
    99
          case Quickcheck.gen_test_term (ProofContext.init_global thy) true true (SOME quickcheck_generator)
bulwahn@34952
   100
                                    size iterations (preprocess thy [] t) of
bulwahn@35380
   101
            (NONE, (time_report, report)) => (NoCex, (time_report, report))
bulwahn@35380
   102
          | (SOME _, (time_report, report)) => (GenuineCex, (time_report, report))) ()
bulwahn@35380
   103
  handle TimeLimit.TimeOut => (Timeout, ([("timelimit", !Auto_Counterexample.time_limit)], NONE))
bulwahn@34952
   104
bulwahn@34952
   105
fun quickcheck_mtd quickcheck_generator =
bulwahn@34952
   106
  ("quickcheck_" ^ quickcheck_generator, invoke_quickcheck quickcheck_generator)
bulwahn@34952
   107
bulwahn@34952
   108
  (*
bulwahn@34952
   109
fun invoke_refute thy t =
bulwahn@34952
   110
  let
bulwahn@34952
   111
    val res = MyRefute.refute_term thy [] t
bulwahn@34952
   112
    val _ = priority ("Refute: " ^ res)
bulwahn@34952
   113
  in
bulwahn@34952
   114
    case res of
bulwahn@34952
   115
      "genuine" => GenuineCex
bulwahn@34952
   116
    | "likely_genuine" => GenuineCex
bulwahn@34952
   117
    | "potential" => PotentialCex
bulwahn@34952
   118
    | "none" => NoCex
bulwahn@34952
   119
    | "unknown" => Donno
bulwahn@34952
   120
    | _ => Error
bulwahn@34952
   121
  end
bulwahn@34952
   122
  handle MyRefute.REFUTE (loc, details) =>
bulwahn@34952
   123
         (error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^
bulwahn@34952
   124
                   "."))
bulwahn@34952
   125
val refute_mtd = ("refute", invoke_refute)
bulwahn@34952
   126
*)
bulwahn@34952
   127
bulwahn@34952
   128
(*
bulwahn@34952
   129
open Nitpick_Util
bulwahn@34952
   130
open Nitpick_Rep
bulwahn@34952
   131
open Nitpick_Nut
bulwahn@34952
   132
bulwahn@34952
   133
fun invoke_nitpick thy t =
bulwahn@34952
   134
  let
wenzelm@36633
   135
    val ctxt = ProofContext.init_global thy
bulwahn@34952
   136
    val state = Proof.init ctxt
bulwahn@34952
   137
  in
bulwahn@34952
   138
    let
bulwahn@34952
   139
      val (res, _) = Nitpick.pick_nits_in_term state (Nitpick_Isar.default_params thy []) false [] t
bulwahn@34952
   140
      val _ = priority ("Nitpick: " ^ res)
bulwahn@34952
   141
    in
bulwahn@34952
   142
      case res of
bulwahn@34952
   143
        "genuine" => GenuineCex
bulwahn@34952
   144
      | "likely_genuine" => GenuineCex
bulwahn@34952
   145
      | "potential" => PotentialCex
bulwahn@34952
   146
      | "none" => NoCex
bulwahn@34952
   147
      | "unknown" => Donno
bulwahn@34952
   148
      | _ => Error
bulwahn@34952
   149
    end
bulwahn@34952
   150
    handle ARG (loc, details) =>
bulwahn@34952
   151
           (error ("Bad argument(s) to " ^ quote loc ^ ": " ^ details ^ "."))
bulwahn@34952
   152
         | BAD (loc, details) =>
bulwahn@34952
   153
           (error ("Internal error (" ^ quote loc ^ "): " ^ details ^ "."))
bulwahn@34952
   154
         | LIMIT (_, details) =>
bulwahn@34952
   155
           (warning ("Limit reached: " ^ details ^ "."); Donno)
bulwahn@34952
   156
         | NOT_SUPPORTED details =>
bulwahn@34952
   157
           (warning ("Unsupported case: " ^ details ^ "."); Donno)
bulwahn@34952
   158
         | NUT (loc, us) =>
bulwahn@34952
   159
           (error ("Invalid nut" ^ plural_s_for_list us ^
bulwahn@34952
   160
                   " (" ^ quote loc ^ "): " ^
bulwahn@34952
   161
                  commas (map (string_for_nut ctxt) us) ^ "."))
bulwahn@34952
   162
         | REP (loc, Rs) =>
bulwahn@34952
   163
           (error ("Invalid representation" ^ plural_s_for_list Rs ^
bulwahn@34952
   164
                   " (" ^ quote loc ^ "): " ^
bulwahn@34952
   165
                   commas (map string_for_rep Rs) ^ "."))
bulwahn@34952
   166
         | TERM (loc, ts) =>
bulwahn@34952
   167
           (error ("Invalid term" ^ plural_s_for_list ts ^
bulwahn@34952
   168
                   " (" ^ quote loc ^ "): " ^
bulwahn@34952
   169
                   commas (map (Syntax.string_of_term ctxt) ts) ^ "."))
bulwahn@34952
   170
         | TYPE (loc, Ts, ts) =>
bulwahn@34952
   171
           (error ("Invalid type" ^ plural_s_for_list Ts ^
bulwahn@34952
   172
                   (if null ts then
bulwahn@34952
   173
                      ""
bulwahn@34952
   174
                    else
bulwahn@34952
   175
                      " for term" ^ plural_s_for_list ts ^ " " ^
bulwahn@34952
   176
                      commas (map (quote o Syntax.string_of_term ctxt) ts)) ^
bulwahn@34952
   177
                   " (" ^ quote loc ^ "): " ^
bulwahn@34952
   178
                   commas (map (Syntax.string_of_typ ctxt) Ts) ^ "."))
bulwahn@34952
   179
         | Kodkod.SYNTAX (_, details) =>
bulwahn@34952
   180
           (warning ("Ill-formed Kodkodi output: " ^ details ^ "."); Error)
bulwahn@34952
   181
         | Refute.REFUTE (loc, details) =>
bulwahn@34952
   182
           (error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^
bulwahn@34952
   183
                   "."))
bulwahn@34952
   184
         | Exn.Interrupt => raise Exn.Interrupt
bulwahn@34952
   185
         | _ => (priority ("Unknown error in Nitpick"); Error)
bulwahn@34952
   186
  end
bulwahn@34952
   187
val nitpick_mtd = ("nitpick", invoke_nitpick)
bulwahn@34952
   188
*)
bulwahn@34952
   189
bulwahn@34952
   190
val comms = [@{const_name "op ="}, @{const_name "op |"}, @{const_name "op &"}]
bulwahn@34952
   191
bulwahn@34952
   192
val forbidden =
bulwahn@34952
   193
 [(* (@{const_name "power"}, "'a"), *)
bulwahn@35325
   194
  (*(@{const_name induct_equal}, "'a"),
bulwahn@35325
   195
  (@{const_name induct_implies}, "'a"),
bulwahn@35325
   196
  (@{const_name induct_conj}, "'a"),*)
bulwahn@34952
   197
  (@{const_name "undefined"}, "'a"),
bulwahn@34952
   198
  (@{const_name "default"}, "'a"),
bulwahn@36255
   199
  (@{const_name "dummy_pattern"}, "'a::{}"),
bulwahn@36255
   200
  (@{const_name "HOL.simp_implies"}, "prop => prop => prop"),
bulwahn@36255
   201
  (@{const_name "bot_fun_inst.bot_fun"}, "'a"),
bulwahn@36255
   202
  (@{const_name "top_fun_inst.top_fun"}, "'a"),
bulwahn@36255
   203
  (@{const_name "Pure.term"}, "'a"),
bulwahn@36255
   204
  (@{const_name "top_class.top"}, "'a"),
bulwahn@36255
   205
  (@{const_name "eq_class.eq"}, "'a"),
bulwahn@36255
   206
  (@{const_name "Quotient.Quot_True"}, "'a")(*,
bulwahn@34952
   207
  (@{const_name "uminus"}, "'a"),
bulwahn@34952
   208
  (@{const_name "Nat.size"}, "'a"),
haftmann@35092
   209
  (@{const_name "Groups.abs"}, "'a") *)]
bulwahn@34952
   210
bulwahn@34952
   211
val forbidden_thms =
bulwahn@34952
   212
 ["finite_intvl_succ_class",
bulwahn@34952
   213
  "nibble"]
bulwahn@34952
   214
bulwahn@34952
   215
val forbidden_consts =
bulwahn@34952
   216
 [@{const_name nibble_pair_of_char}]
bulwahn@34952
   217
bulwahn@34952
   218
fun is_forbidden_theorem (s, th) =
bulwahn@34952
   219
  let val consts = Term.add_const_names (prop_of th) [] in
haftmann@36677
   220
    exists (member (op =) (space_explode "." s)) forbidden_thms orelse
haftmann@36677
   221
    exists (member (op =) forbidden_consts) consts orelse
bulwahn@34952
   222
    length (space_explode "." s) <> 2 orelse
bulwahn@34952
   223
    String.isPrefix "type_definition" (List.last (space_explode "." s)) orelse
bulwahn@34952
   224
    String.isSuffix "_def" s orelse
bulwahn@34952
   225
    String.isSuffix "_raw" s
bulwahn@34952
   226
  end
bulwahn@34952
   227
bulwahn@36255
   228
val forbidden_mutant_constnames =
bulwahn@36255
   229
 ["HOL.induct_equal",
bulwahn@36255
   230
  "HOL.induct_implies",
bulwahn@36255
   231
  "HOL.induct_conj",
bulwahn@36255
   232
 @{const_name undefined},
bulwahn@36255
   233
 @{const_name default},
bulwahn@36255
   234
 @{const_name dummy_pattern},
bulwahn@36255
   235
 @{const_name "HOL.simp_implies"},
bulwahn@36255
   236
 @{const_name "bot_fun_inst.bot_fun"},
bulwahn@36255
   237
 @{const_name "top_fun_inst.top_fun"},
bulwahn@36255
   238
 @{const_name "Pure.term"},
bulwahn@36255
   239
 @{const_name "top_class.top"},
bulwahn@36255
   240
 @{const_name "eq_class.eq"},
bulwahn@36255
   241
 @{const_name "Quotient.Quot_True"}]
bulwahn@36255
   242
bulwahn@34952
   243
fun is_forbidden_mutant t =
bulwahn@36255
   244
  let
bulwahn@36255
   245
    val consts = Term.add_const_names t []
bulwahn@36255
   246
  in
bulwahn@34952
   247
    exists (String.isPrefix "Nitpick") consts orelse
bulwahn@36255
   248
    exists (String.isSubstring "_sumC") consts orelse
bulwahn@36255
   249
    exists (member (op =) forbidden_mutant_constnames) consts
bulwahn@34952
   250
  end
bulwahn@34952
   251
bulwahn@34952
   252
fun is_executable_term thy t = can (TimeLimit.timeLimit (Time.fromMilliseconds 2000) (Quickcheck.test_term
wenzelm@36633
   253
 (ProofContext.init_global thy) false (SOME "SML") 1 0)) (preprocess thy [] t)
bulwahn@34952
   254
fun is_executable_thm thy th = is_executable_term thy (prop_of th)
bulwahn@34952
   255
bulwahn@34952
   256
val freezeT =
bulwahn@34952
   257
  map_types (map_type_tvar (fn ((a, i), S) =>
bulwahn@34952
   258
    TFree (if i = 0 then a else a ^ "_" ^ string_of_int i, S)))
bulwahn@34952
   259
bulwahn@34952
   260
fun thms_of all thy =
bulwahn@34952
   261
  filter
bulwahn@34952
   262
    (fn th => (all orelse Context.theory_name (theory_of_thm th) = Context.theory_name thy)
bulwahn@34952
   263
      (* andalso is_executable_thm thy th *))
bulwahn@34952
   264
    (map snd (filter_out is_forbidden_theorem (Mutabelle.all_unconcealed_thms_of thy)))
bulwahn@34952
   265
bulwahn@34952
   266
val count = length oo filter o equal
bulwahn@34952
   267
bulwahn@34952
   268
fun take_random 0 _ = []
bulwahn@34952
   269
  | take_random _ [] = []
bulwahn@34952
   270
  | take_random n xs =
bulwahn@34952
   271
    let val j = random_range 0 (length xs - 1) in
bulwahn@34952
   272
      Library.nth xs j :: take_random (n - 1) (nth_drop j xs)
bulwahn@34952
   273
    end
bulwahn@34952
   274
bulwahn@35324
   275
fun cpu_time description f =
bulwahn@35324
   276
  let
bulwahn@35324
   277
    val start = start_timing ()
bulwahn@35324
   278
    val result = Exn.capture f ()
bulwahn@35324
   279
    val time = Time.toMilliseconds (#cpu (end_timing start))
bulwahn@35324
   280
  in (Exn.release result, (description, time)) end
bulwahn@35324
   281
bulwahn@34952
   282
fun safe_invoke_mtd thy (mtd_name, invoke_mtd) t =
bulwahn@34952
   283
  let
bulwahn@34952
   284
    val _ = priority ("Invoking " ^ mtd_name)
bulwahn@35380
   285
    val ((res, (timing, reports)), time) = cpu_time "total time"
bulwahn@35324
   286
      (fn () => case try (invoke_mtd thy) t of
bulwahn@35380
   287
          SOME (res, (timing, reports)) => (res, (timing, reports))
bulwahn@35324
   288
        | NONE => (priority ("**** PROBLEMS WITH " ^ Syntax.string_of_term_global thy t);
bulwahn@35380
   289
           (Error , ([], NONE))))
bulwahn@34952
   290
    val _ = priority (" Done")
bulwahn@35380
   291
  in (res, (time :: timing, reports)) end
bulwahn@34952
   292
bulwahn@34952
   293
(* theory -> term list -> mtd -> subentry *)
bulwahn@35324
   294
(*
bulwahn@34952
   295
fun test_mutants_using_one_method thy mutants (mtd_name, invoke_mtd) =
bulwahn@34952
   296
  let
bulwahn@34952
   297
     val res = map (safe_invoke_mtd thy (mtd_name, invoke_mtd)) mutants
bulwahn@34952
   298
  in
bulwahn@34952
   299
    (mtd_name, count GenuineCex res, count PotentialCex res, count NoCex res,
bulwahn@34952
   300
     count Donno res, count Timeout res, count Error res)
bulwahn@34952
   301
  end
bulwahn@34952
   302
bulwahn@34952
   303
fun create_entry thy thm exec mutants mtds =
wenzelm@36752
   304
  (Thm.get_name_hint thm, exec, map (test_mutants_using_one_method thy mutants) mtds)
bulwahn@35324
   305
*)
bulwahn@34952
   306
fun create_detailed_entry thy thm exec mutants mtds =
bulwahn@34952
   307
  let
bulwahn@34952
   308
    fun create_mutant_subentry mutant = (mutant,
bulwahn@34952
   309
      map (fn (mtd_name, invoke_mtd) =>
bulwahn@34952
   310
        (mtd_name, safe_invoke_mtd thy (mtd_name, invoke_mtd) mutant)) mtds)
bulwahn@34952
   311
  in
wenzelm@36752
   312
    (Thm.get_name_hint thm, exec, prop_of thm, map create_mutant_subentry mutants)
bulwahn@34952
   313
  end
bulwahn@34952
   314
bulwahn@34952
   315
(* (theory -> thm -> bool -> term list -> mtd list -> 'a) -> theory -> mtd list -> thm -> 'a *)
bulwahn@34952
   316
fun mutate_theorem create_entry thy mtds thm =
bulwahn@34952
   317
  let
bulwahn@34952
   318
    val pp = Syntax.pp_global thy
bulwahn@34952
   319
    val exec = is_executable_thm thy thm
bulwahn@34952
   320
    val _ = priority (if exec then "EXEC" else "NOEXEC")
bulwahn@34952
   321
    val mutants =
bulwahn@34952
   322
          (if num_mutations = 0 then
bulwahn@34952
   323
             [Thm.prop_of thm]
bulwahn@34952
   324
           else
bulwahn@34952
   325
             Mutabelle.mutate_mix (Thm.prop_of thm) thy comms forbidden
bulwahn@34952
   326
                                  num_mutations)
bulwahn@34952
   327
             |> filter_out is_forbidden_mutant
bulwahn@34952
   328
    val mutants =
bulwahn@34952
   329
      if exec then
bulwahn@34952
   330
        let
bulwahn@34952
   331
          val _ = priority ("BEFORE PARTITION OF " ^
bulwahn@34952
   332
                            Int.toString (length mutants) ^ " MUTANTS")
bulwahn@34952
   333
          val (execs, noexecs) = List.partition (is_executable_term thy) (take_random (20 * max_mutants) mutants)
bulwahn@34952
   334
          val _ = tracing ("AFTER PARTITION (" ^ Int.toString (length execs) ^
bulwahn@34952
   335
                           " vs " ^ Int.toString (length noexecs) ^ ")")
bulwahn@34952
   336
        in
bulwahn@34952
   337
          execs @ take_random (Int.max (0, max_mutants - length execs)) noexecs
bulwahn@34952
   338
        end
bulwahn@34952
   339
      else
bulwahn@34952
   340
        mutants
bulwahn@34952
   341
    val mutants = mutants
bulwahn@34952
   342
          |> take_random max_mutants
bulwahn@34952
   343
          |> map Mutabelle.freeze |> map freezeT
bulwahn@34952
   344
(*          |> filter (not o is_forbidden_mutant) *)
bulwahn@34952
   345
          |> List.mapPartial (try (Sign.cert_term thy))
bulwahn@34952
   346
    val _ = map (fn t => priority ("MUTANT: " ^ Pretty.string_of (Pretty.term pp t))) mutants
bulwahn@34952
   347
  in
bulwahn@34952
   348
    create_entry thy thm exec mutants mtds
bulwahn@34952
   349
  end
bulwahn@34952
   350
bulwahn@34952
   351
(* theory -> mtd list -> thm list -> report *)
bulwahn@34952
   352
val mutate_theorems = map ooo mutate_theorem
bulwahn@34952
   353
bulwahn@34952
   354
fun string_of_outcome GenuineCex = "GenuineCex"
bulwahn@34952
   355
  | string_of_outcome PotentialCex = "PotentialCex"
bulwahn@34952
   356
  | string_of_outcome NoCex = "NoCex"
bulwahn@34952
   357
  | string_of_outcome Donno = "Donno"
bulwahn@34952
   358
  | string_of_outcome Timeout = "Timeout"
bulwahn@34952
   359
  | string_of_outcome Error = "Error"
bulwahn@34952
   360
bulwahn@35324
   361
fun string_of_mutant_subentry thy thm_name (t, results) =
bulwahn@34952
   362
  "mutant: " ^ Syntax.string_of_term_global thy t ^ "\n" ^
bulwahn@35324
   363
  space_implode "; "
bulwahn@35324
   364
    (map (fn (mtd_name, (outcome, timing)) => mtd_name ^ ": " ^ string_of_outcome outcome) results) ^
bulwahn@34952
   365
  "\n"
bulwahn@34952
   366
bulwahn@36255
   367
(* XML.tree -> string *)
bulwahn@36255
   368
fun plain_string_from_xml_tree t =
bulwahn@36255
   369
  Buffer.empty |> XML.add_content t |> Buffer.content
bulwahn@36255
   370
(* string -> string *)
bulwahn@36255
   371
val unyxml = plain_string_from_xml_tree o YXML.parse
bulwahn@36255
   372
bulwahn@35324
   373
fun string_of_mutant_subentry' thy thm_name (t, results) =
bulwahn@35380
   374
  let
bulwahn@35380
   375
    fun string_of_report (Quickcheck.Report {iterations = i, raised_match_errors = e,
bulwahn@35380
   376
      satisfied_assms = s, positive_concl_tests = p}) =
bulwahn@35380
   377
      "errors: " ^ string_of_int e ^ "; conclusion tests: " ^ string_of_int p
bulwahn@35380
   378
    fun string_of_reports NONE = ""
bulwahn@35380
   379
      | string_of_reports (SOME reports) =
bulwahn@35380
   380
        cat_lines (map (fn (size, [report]) =>
bulwahn@35380
   381
          "size " ^ string_of_int size ^ ": " ^ string_of_report report) (rev reports))
bulwahn@35380
   382
    fun string_of_mtd_result (mtd_name, (outcome, (timing, reports))) =
bulwahn@35324
   383
      mtd_name ^ ": " ^ string_of_outcome outcome ^ "; " ^
bulwahn@35380
   384
      space_implode "; " (map (fn (s, t) => (s ^ ": " ^ string_of_int t)) timing)
bulwahn@36255
   385
      (*^ "\n" ^ string_of_reports reports*)
bulwahn@35380
   386
  in
bulwahn@36255
   387
    "mutant of " ^ thm_name ^ ":\n"
bulwahn@36255
   388
    ^ unyxml (Syntax.string_of_term_global thy t) ^ "\n" ^ cat_lines (map string_of_mtd_result results)
bulwahn@35380
   389
  end
bulwahn@35324
   390
bulwahn@34952
   391
fun string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) = 
bulwahn@34952
   392
   thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ": " ^
bulwahn@36255
   393
   Syntax.string_of_term_global thy t ^ "\n" ^                                    
bulwahn@35324
   394
   cat_lines (map (string_of_mutant_subentry' thy thm_name) mutant_subentries) ^ "\n"
bulwahn@34952
   395
bulwahn@36255
   396
fun theoryfile_string_of_mutant_subentry thy thm_name (i, (t, results)) =
bulwahn@36255
   397
  "lemma " ^ thm_name ^ "_" ^ string_of_int (i + 1) ^ ":\n" ^
bulwahn@36255
   398
  "\"" ^ unyxml (Syntax.string_of_term_global thy t) ^
bulwahn@36255
   399
  "\" \nquickcheck[generator = SML]\nquickcheck[generator = predicate_compile_wo_ff]\n" ^
bulwahn@36255
   400
  "quickcheck[generator = predicate_compile_ff_nofs]\noops\n"
bulwahn@36255
   401
bulwahn@36255
   402
fun theoryfile_string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) =
bulwahn@36255
   403
  "subsubsection {* mutants of " ^ thm_name ^ " *}\n\n" ^
bulwahn@36255
   404
  cat_lines (map_index
bulwahn@36255
   405
    (theoryfile_string_of_mutant_subentry thy thm_name) mutant_subentries) ^ "\n"
bulwahn@36255
   406
bulwahn@34952
   407
(* subentry -> string *)
bulwahn@34952
   408
fun string_for_subentry (mtd_name, genuine_cex, potential_cex, no_cex, donno,
bulwahn@34952
   409
                         timeout, error) =
bulwahn@34952
   410
  "    " ^ mtd_name ^ ": " ^ Int.toString genuine_cex ^ "+ " ^
bulwahn@34952
   411
  Int.toString potential_cex ^ "= " ^ Int.toString no_cex ^ "- " ^
bulwahn@34952
   412
  Int.toString donno ^ "? " ^ Int.toString timeout ^ "T " ^
bulwahn@34952
   413
  Int.toString error ^ "!"
bulwahn@34952
   414
(* entry -> string *)
bulwahn@34952
   415
fun string_for_entry (thm_name, exec, subentries) =
bulwahn@34952
   416
  thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ":\n" ^
bulwahn@34952
   417
  cat_lines (map string_for_subentry subentries) ^ "\n"
bulwahn@34952
   418
(* report -> string *)
bulwahn@34952
   419
fun string_for_report report = cat_lines (map string_for_entry report)
bulwahn@34952
   420
bulwahn@34952
   421
(* string -> report -> unit *)
bulwahn@34952
   422
fun write_report file_name =
bulwahn@34952
   423
  File.write (Path.explode file_name) o string_for_report
bulwahn@34952
   424
bulwahn@34952
   425
(* theory -> mtd list -> thm list -> string -> unit *)
bulwahn@34952
   426
fun mutate_theorems_and_write_report thy mtds thms file_name =
bulwahn@34952
   427
  let
bulwahn@34952
   428
    val _ = priority "Starting Mutabelle..."
bulwahn@34952
   429
    val path = Path.explode file_name
bulwahn@34952
   430
    (* for normal report: *)
bulwahn@34952
   431
    (*val (gen_create_entry, gen_string_for_entry) = (create_entry, string_for_entry)*)
bulwahn@34952
   432
    (*for detailled report: *)
bulwahn@34952
   433
    val (gen_create_entry, gen_string_for_entry) =
bulwahn@34952
   434
      (create_detailed_entry, string_of_detailed_entry thy)
bulwahn@36255
   435
    val (gen_create_entry, gen_string_for_entry) =
bulwahn@36255
   436
      (create_detailed_entry, theoryfile_string_of_detailed_entry thy)
bulwahn@34952
   437
  in
bulwahn@34952
   438
    File.write path (
bulwahn@34952
   439
    "Mutation options = "  ^
bulwahn@34952
   440
      "max_mutants: " ^ string_of_int max_mutants ^
bulwahn@34952
   441
      "; num_mutations: " ^ string_of_int num_mutations ^ "\n" ^
bulwahn@34952
   442
    "QC options = " ^
bulwahn@34952
   443
      (*"quickcheck_generator: " ^ quickcheck_generator ^ ";*)
bulwahn@34952
   444
      "size: " ^ string_of_int size ^
bulwahn@34952
   445
      "; iterations: " ^ string_of_int iterations ^ "\n");
bulwahn@34952
   446
    map (File.append path o gen_string_for_entry o mutate_theorem gen_create_entry thy mtds) thms;
bulwahn@34952
   447
    ()
bulwahn@34952
   448
  end
bulwahn@34952
   449
bulwahn@34952
   450
end;