src/HOL/Mutabelle/mutabelle_extra.ML
author bulwahn
Mon, 06 Dec 2010 10:52:44 +0100
changeset 41218 ce78ef6a909b
parent 41217 6604115019bf
child 41220 29e5cae93584
permissions -rw-r--r--
adding timeout to try invocation in mutabelle
     1 (*  Title:      HOL/Mutabelle/mutabelle_extra.ML
     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 
    11 datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error | Solved | Unsolved
    12 type timing = (string * int) list
    13 
    14 type mtd = string * (theory -> term -> outcome * (timing * (int * Quickcheck.report list) list option))
    15 
    16 type mutant_subentry = term * (string * (outcome * (timing * Quickcheck.report option))) list
    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 
    23 val quickcheck_mtd : (Proof.context -> Proof.context) -> string -> mtd
    24 
    25 val solve_direct_mtd : mtd
    26 val try_mtd : mtd
    27 
    28 val sledgehammer_mtd : mtd
    29 
    30 (*
    31 val refute_mtd : mtd
    32 val nitpick_mtd : mtd
    33 *)
    34 
    35 val freezeT : term -> term
    36 val thms_of : bool -> theory -> thm list
    37 
    38 val string_for_report : report -> string
    39 val write_report : string -> report -> unit
    40 val mutate_theorems_and_write_report :
    41   theory -> mtd list -> thm list -> string -> unit
    42 
    43 val random_seed : real Unsynchronized.ref
    44 end;
    45 
    46 structure MutabelleExtra : MUTABELLE_EXTRA =
    47 struct
    48 
    49 (* Own seed; can't rely on the Isabelle one to stay the same *)
    50 val random_seed = Unsynchronized.ref 1.0;
    51 
    52 
    53 (* mutation options *)
    54 (*val max_mutants = 4
    55 val num_mutations = 1*)
    56 (* soundness check: *)
    57 val max_mutants =  10
    58 val num_mutations = 1
    59 
    60 (* quickcheck options *)
    61 (*val quickcheck_generator = "SML"*)
    62 
    63 (* Another Random engine *)
    64 
    65 exception RANDOM;
    66 
    67 fun rmod x y = x - y * Real.realFloor (x / y);
    68 
    69 local
    70   val a = 16807.0;
    71   val m = 2147483647.0;
    72 in
    73 
    74 fun random () = CRITICAL (fn () =>
    75   let val r = rmod (a * ! random_seed) m
    76   in (random_seed := r; r) end);
    77 
    78 end;
    79 
    80 fun random_range l h =
    81   if h < l orelse l < 0 then raise RANDOM
    82   else l + Real.floor (rmod (random ()) (real (h - l + 1)));
    83 
    84 fun take_random 0 _ = []
    85   | take_random _ [] = []
    86   | take_random n xs =
    87     let val j = random_range 0 (length xs - 1) in
    88       Library.nth xs j :: take_random (n - 1) (nth_drop j xs)
    89     end
    90   
    91 (* possible outcomes *)
    92 
    93 datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error | Solved | Unsolved
    94 
    95 fun string_of_outcome GenuineCex = "GenuineCex"
    96   | string_of_outcome PotentialCex = "PotentialCex"
    97   | string_of_outcome NoCex = "NoCex"
    98   | string_of_outcome Donno = "Donno"
    99   | string_of_outcome Timeout = "Timeout"
   100   | string_of_outcome Error = "Error"
   101   | string_of_outcome Solved = "Solved"
   102   | string_of_outcome Unsolved = "Unsolved"
   103 
   104 type timing = (string * int) list
   105 
   106 type mtd = string * (theory -> term -> outcome * (timing * (int * Quickcheck.report list) list option))
   107 
   108 type mutant_subentry = term * (string * (outcome * (timing * Quickcheck.report option))) list
   109 type detailed_entry = string * bool * term * mutant_subentry list
   110 
   111 type subentry = string * int * int * int * int * int * int
   112 type entry = string * bool * subentry list
   113 type report = entry list
   114 
   115 (* possible invocations *)
   116 
   117 (** quickcheck **)
   118 
   119 fun invoke_quickcheck change_options quickcheck_generator thy t =
   120   TimeLimit.timeLimit (seconds (!Auto_Tools.time_limit))
   121       (fn _ =>
   122           case Quickcheck.test_goal_terms (change_options (ProofContext.init_global thy))
   123               false [] [t] of
   124             (NONE, _) => (NoCex, ([], NONE))
   125           | (SOME _, _) => (GenuineCex, ([], NONE))) ()
   126   handle TimeLimit.TimeOut =>
   127          (Timeout, ([("timelimit", Real.floor (!Auto_Tools.time_limit))], NONE))
   128 
   129 fun quickcheck_mtd change_options quickcheck_generator =
   130   ("quickcheck_" ^ quickcheck_generator, invoke_quickcheck change_options quickcheck_generator)
   131 
   132 (** solve direct **)
   133  
   134 fun invoke_solve_direct thy t =
   135   let
   136     val state = Proof.theorem NONE (K I) (map (single o rpair []) [t]) (ProofContext.init_global thy) 
   137   in
   138     case Solve_Direct.solve_direct false state of
   139       (true, _) => (Solved, ([], NONE))
   140     | (false, _) => (Unsolved, ([], NONE))
   141   end
   142 
   143 val solve_direct_mtd = ("solve_direct", invoke_solve_direct) 
   144 
   145 (** try **)
   146 
   147 fun invoke_try thy t =
   148   let
   149     val state = Proof.theorem NONE (K I) (map (single o rpair []) [t]) (ProofContext.init_global thy)
   150   in
   151     case Try.invoke_try (SOME (seconds 5.0)) state of
   152       true => (Solved, ([], NONE))
   153     | false => (Unsolved, ([], NONE))
   154   end
   155 
   156 val try_mtd = ("try", invoke_try)
   157 
   158 (** sledgehammer **)
   159 
   160 fun invoke_sledgehammer thy t =
   161   if can (Goal.prove_global thy (Term.add_free_names t [])  [] t)
   162       (fn {context, ...} => Sledgehammer_Tactics.sledgehammer_with_metis_tac context 1) then
   163     (Solved, ([], NONE))
   164   else
   165     (Unsolved, ([], NONE))
   166 
   167 val sledgehammer_mtd = ("sledgehammer", invoke_sledgehammer)
   168 
   169 (*
   170 fun invoke_refute thy t =
   171   let
   172     val res = MyRefute.refute_term thy [] t
   173     val _ = Output.urgent_message ("Refute: " ^ res)
   174   in
   175     case res of
   176       "genuine" => GenuineCex
   177     | "likely_genuine" => GenuineCex
   178     | "potential" => PotentialCex
   179     | "none" => NoCex
   180     | "unknown" => Donno
   181     | _ => Error
   182   end
   183   handle MyRefute.REFUTE (loc, details) =>
   184          (error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^
   185                    "."))
   186 val refute_mtd = ("refute", invoke_refute)
   187 *)
   188 
   189 (*
   190 open Nitpick_Util
   191 open Nitpick_Rep
   192 open Nitpick_Nut
   193 
   194 fun invoke_nitpick thy t =
   195   let
   196     val ctxt = ProofContext.init_global thy
   197     val state = Proof.init ctxt
   198   in
   199     let
   200       val (res, _) = Nitpick.pick_nits_in_term state (Nitpick_Isar.default_params thy []) false [] t
   201       val _ = Output.urgent_message ("Nitpick: " ^ res)
   202     in
   203       case res of
   204         "genuine" => GenuineCex
   205       | "likely_genuine" => GenuineCex
   206       | "potential" => PotentialCex
   207       | "none" => NoCex
   208       | "unknown" => Donno
   209       | _ => Error
   210     end
   211     handle ARG (loc, details) =>
   212            (error ("Bad argument(s) to " ^ quote loc ^ ": " ^ details ^ "."))
   213          | BAD (loc, details) =>
   214            (error ("Internal error (" ^ quote loc ^ "): " ^ details ^ "."))
   215          | LIMIT (_, details) =>
   216            (warning ("Limit reached: " ^ details ^ "."); Donno)
   217          | NOT_SUPPORTED details =>
   218            (warning ("Unsupported case: " ^ details ^ "."); Donno)
   219          | NUT (loc, us) =>
   220            (error ("Invalid nut" ^ plural_s_for_list us ^
   221                    " (" ^ quote loc ^ "): " ^
   222                   commas (map (string_for_nut ctxt) us) ^ "."))
   223          | REP (loc, Rs) =>
   224            (error ("Invalid representation" ^ plural_s_for_list Rs ^
   225                    " (" ^ quote loc ^ "): " ^
   226                    commas (map string_for_rep Rs) ^ "."))
   227          | TERM (loc, ts) =>
   228            (error ("Invalid term" ^ plural_s_for_list ts ^
   229                    " (" ^ quote loc ^ "): " ^
   230                    commas (map (Syntax.string_of_term ctxt) ts) ^ "."))
   231          | TYPE (loc, Ts, ts) =>
   232            (error ("Invalid type" ^ plural_s_for_list Ts ^
   233                    (if null ts then
   234                       ""
   235                     else
   236                       " for term" ^ plural_s_for_list ts ^ " " ^
   237                       commas (map (quote o Syntax.string_of_term ctxt) ts)) ^
   238                    " (" ^ quote loc ^ "): " ^
   239                    commas (map (Syntax.string_of_typ ctxt) Ts) ^ "."))
   240          | Kodkod.SYNTAX (_, details) =>
   241            (warning ("Ill-formed Kodkodi output: " ^ details ^ "."); Error)
   242          | Refute.REFUTE (loc, details) =>
   243            (error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^
   244                    "."))
   245          | Exn.Interrupt => raise Exn.Interrupt  (* FIXME violates Isabelle/ML exception model *)
   246          | _ => (Output.urgent_message ("Unknown error in Nitpick"); Error)
   247   end
   248 val nitpick_mtd = ("nitpick", invoke_nitpick)
   249 *)
   250 
   251 (* filtering forbidden theorems and mutants *)
   252 
   253 val comms = [@{const_name HOL.eq}, @{const_name HOL.disj}, @{const_name HOL.conj}]
   254 
   255 val forbidden =
   256  [(* (@{const_name "power"}, "'a"), *)
   257   (*(@{const_name induct_equal}, "'a"),
   258   (@{const_name induct_implies}, "'a"),
   259   (@{const_name induct_conj}, "'a"),*)
   260   (@{const_name "undefined"}, "'a"),
   261   (@{const_name "default"}, "'a"),
   262   (@{const_name "dummy_pattern"}, "'a::{}"),
   263   (@{const_name "HOL.simp_implies"}, "prop => prop => prop"),
   264   (@{const_name "bot_fun_inst.bot_fun"}, "'a"),
   265   (@{const_name "top_fun_inst.top_fun"}, "'a"),
   266   (@{const_name "Pure.term"}, "'a"),
   267   (@{const_name "top_class.top"}, "'a"),
   268   (@{const_name "Quotient.Quot_True"}, "'a")(*,
   269   (@{const_name "uminus"}, "'a"),
   270   (@{const_name "Nat.size"}, "'a"),
   271   (@{const_name "Groups.abs"}, "'a") *)]
   272 
   273 val forbidden_thms =
   274  ["finite_intvl_succ_class",
   275   "nibble"]
   276 
   277 val forbidden_consts =
   278  [@{const_name nibble_pair_of_char}, @{const_name "TYPE"}]
   279 
   280 fun is_forbidden_theorem (s, th) =
   281   let val consts = Term.add_const_names (prop_of th) [] in
   282     exists (member (op =) (space_explode "." s)) forbidden_thms orelse
   283     exists (member (op =) forbidden_consts) consts orelse
   284     length (space_explode "." s) <> 2 orelse
   285     String.isPrefix "type_definition" (List.last (space_explode "." s)) orelse
   286     String.isSuffix "_def" s orelse
   287     String.isSuffix "_raw" s orelse
   288     String.isPrefix "term_of" (List.last (space_explode "." s))
   289   end
   290 
   291 val forbidden_mutant_constnames =
   292  ["HOL.induct_equal",
   293   "HOL.induct_implies",
   294   "HOL.induct_conj",
   295  @{const_name undefined},
   296  @{const_name default},
   297  @{const_name dummy_pattern},
   298  @{const_name "HOL.simp_implies"},
   299  @{const_name "bot_fun_inst.bot_fun"},
   300  @{const_name "top_fun_inst.top_fun"},
   301  @{const_name "Pure.term"},
   302  @{const_name "top_class.top"},
   303  (*@{const_name "HOL.equal"},*)
   304  @{const_name "Quotient.Quot_True"},
   305  @{const_name "equal_fun_inst.equal_fun"},
   306  @{const_name "equal_bool_inst.equal_bool"},
   307  @{const_name "ord_fun_inst.less_eq_fun"},
   308  @{const_name "ord_fun_inst.less_fun"},
   309  @{const_name Metis.fequal},
   310  @{const_name Meson.skolem},
   311  @{const_name transfer_morphism}
   312  (*@{const_name "==>"}, @{const_name "=="}*)]
   313 
   314 val forbidden_mutant_consts =
   315   [
   316    (@{const_name "Groups.zero_class.zero"}, @{typ "prop => prop => prop"}),
   317    (@{const_name "Groups.one_class.one"}, @{typ "prop => prop => prop"}),
   318    (@{const_name "Groups.plus_class.plus"}, @{typ "prop => prop => prop"}),
   319    (@{const_name "Groups.minus_class.minus"}, @{typ "prop => prop => prop"}),
   320    (@{const_name "Groups.times_class.times"}, @{typ "prop => prop => prop"}),
   321    (@{const_name "Rings.inverse_class.divide"}, @{typ "prop => prop => prop"}),
   322    (@{const_name "Lattices.semilattice_inf_class.inf"}, @{typ "prop => prop => prop"}),
   323    (@{const_name "Lattices.semilattice_sup_class.sup"}, @{typ "prop => prop => prop"}),
   324    (@{const_name "Orderings.bot_class.bot"}, @{typ "prop => prop => prop"}),
   325    (@{const_name "Orderings.ord_class.min"}, @{typ "prop => prop => prop"}),
   326    (@{const_name "Orderings.ord_class.max"}, @{typ "prop => prop => prop"}),
   327    (@{const_name "Divides.div_class.mod"}, @{typ "prop => prop => prop"}),
   328    (@{const_name "Divides.div_class.div"}, @{typ "prop => prop => prop"}),
   329    (@{const_name "GCD.gcd_class.gcd"}, @{typ "prop => prop => prop"}),
   330    (@{const_name "GCD.gcd_class.lcm"}, @{typ "prop => prop => prop"}),
   331    (@{const_name "Orderings.bot_class.bot"}, @{typ "bool => prop"}),
   332    (@{const_name "Groups.one_class.one"}, @{typ "bool => prop"}),
   333    (@{const_name "Groups.zero_class.zero"},@{typ "bool => prop"})]
   334 
   335 fun is_forbidden_mutant t =
   336   let
   337     val const_names = Term.add_const_names t []
   338     val consts = Term.add_consts t []
   339   in
   340     exists (String.isPrefix "Nitpick") const_names orelse
   341     exists (String.isSubstring "_sumC") const_names orelse
   342     exists (member (op =) forbidden_mutant_constnames) const_names orelse
   343     exists (member (op =) forbidden_mutant_consts) consts
   344   end
   345 
   346 (* executable via quickcheck *)
   347 
   348 fun is_executable_term thy t =
   349   let
   350     val ctxt = ProofContext.init_global thy
   351   in
   352     can (TimeLimit.timeLimit (seconds 2.0)
   353       (Quickcheck.test_goal_terms
   354         ((Config.put Quickcheck.finite_types true #>
   355           Config.put Quickcheck.size 1 #> Config.put Quickcheck.iterations 1) ctxt)
   356         false [])) (map (Object_Logic.atomize_term thy) (fst (Variable.import_terms true [t] ctxt)))
   357   end
   358 
   359 fun is_executable_thm thy th = is_executable_term thy (prop_of th)
   360 
   361 val freezeT =
   362   map_types (map_type_tvar (fn ((a, i), S) =>
   363     TFree (if i = 0 then a else a ^ "_" ^ string_of_int i, S)))
   364 
   365 fun thms_of all thy =
   366   filter
   367     (fn th => (all orelse Context.theory_name (theory_of_thm th) = Context.theory_name thy)
   368       (* andalso is_executable_thm thy th *))
   369     (map snd (filter_out is_forbidden_theorem (Mutabelle.all_unconcealed_thms_of thy)))
   370 
   371 val count = length oo filter o equal
   372 
   373 fun cpu_time description f =
   374   let
   375     val start = start_timing ()
   376     val result = Exn.capture f ()
   377     val time = Time.toMilliseconds (#cpu (end_timing start))
   378   in (Exn.release result, (description, time)) end
   379 (*
   380 fun unsafe_invoke_mtd thy (mtd_name, invoke_mtd) t =
   381   let
   382     val _ = Output.urgent_message ("Invoking " ^ mtd_name)
   383     val ((res, (timing, reports)), time) = cpu_time "total time" (fn () => invoke_mtd thy t
   384       handle ERROR s => (tracing s; (Error, ([], NONE))))
   385     val _ = Output.urgent_message (" Done")
   386   in (res, (time :: timing, reports)) end
   387 *)  
   388 fun safe_invoke_mtd thy (mtd_name, invoke_mtd) t =
   389   let
   390     val _ = Output.urgent_message ("Invoking " ^ mtd_name)
   391     val (res, (timing, reports)) = (*cpu_time "total time"
   392       (fn () => *)case try (invoke_mtd thy) t of
   393           SOME (res, (timing, reports)) => (res, (timing, reports))
   394         | NONE => (Output.urgent_message ("**** PROBLEMS WITH " ^ Syntax.string_of_term_global thy t);
   395            (Error, ([], NONE)))
   396     val _ = Output.urgent_message (" Done")
   397   in (res, (timing, reports)) end
   398 
   399 (* theory -> term list -> mtd -> subentry *)
   400 
   401 fun test_mutants_using_one_method thy mutants (mtd_name, invoke_mtd) =
   402   let
   403      val res = map (fst o safe_invoke_mtd thy (mtd_name, invoke_mtd)) mutants
   404   in
   405     (mtd_name, count GenuineCex res, count PotentialCex res, count NoCex res,
   406      count Donno res, count Timeout res, count Error res)
   407   end
   408 
   409 (* creating entries *)
   410 
   411 fun create_entry thy thm exec mutants mtds =
   412   (Thm.get_name_hint thm, exec, map (test_mutants_using_one_method thy mutants) mtds)
   413 
   414 fun create_detailed_entry thy thm exec mutants mtds =
   415   let
   416     fun create_mutant_subentry mutant = (mutant,
   417       map (fn (mtd_name, invoke_mtd) =>
   418         (mtd_name, safe_invoke_mtd thy (mtd_name, invoke_mtd) mutant)) mtds)
   419   in
   420     (Thm.get_name_hint thm, exec, prop_of thm, map create_mutant_subentry mutants)
   421   end
   422 
   423 (* (theory -> thm -> bool -> term list -> mtd list -> 'a) -> theory -> mtd list -> thm -> 'a *)
   424 fun mutate_theorem create_entry thy mtds thm =
   425   let
   426     val exec = is_executable_thm thy thm
   427     val _ = Output.tracing (if exec then "EXEC" else "NOEXEC")
   428     val mutants =
   429           (if num_mutations = 0 then
   430              [Thm.prop_of thm]
   431            else
   432              Mutabelle.mutate_mix (Thm.prop_of thm) thy comms forbidden
   433                                   num_mutations)
   434              |> tap (fn muts => tracing ("mutants: " ^ string_of_int (length muts)))
   435              |> filter_out is_forbidden_mutant
   436     val mutants =
   437       if exec then
   438         let
   439           val _ = Output.urgent_message ("BEFORE PARTITION OF " ^
   440                             Int.toString (length mutants) ^ " MUTANTS")
   441           val (execs, noexecs) = List.partition (is_executable_term thy) (take_random (20 * max_mutants) mutants)
   442           val _ = tracing ("AFTER PARTITION (" ^ Int.toString (length execs) ^
   443                            " vs " ^ Int.toString (length noexecs) ^ ")")
   444         in
   445           execs @ take_random (Int.max (0, max_mutants - length execs)) noexecs
   446         end
   447       else
   448         mutants
   449     val mutants = mutants
   450           |> map Mutabelle.freeze |> map freezeT
   451 (*          |> filter (not o is_forbidden_mutant) *)
   452           |> List.mapPartial (try (Sign.cert_term thy))
   453           |> List.filter (is_some o try (Thm.cterm_of thy))
   454           |> List.filter (is_some o try (Syntax.check_term (ProofContext.init_global thy)))
   455           |> take_random max_mutants
   456     val _ = map (fn t => Output.urgent_message ("MUTANT: " ^ Syntax.string_of_term_global thy t)) mutants
   457   in
   458     create_entry thy thm exec mutants mtds
   459   end
   460 
   461 (* theory -> mtd list -> thm list -> report *)
   462 val mutate_theorems = map ooo mutate_theorem
   463 
   464 fun string_of_mutant_subentry thy thm_name (t, results) =
   465   "mutant: " ^ Syntax.string_of_term_global thy t ^ "\n" ^
   466   space_implode "; "
   467     (map (fn (mtd_name, (outcome, timing)) => mtd_name ^ ": " ^ string_of_outcome outcome) results) ^
   468   "\n"
   469 
   470 (* string -> string *)
   471 val unyxml = XML.content_of o YXML.parse_body
   472 
   473 fun string_of_mutant_subentry' thy thm_name (t, results) =
   474   let
   475     fun string_of_report (Quickcheck.Report {iterations = i, raised_match_errors = e,
   476       satisfied_assms = s, positive_concl_tests = p}) =
   477       "errors: " ^ string_of_int e ^ "; conclusion tests: " ^ string_of_int p
   478     fun string_of_reports NONE = ""
   479       | string_of_reports (SOME reports) =
   480         cat_lines (map (fn (size, [report]) =>
   481           "size " ^ string_of_int size ^ ": " ^ string_of_report report) (rev reports))
   482     fun string_of_mtd_result (mtd_name, (outcome, (timing, reports))) =
   483       mtd_name ^ ": " ^ string_of_outcome outcome
   484       (*" with time " ^ " (" ^ space_implode "; " (map (fn (s, t) => (s ^ ": " ^ string_of_int t)) timing) ^ ")"*)
   485       (*^ "\n" ^ string_of_reports reports*)
   486   in
   487     "mutant of " ^ thm_name ^ ":\n"
   488     ^ unyxml (Syntax.string_of_term_global thy t) ^ "\n" ^ space_implode "; " (map string_of_mtd_result results)
   489   end
   490 
   491 fun string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) = 
   492    thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ": " ^
   493    Syntax.string_of_term_global thy t ^ "\n" ^                                    
   494    cat_lines (map (string_of_mutant_subentry' thy thm_name) mutant_subentries) ^ "\n"
   495 
   496 fun theoryfile_string_of_mutant_subentry thy thm_name (i, (t, results)) =
   497   "lemma " ^ thm_name ^ "_" ^ string_of_int (i + 1) ^ ":\n" ^
   498   "\"" ^ unyxml (Syntax.string_of_term_global thy t) ^
   499   "\" \nquickcheck\noops\n"
   500 
   501 fun theoryfile_string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) =
   502   "subsubsection {* mutants of " ^ thm_name ^ " *}\n\n" ^
   503   cat_lines (map_index
   504     (theoryfile_string_of_mutant_subentry thy thm_name) mutant_subentries) ^ "\n"
   505 
   506 (* subentry -> string *)
   507 fun string_for_subentry (mtd_name, genuine_cex, potential_cex, no_cex, donno,
   508                          timeout, error) =
   509   "    " ^ mtd_name ^ ": " ^ Int.toString genuine_cex ^ "+ " ^
   510   Int.toString potential_cex ^ "= " ^ Int.toString no_cex ^ "- " ^
   511   Int.toString donno ^ "? " ^ Int.toString timeout ^ "T " ^
   512   Int.toString error ^ "!"
   513 
   514 (* entry -> string *)
   515 fun string_for_entry (thm_name, exec, subentries) =
   516   thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ":\n" ^
   517   cat_lines (map string_for_subentry subentries) ^ "\n"
   518 
   519 (* report -> string *)
   520 fun string_for_report report = cat_lines (map string_for_entry report)
   521 
   522 (* string -> report -> unit *)
   523 fun write_report file_name =
   524   File.write (Path.explode file_name) o string_for_report
   525 
   526 (* theory -> mtd list -> thm list -> string -> unit *)
   527 fun mutate_theorems_and_write_report thy mtds thms file_name =
   528   let
   529     val _ = Output.urgent_message "Starting Mutabelle..."
   530     val ctxt = ProofContext.init_global thy
   531     val path = Path.explode file_name
   532     (* for normal report: *)
   533     (*
   534     val (gen_create_entry, gen_string_for_entry) = (create_entry, string_for_entry)
   535     *)
   536     (* for detailled report: *)
   537     val (gen_create_entry, gen_string_for_entry) = (create_detailed_entry, string_of_detailed_entry thy)
   538     (* for theory creation: *)
   539     (*val (gen_create_entry, gen_string_for_entry) = (create_detailed_entry, theoryfile_string_of_detailed_entry thy)*)
   540   in
   541     File.write path (
   542     "Mutation options = "  ^
   543       "max_mutants: " ^ string_of_int max_mutants ^
   544       "; num_mutations: " ^ string_of_int num_mutations ^ "\n" ^
   545     "QC options = " ^
   546       (*"quickcheck_generator: " ^ quickcheck_generator ^ ";*)
   547       "size: " ^ string_of_int (Config.get ctxt Quickcheck.size) ^
   548       "; iterations: " ^ string_of_int (Config.get ctxt Quickcheck.iterations) ^ "\n");
   549     map (File.append path o gen_string_for_entry o mutate_theorem gen_create_entry thy mtds) thms;
   550     ()
   551   end
   552 
   553 end;