src/HOL/Mutabelle/mutabelle_extra.ML
author bulwahn
Mon, 25 Jan 2010 16:19:42 +0100
changeset 34965 3b4762c1052c
child 34974 18b41bba42b5
permissions -rw-r--r--
adding Mutabelle to repository
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     1
(*
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     2
    Title:      HOL/Mutabelle/mutabelle_extra.ML
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     3
    Author:     Stefan Berghofer, Jasmin Blanchette, Lukas Bulwahn, TU Muenchen
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     4
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     5
    Invokation of Counterexample generators
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     6
*)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     7
signature MUTABELLE_EXTRA =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     8
sig
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     9
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    10
val take_random : int -> 'a list -> 'a list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    11
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    12
datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    13
type mtd = string * (theory -> term -> outcome)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    14
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    15
type mutant_subentry = term * (string * outcome) list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    16
type detailed_entry = string * bool * term * mutant_subentry list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    17
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    18
type subentry = string * int * int * int * int * int * int
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    19
type entry = string * bool * subentry list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    20
type report = entry list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    21
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    22
val quickcheck_mtd : string -> mtd
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    23
(*
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    24
val refute_mtd : mtd
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    25
val nitpick_mtd : mtd
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    26
*)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    27
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    28
val freezeT : term -> term
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    29
val thms_of : bool -> theory -> thm list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    30
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    31
val string_for_report : report -> string
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    32
val write_report : string -> report -> unit
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    33
val mutate_theorems_and_write_report :
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    34
  theory -> mtd list -> thm list -> string -> unit
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    35
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    36
val random_seed : real Unsynchronized.ref
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    37
end;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    38
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    39
structure MutabelleExtra : MUTABELLE_EXTRA =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    40
struct
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    41
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    42
(* Own seed; can't rely on the Isabelle one to stay the same *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    43
val random_seed = Unsynchronized.ref 1.0;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    44
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    45
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    46
(* mutation options *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    47
val max_mutants = 4
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    48
val num_mutations = 1
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    49
(* soundness check: *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    50
val max_mutants = 1
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    51
val num_mutations = 0
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    52
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    53
(* quickcheck options *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    54
(*val quickcheck_generator = "SML"*)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    55
val iterations = 100
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    56
val size = 5
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    57
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    58
exception RANDOM;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    59
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    60
fun rmod x y = x - y * Real.realFloor (x / y);
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    61
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    62
local
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    63
  val a = 16807.0;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    64
  val m = 2147483647.0;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    65
in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    66
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    67
fun random () = CRITICAL (fn () =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    68
  let val r = rmod (a * ! random_seed) m
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    69
  in (random_seed := r; r) end);
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    70
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    71
end;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    72
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    73
fun random_range l h =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    74
  if h < l orelse l < 0 then raise RANDOM
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    75
  else l + Real.floor (rmod (random ()) (real (h - l + 1)));
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    76
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    77
datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    78
type mtd = string * (theory -> term -> outcome)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    79
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    80
type mutant_subentry = term * (string * outcome) list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    81
type detailed_entry = string * bool * term * mutant_subentry list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    82
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    83
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    84
type subentry = string * int * int * int * int * int * int
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    85
type entry = string * bool * subentry list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    86
type report = entry list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    87
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    88
fun inst_type insts (Type (s, Ts)) = Type (s, map (inst_type insts) Ts)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    89
  | inst_type insts T = the_default HOLogic.intT (AList.lookup op = insts T)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    90
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    91
fun preprocess thy insts t = ObjectLogic.atomize_term thy
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    92
 (map_types (inst_type insts) (Mutabelle.freeze t));
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    93
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    94
fun invoke_quickcheck quickcheck_generator thy t =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    95
  TimeLimit.timeLimit (Time.fromSeconds (! Auto_Counterexample.time_limit))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    96
      (fn _ =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    97
          case Quickcheck.test_term (ProofContext.init thy) false (SOME quickcheck_generator)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    98
                                    size iterations (preprocess thy [] t) of
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    99
            NONE => NoCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   100
          | SOME _ => GenuineCex) ()
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   101
  handle TimeLimit.TimeOut => Timeout
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   102
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   103
fun quickcheck_mtd quickcheck_generator =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   104
  ("quickcheck_" ^ quickcheck_generator, invoke_quickcheck quickcheck_generator)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   105
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   106
  (*
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   107
fun invoke_refute thy t =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   108
  let
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   109
    val res = MyRefute.refute_term thy [] t
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   110
    val _ = priority ("Refute: " ^ res)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   111
  in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   112
    case res of
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   113
      "genuine" => GenuineCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   114
    | "likely_genuine" => GenuineCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   115
    | "potential" => PotentialCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   116
    | "none" => NoCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   117
    | "unknown" => Donno
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   118
    | _ => Error
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   119
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   120
  handle MyRefute.REFUTE (loc, details) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   121
         (error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   122
                   "."))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   123
val refute_mtd = ("refute", invoke_refute)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   124
*)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   125
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   126
(*
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   127
open Nitpick_Util
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   128
open Nitpick_Rep
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   129
open Nitpick_Nut
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   130
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   131
fun invoke_nitpick thy t =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   132
  let
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   133
    val ctxt = ProofContext.init thy
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   134
    val state = Proof.init ctxt
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   135
  in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   136
    let
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   137
      val (res, _) = Nitpick.pick_nits_in_term state (Nitpick_Isar.default_params thy []) false [] t
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   138
      val _ = priority ("Nitpick: " ^ res)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   139
    in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   140
      case res of
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   141
        "genuine" => GenuineCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   142
      | "likely_genuine" => GenuineCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   143
      | "potential" => PotentialCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   144
      | "none" => NoCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   145
      | "unknown" => Donno
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   146
      | _ => Error
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   147
    end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   148
    handle ARG (loc, details) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   149
           (error ("Bad argument(s) to " ^ quote loc ^ ": " ^ details ^ "."))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   150
         | BAD (loc, details) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   151
           (error ("Internal error (" ^ quote loc ^ "): " ^ details ^ "."))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   152
         | LIMIT (_, details) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   153
           (warning ("Limit reached: " ^ details ^ "."); Donno)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   154
         | NOT_SUPPORTED details =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   155
           (warning ("Unsupported case: " ^ details ^ "."); Donno)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   156
         | NUT (loc, us) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   157
           (error ("Invalid nut" ^ plural_s_for_list us ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   158
                   " (" ^ quote loc ^ "): " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   159
                  commas (map (string_for_nut ctxt) us) ^ "."))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   160
         | REP (loc, Rs) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   161
           (error ("Invalid representation" ^ plural_s_for_list Rs ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   162
                   " (" ^ quote loc ^ "): " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   163
                   commas (map string_for_rep Rs) ^ "."))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   164
         | TERM (loc, ts) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   165
           (error ("Invalid term" ^ plural_s_for_list ts ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   166
                   " (" ^ quote loc ^ "): " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   167
                   commas (map (Syntax.string_of_term ctxt) ts) ^ "."))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   168
         | TYPE (loc, Ts, ts) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   169
           (error ("Invalid type" ^ plural_s_for_list Ts ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   170
                   (if null ts then
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   171
                      ""
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   172
                    else
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   173
                      " for term" ^ plural_s_for_list ts ^ " " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   174
                      commas (map (quote o Syntax.string_of_term ctxt) ts)) ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   175
                   " (" ^ quote loc ^ "): " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   176
                   commas (map (Syntax.string_of_typ ctxt) Ts) ^ "."))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   177
         | Kodkod.SYNTAX (_, details) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   178
           (warning ("Ill-formed Kodkodi output: " ^ details ^ "."); Error)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   179
         | Refute.REFUTE (loc, details) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   180
           (error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   181
                   "."))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   182
         | Exn.Interrupt => raise Exn.Interrupt
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   183
         | _ => (priority ("Unknown error in Nitpick"); Error)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   184
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   185
val nitpick_mtd = ("nitpick", invoke_nitpick)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   186
*)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   187
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   188
val comms = [@{const_name "op ="}, @{const_name "op |"}, @{const_name "op &"}]
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   189
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   190
val forbidden =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   191
 [(* (@{const_name "power"}, "'a"), *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   192
  ("HOL.induct_equal", "'a"),
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   193
  ("HOL.induct_implies", "'a"),
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   194
  ("HOL.induct_conj", "'a"),
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   195
  (@{const_name "undefined"}, "'a"),
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   196
  (@{const_name "default"}, "'a"),
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   197
  (@{const_name "dummy_pattern"}, "'a::{}") (*,
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   198
  (@{const_name "uminus"}, "'a"),
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   199
  (@{const_name "Nat.size"}, "'a"),
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   200
  (@{const_name "HOL.abs"}, "'a") *)]
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   201
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   202
val forbidden_thms =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   203
 ["finite_intvl_succ_class",
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   204
  "nibble"]
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   205
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   206
val forbidden_consts =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   207
 [@{const_name nibble_pair_of_char}]
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   208
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   209
fun is_forbidden_theorem (s, th) =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   210
  let val consts = Term.add_const_names (prop_of th) [] in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   211
    exists (fn s' => s' mem space_explode "." s) forbidden_thms orelse
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   212
    exists (fn s' => s' mem forbidden_consts) consts orelse
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   213
    length (space_explode "." s) <> 2 orelse
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   214
    String.isPrefix "type_definition" (List.last (space_explode "." s)) orelse
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   215
    String.isSuffix "_def" s orelse
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   216
    String.isSuffix "_raw" s
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   217
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   218
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   219
fun is_forbidden_mutant t =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   220
  let val consts = Term.add_const_names t [] in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   221
    exists (String.isPrefix "Nitpick") consts orelse
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   222
    exists (String.isSubstring "_sumC") consts (* internal function *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   223
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   224
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   225
fun is_executable_term thy t = can (TimeLimit.timeLimit (Time.fromMilliseconds 2000) (Quickcheck.test_term
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   226
 (ProofContext.init thy) false (SOME "SML") 1 0)) (preprocess thy [] t)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   227
fun is_executable_thm thy th = is_executable_term thy (prop_of th)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   228
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   229
val freezeT =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   230
  map_types (map_type_tvar (fn ((a, i), S) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   231
    TFree (if i = 0 then a else a ^ "_" ^ string_of_int i, S)))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   232
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   233
fun thms_of all thy =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   234
  filter
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   235
    (fn th => (all orelse Context.theory_name (theory_of_thm th) = Context.theory_name thy)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   236
      (* andalso is_executable_thm thy th *))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   237
    (map snd (filter_out is_forbidden_theorem (Mutabelle.all_unconcealed_thms_of thy)))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   238
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   239
val count = length oo filter o equal
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   240
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   241
fun take_random 0 _ = []
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   242
  | take_random _ [] = []
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   243
  | take_random n xs =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   244
    let val j = random_range 0 (length xs - 1) in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   245
      Library.nth xs j :: take_random (n - 1) (nth_drop j xs)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   246
    end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   247
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   248
fun safe_invoke_mtd thy (mtd_name, invoke_mtd) t =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   249
  let
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   250
    val _ = priority ("Invoking " ^ mtd_name)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   251
    val res = case try (invoke_mtd thy) t of
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   252
                SOME res => res
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   253
              | NONE => (priority ("**** PROBLEMS WITH " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   254
                                 Syntax.string_of_term_global thy t); Error)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   255
    val _ = priority (" Done")
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   256
  in res end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   257
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   258
(* theory -> term list -> mtd -> subentry *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   259
fun test_mutants_using_one_method thy mutants (mtd_name, invoke_mtd) =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   260
  let
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   261
     val res = map (safe_invoke_mtd thy (mtd_name, invoke_mtd)) mutants
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   262
  in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   263
    (mtd_name, count GenuineCex res, count PotentialCex res, count NoCex res,
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   264
     count Donno res, count Timeout res, count Error res)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   265
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   266
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   267
fun create_entry thy thm exec mutants mtds =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   268
  (Thm.get_name thm, exec, map (test_mutants_using_one_method thy mutants) mtds)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   269
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   270
fun create_detailed_entry thy thm exec mutants mtds =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   271
  let
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   272
    fun create_mutant_subentry mutant = (mutant,
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   273
      map (fn (mtd_name, invoke_mtd) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   274
        (mtd_name, safe_invoke_mtd thy (mtd_name, invoke_mtd) mutant)) mtds)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   275
  in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   276
    (Thm.get_name thm, exec, prop_of thm, map create_mutant_subentry mutants)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   277
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   278
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   279
(* (theory -> thm -> bool -> term list -> mtd list -> 'a) -> theory -> mtd list -> thm -> 'a *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   280
fun mutate_theorem create_entry thy mtds thm =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   281
  let
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   282
    val pp = Syntax.pp_global thy
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   283
    val exec = is_executable_thm thy thm
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   284
    val _ = priority (if exec then "EXEC" else "NOEXEC")
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   285
    val mutants =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   286
          (if num_mutations = 0 then
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   287
             [Thm.prop_of thm]
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   288
           else
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   289
             Mutabelle.mutate_mix (Thm.prop_of thm) thy comms forbidden
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   290
                                  num_mutations)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   291
             |> filter_out is_forbidden_mutant
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   292
    val mutants =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   293
      if exec then
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   294
        let
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   295
          val _ = priority ("BEFORE PARTITION OF " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   296
                            Int.toString (length mutants) ^ " MUTANTS")
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   297
          val (execs, noexecs) = List.partition (is_executable_term thy) (take_random (20 * max_mutants) mutants)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   298
          val _ = tracing ("AFTER PARTITION (" ^ Int.toString (length execs) ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   299
                           " vs " ^ Int.toString (length noexecs) ^ ")")
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   300
        in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   301
          execs @ take_random (Int.max (0, max_mutants - length execs)) noexecs
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   302
        end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   303
      else
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   304
        mutants
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   305
    val mutants = mutants
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   306
          |> take_random max_mutants
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   307
          |> map Mutabelle.freeze |> map freezeT
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   308
(*          |> filter (not o is_forbidden_mutant) *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   309
          |> List.mapPartial (try (Sign.cert_term thy))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   310
    val _ = map (fn t => priority ("MUTANT: " ^ Pretty.string_of (Pretty.term pp t))) mutants
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   311
  in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   312
    create_entry thy thm exec mutants mtds
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   313
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   314
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   315
(* theory -> mtd list -> thm list -> report *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   316
val mutate_theorems = map ooo mutate_theorem
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   317
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   318
fun string_of_outcome GenuineCex = "GenuineCex"
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   319
  | string_of_outcome PotentialCex = "PotentialCex"
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   320
  | string_of_outcome NoCex = "NoCex"
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   321
  | string_of_outcome Donno = "Donno"
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   322
  | string_of_outcome Timeout = "Timeout"
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   323
  | string_of_outcome Error = "Error"
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   324
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   325
fun string_of_mutant_subentry thy (t, results) =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   326
  "mutant: " ^ Syntax.string_of_term_global thy t ^ "\n" ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   327
  space_implode "; " (map (fn (mtd_name, outcome) => mtd_name ^ ": " ^ string_of_outcome outcome) results) ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   328
  "\n"
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   329
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   330
fun string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) = 
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   331
   thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ": " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   332
   Syntax.string_of_term_global thy t ^ "\n" ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   333
   cat_lines (map (string_of_mutant_subentry thy) mutant_subentries) ^ "\n"
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   334
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   335
(* subentry -> string *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   336
fun string_for_subentry (mtd_name, genuine_cex, potential_cex, no_cex, donno,
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   337
                         timeout, error) =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   338
  "    " ^ mtd_name ^ ": " ^ Int.toString genuine_cex ^ "+ " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   339
  Int.toString potential_cex ^ "= " ^ Int.toString no_cex ^ "- " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   340
  Int.toString donno ^ "? " ^ Int.toString timeout ^ "T " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   341
  Int.toString error ^ "!"
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   342
(* entry -> string *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   343
fun string_for_entry (thm_name, exec, subentries) =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   344
  thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ":\n" ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   345
  cat_lines (map string_for_subentry subentries) ^ "\n"
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   346
(* report -> string *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   347
fun string_for_report report = cat_lines (map string_for_entry report)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   348
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   349
(* string -> report -> unit *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   350
fun write_report file_name =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   351
  File.write (Path.explode file_name) o string_for_report
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   352
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   353
(* theory -> mtd list -> thm list -> string -> unit *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   354
fun mutate_theorems_and_write_report thy mtds thms file_name =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   355
  let
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   356
    val _ = priority "Starting Mutabelle..."
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   357
    val path = Path.explode file_name
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   358
    (* for normal report: *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   359
    (*val (gen_create_entry, gen_string_for_entry) = (create_entry, string_for_entry)*)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   360
    (*for detailled report: *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   361
    val (gen_create_entry, gen_string_for_entry) =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   362
      (create_detailed_entry, string_of_detailed_entry thy)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   363
  in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   364
    File.write path (
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   365
    "Mutation options = "  ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   366
      "max_mutants: " ^ string_of_int max_mutants ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   367
      "; num_mutations: " ^ string_of_int num_mutations ^ "\n" ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   368
    "QC options = " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   369
      (*"quickcheck_generator: " ^ quickcheck_generator ^ ";*)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   370
      "size: " ^ string_of_int size ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   371
      "; iterations: " ^ string_of_int iterations ^ "\n");
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   372
    map (File.append path o gen_string_for_entry o mutate_theorem gen_create_entry thy mtds) thms;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   373
    ()
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   374
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   375
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   376
end;