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