src/HOL/Mutabelle/mutabelle_extra.ML
author bulwahn
Sun, 05 Feb 2012 17:43:14 +0100
changeset 46433 b67bb064de1e
parent 46376 110ba6344446
child 46434 6d2af424d0f8
permissions -rw-r--r--
mutabelle ignores theorems with internal constants
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37744
3daaf23b9ab4 tuned titles
haftmann
parents: 36743
diff changeset
     1
(*  Title:      HOL/Mutabelle/mutabelle_extra.ML
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     2
    Author:     Stefan Berghofer, Jasmin Blanchette, Lukas Bulwahn, TU Muenchen
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     3
41408
08a072ca6348 tuned comments;
wenzelm
parents: 41190
diff changeset
     4
Invokation of Counterexample generators.
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     5
*)
41408
08a072ca6348 tuned comments;
wenzelm
parents: 41190
diff changeset
     6
34965
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
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    12
datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error | Solved | Unsolved
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
    13
type timings = (string * int) list
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    14
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
    15
type mtd = string * (theory -> term -> outcome * timings)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35092
diff changeset
    16
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
    17
type mutant_subentry = term * (string * (outcome * timings)) list
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    18
type detailed_entry = string * bool * term * mutant_subentry list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    19
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    20
type subentry = string * int * int * int * int * int * int
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    21
type entry = string * bool * subentry list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    22
type report = entry list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    23
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    24
val quickcheck_mtd : (Proof.context -> Proof.context) -> string -> mtd
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    25
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    26
val solve_direct_mtd : mtd
43030
e1172791ad0d compile
blanchet
parents: 43018
diff changeset
    27
val try_methods_mtd : mtd
40974
29e5cae93584 commenting out sledgehammer_mtd in Mutabelle
bulwahn
parents: 40972
diff changeset
    28
(*
40971
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
    29
val sledgehammer_mtd : mtd
40974
29e5cae93584 commenting out sledgehammer_mtd in Mutabelle
bulwahn
parents: 40972
diff changeset
    30
*)
41190
0bdc6fac5f48 reactivating nitpick in Mutabelle
bulwahn
parents: 41106
diff changeset
    31
val nitpick_mtd : mtd
0bdc6fac5f48 reactivating nitpick in Mutabelle
bulwahn
parents: 41106
diff changeset
    32
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    33
val refute_mtd : mtd
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    34
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    35
val freezeT : term -> term
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    36
val thms_of : bool -> theory -> thm list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    37
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    38
val string_for_report : report -> string
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    39
val write_report : string -> report -> unit
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    40
val mutate_theorems_and_write_report :
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    41
  theory -> mtd list -> thm list -> string -> unit
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    42
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    43
val random_seed : real Unsynchronized.ref
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    44
end;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    45
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    46
structure MutabelleExtra : MUTABELLE_EXTRA =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    47
struct
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    48
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    49
(* Own seed; can't rely on the Isabelle one to stay the same *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    50
val random_seed = Unsynchronized.ref 1.0;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    51
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    52
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    53
(* mutation options *)
46310
8af202923906 more configurations to mutabelle
bulwahn
parents: 45428
diff changeset
    54
val max_mutants = 4
8af202923906 more configurations to mutabelle
bulwahn
parents: 45428
diff changeset
    55
val num_mutations = 1
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    56
(* soundness check: *)
46310
8af202923906 more configurations to mutabelle
bulwahn
parents: 45428
diff changeset
    57
(*val max_mutants =  10
8af202923906 more configurations to mutabelle
bulwahn
parents: 45428
diff changeset
    58
val num_mutations = 1*)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    59
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    60
(* quickcheck options *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    61
(*val quickcheck_generator = "SML"*)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    62
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    63
(* Another Random engine *)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    64
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    65
exception RANDOM;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    66
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    67
fun rmod x y = x - y * Real.realFloor (x / y);
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    68
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    69
local
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    70
  val a = 16807.0;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    71
  val m = 2147483647.0;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    72
in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    73
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    74
fun random () = CRITICAL (fn () =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    75
  let val r = rmod (a * ! random_seed) m
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    76
  in (random_seed := r; r) end);
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    77
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    78
end;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    79
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    80
fun random_range l h =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    81
  if h < l orelse l < 0 then raise RANDOM
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    82
  else l + Real.floor (rmod (random ()) (real (h - l + 1)));
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    83
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    84
fun take_random 0 _ = []
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    85
  | take_random _ [] = []
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    86
  | take_random n xs =
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    87
    let val j = random_range 0 (length xs - 1) in
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    88
      Library.nth xs j :: take_random (n - 1) (nth_drop j xs)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    89
    end
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    90
  
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    91
(* possible outcomes *)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    92
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    93
datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error | Solved | Unsolved
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    94
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    95
fun string_of_outcome GenuineCex = "GenuineCex"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    96
  | string_of_outcome PotentialCex = "PotentialCex"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    97
  | string_of_outcome NoCex = "NoCex"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    98
  | string_of_outcome Donno = "Donno"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
    99
  | string_of_outcome Timeout = "Timeout"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   100
  | string_of_outcome Error = "Error"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   101
  | string_of_outcome Solved = "Solved"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   102
  | string_of_outcome Unsolved = "Unsolved"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   103
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   104
type timings = (string * int) list
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   105
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   106
type mtd = string * (theory -> term -> outcome * timings)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35092
diff changeset
   107
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   108
type mutant_subentry = term * (string * (outcome * timings)) list
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   109
type detailed_entry = string * bool * term * mutant_subentry list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   110
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   111
type subentry = string * int * int * int * int * int * int
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   112
type entry = string * bool * subentry list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   113
type report = entry list
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   114
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   115
(* possible invocations *)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   116
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   117
(** quickcheck **)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   118
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   119
fun invoke_quickcheck change_options quickcheck_generator thy t =
43018
121aa59b4d17 renamed "Auto_Tools" "Try"
blanchet
parents: 42438
diff changeset
   120
  TimeLimit.timeLimit (seconds (!Try.auto_time_limit))
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   121
      (fn _ =>
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   122
          let
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 44845
diff changeset
   123
            val ctxt' = change_options (Proof_Context.init_global thy)
46376
110ba6344446 mutabelle must handle the case where quickcheck returns multiple results
bulwahn
parents: 46326
diff changeset
   124
            val (result :: _) = case Quickcheck.active_testers ctxt' of
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 44845
diff changeset
   125
              [] => error "No active testers for quickcheck"
45428
aa35c9454a95 quickcheck invocations in mutabelle must not catch codegenerator errors internally
bulwahn
parents: 45397
diff changeset
   126
            | [tester] => tester ctxt' false [] [(t, [])]
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 44845
diff changeset
   127
            | _ => error "Multiple active testers for quickcheck"
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   128
          in
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   129
            case Quickcheck.counterexample_of result of 
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   130
              NONE => (NoCex, Quickcheck.timings_of result)
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   131
            | SOME _ => (GenuineCex, Quickcheck.timings_of result)
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   132
          end) ()
40931
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 40920
diff changeset
   133
  handle TimeLimit.TimeOut =>
43018
121aa59b4d17 renamed "Auto_Tools" "Try"
blanchet
parents: 42438
diff changeset
   134
         (Timeout, [("timelimit", Real.floor (!Try.auto_time_limit))])
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   135
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   136
fun quickcheck_mtd change_options quickcheck_generator =
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   137
  ("quickcheck_" ^ quickcheck_generator, invoke_quickcheck change_options quickcheck_generator)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   138
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   139
(** solve direct **)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   140
 
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   141
fun invoke_solve_direct thy t =
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   142
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42179
diff changeset
   143
    val state = Proof.theorem NONE (K I) (map (single o rpair []) [t]) (Proof_Context.init_global thy) 
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   144
  in
43030
e1172791ad0d compile
blanchet
parents: 43018
diff changeset
   145
    case Solve_Direct.solve_direct state of
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   146
      (true, _) => (Solved, [])
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   147
    | (false, _) => (Unsolved, [])
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   148
  end
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   149
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   150
val solve_direct_mtd = ("solve_direct", invoke_solve_direct) 
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   151
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   152
(** try **)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   153
43030
e1172791ad0d compile
blanchet
parents: 43018
diff changeset
   154
fun invoke_try_methods thy t =
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   155
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42179
diff changeset
   156
    val state = Proof.theorem NONE (K I) (map (single o rpair []) [t]) (Proof_Context.init_global thy)
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   157
  in
43030
e1172791ad0d compile
blanchet
parents: 43018
diff changeset
   158
    case Try_Methods.try_methods (SOME (seconds 5.0)) ([], [], [], []) state of
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   159
      true => (Solved, [])
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   160
    | false => (Unsolved, [])
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   161
  end
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   162
43030
e1172791ad0d compile
blanchet
parents: 43018
diff changeset
   163
val try_methods_mtd = ("try_methods", invoke_try_methods)
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   164
40971
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   165
(** sledgehammer **)
40974
29e5cae93584 commenting out sledgehammer_mtd in Mutabelle
bulwahn
parents: 40972
diff changeset
   166
(*
40971
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   167
fun invoke_sledgehammer thy t =
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   168
  if can (Goal.prove_global thy (Term.add_free_names t [])  [] t)
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   169
      (fn {context, ...} => Sledgehammer_Tactics.sledgehammer_with_metis_tac context 1) then
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   170
    (Solved, ([], NONE))
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   171
  else
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   172
    (Unsolved, ([], NONE))
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   173
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   174
val sledgehammer_mtd = ("sledgehammer", invoke_sledgehammer)
40974
29e5cae93584 commenting out sledgehammer_mtd in Mutabelle
bulwahn
parents: 40972
diff changeset
   175
*)
45397
20128348e9b9 revived Refute in Mutabelle
blanchet
parents: 45165
diff changeset
   176
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   177
fun invoke_refute thy t =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   178
  let
45397
20128348e9b9 revived Refute in Mutabelle
blanchet
parents: 45165
diff changeset
   179
    val params = []
20128348e9b9 revived Refute in Mutabelle
blanchet
parents: 45165
diff changeset
   180
    val res = Refute.refute_term (Proof_Context.init_global thy) params [] t
40132
7ee65dbffa31 renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents: 39555
diff changeset
   181
    val _ = Output.urgent_message ("Refute: " ^ res)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   182
  in
45397
20128348e9b9 revived Refute in Mutabelle
blanchet
parents: 45165
diff changeset
   183
    (rpair []) (case res of
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   184
      "genuine" => GenuineCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   185
    | "likely_genuine" => GenuineCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   186
    | "potential" => PotentialCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   187
    | "none" => NoCex
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   188
    | "unknown" => Donno
45397
20128348e9b9 revived Refute in Mutabelle
blanchet
parents: 45165
diff changeset
   189
    | _ => Error)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   190
  end
45397
20128348e9b9 revived Refute in Mutabelle
blanchet
parents: 45165
diff changeset
   191
  handle Refute.REFUTE (loc, details) =>
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   192
         (error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   193
                   "."))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   194
val refute_mtd = ("refute", invoke_refute)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   195
41190
0bdc6fac5f48 reactivating nitpick in Mutabelle
bulwahn
parents: 41106
diff changeset
   196
(** nitpick **)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   197
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   198
fun invoke_nitpick thy t =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   199
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42179
diff changeset
   200
    val ctxt = Proof_Context.init_global thy
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   201
    val state = Proof.init ctxt
41190
0bdc6fac5f48 reactivating nitpick in Mutabelle
bulwahn
parents: 41106
diff changeset
   202
    val (res, _) = Nitpick.pick_nits_in_term state
43030
e1172791ad0d compile
blanchet
parents: 43018
diff changeset
   203
      (Nitpick_Isar.default_params thy []) Nitpick.Normal 1 1 1 [] [] t
41190
0bdc6fac5f48 reactivating nitpick in Mutabelle
bulwahn
parents: 41106
diff changeset
   204
    val _ = Output.urgent_message ("Nitpick: " ^ res)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   205
  in
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   206
    (rpair []) (case res of
41190
0bdc6fac5f48 reactivating nitpick in Mutabelle
bulwahn
parents: 41106
diff changeset
   207
      "genuine" => GenuineCex
0bdc6fac5f48 reactivating nitpick in Mutabelle
bulwahn
parents: 41106
diff changeset
   208
    | "likely_genuine" => GenuineCex
0bdc6fac5f48 reactivating nitpick in Mutabelle
bulwahn
parents: 41106
diff changeset
   209
    | "potential" => PotentialCex
0bdc6fac5f48 reactivating nitpick in Mutabelle
bulwahn
parents: 41106
diff changeset
   210
    | "none" => NoCex
0bdc6fac5f48 reactivating nitpick in Mutabelle
bulwahn
parents: 41106
diff changeset
   211
    | "unknown" => Donno
0bdc6fac5f48 reactivating nitpick in Mutabelle
bulwahn
parents: 41106
diff changeset
   212
    | _ => Error)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   213
  end
41190
0bdc6fac5f48 reactivating nitpick in Mutabelle
bulwahn
parents: 41106
diff changeset
   214
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   215
val nitpick_mtd = ("nitpick", invoke_nitpick)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   216
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   217
(* filtering forbidden theorems and mutants *)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   218
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38857
diff changeset
   219
val comms = [@{const_name HOL.eq}, @{const_name HOL.disj}, @{const_name HOL.conj}]
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   220
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   221
val forbidden =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   222
 [(* (@{const_name "power"}, "'a"), *)
35325
4123977b469d adding ROOT.ML to HOL-Mutabelle session; uncommenting HOL.induct constants in Mutabelle session
bulwahn
parents: 35324
diff changeset
   223
  (*(@{const_name induct_equal}, "'a"),
4123977b469d adding ROOT.ML to HOL-Mutabelle session; uncommenting HOL.induct constants in Mutabelle session
bulwahn
parents: 35324
diff changeset
   224
  (@{const_name induct_implies}, "'a"),
4123977b469d adding ROOT.ML to HOL-Mutabelle session; uncommenting HOL.induct constants in Mutabelle session
bulwahn
parents: 35324
diff changeset
   225
  (@{const_name induct_conj}, "'a"),*)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   226
  (@{const_name "undefined"}, "'a"),
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   227
  (@{const_name "default"}, "'a"),
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   228
  (@{const_name "dummy_pattern"}, "'a::{}"),
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   229
  (@{const_name "HOL.simp_implies"}, "prop => prop => prop"),
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   230
  (@{const_name "bot_fun_inst.bot_fun"}, "'a"),
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   231
  (@{const_name "top_fun_inst.top_fun"}, "'a"),
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   232
  (@{const_name "Pure.term"}, "'a"),
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   233
  (@{const_name "top_class.top"}, "'a"),
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   234
  (@{const_name "Quotient.Quot_True"}, "'a")(*,
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   235
  (@{const_name "uminus"}, "'a"),
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   236
  (@{const_name "Nat.size"}, "'a"),
35092
cfe605c54e50 moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents: 34974
diff changeset
   237
  (@{const_name "Groups.abs"}, "'a") *)]
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   238
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   239
val forbidden_thms =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   240
 ["finite_intvl_succ_class",
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   241
  "nibble"]
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   242
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   243
val forbidden_consts =
46433
b67bb064de1e mutabelle ignores theorems with internal constants
bulwahn
parents: 46376
diff changeset
   244
 [@{const_name nibble_pair_of_char}, @{const_name "TYPE"},
b67bb064de1e mutabelle ignores theorems with internal constants
bulwahn
parents: 46376
diff changeset
   245
  @{const_name Datatype.dsum}, @{const_name Datatype.usum}]
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   246
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   247
fun is_forbidden_theorem (s, th) =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   248
  let val consts = Term.add_const_names (prop_of th) [] in
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   249
    exists (member (op =) (space_explode "." s)) forbidden_thms orelse
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   250
    exists (member (op =) forbidden_consts) consts orelse
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   251
    length (space_explode "." s) <> 2 orelse
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   252
    String.isPrefix "type_definition" (List.last (space_explode "." s)) orelse
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   253
    String.isSuffix "_def" s orelse
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   254
    String.isSuffix "_raw" s orelse
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   255
    String.isPrefix "term_of" (List.last (space_explode "." s))
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   256
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   257
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   258
val forbidden_mutant_constnames =
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   259
 ["HOL.induct_equal",
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   260
  "HOL.induct_implies",
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   261
  "HOL.induct_conj",
46314
f6532c597300 adding some more forbidden constant names for the mutated conjecture generation
bulwahn
parents: 46310
diff changeset
   262
  "HOL.induct_forall",
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   263
 @{const_name undefined},
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   264
 @{const_name default},
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   265
 @{const_name dummy_pattern},
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   266
 @{const_name "HOL.simp_implies"},
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   267
 @{const_name "bot_fun_inst.bot_fun"},
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   268
 @{const_name "top_fun_inst.top_fun"},
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   269
 @{const_name "Pure.term"},
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   270
 @{const_name "top_class.top"},
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   271
 (*@{const_name "HOL.equal"},*)
40971
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   272
 @{const_name "Quotient.Quot_True"},
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   273
 @{const_name "equal_fun_inst.equal_fun"},
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   274
 @{const_name "equal_bool_inst.equal_bool"},
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   275
 @{const_name "ord_fun_inst.less_eq_fun"},
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   276
 @{const_name "ord_fun_inst.less_fun"},
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   277
 @{const_name Meson.skolem},
43111
61faa204c810 compile
blanchet
parents: 43030
diff changeset
   278
 @{const_name ATP.fequal},
46326
9a5d8e7684e5 some more constants on mutabelle's blacklist
bulwahn
parents: 46315
diff changeset
   279
 @{const_name ATP.fEx},
42435
c3d880b13aa9 adding two further code-generator internal constants to the blacklist of Mutabelle
bulwahn
parents: 42361
diff changeset
   280
 @{const_name transfer_morphism},
c3d880b13aa9 adding two further code-generator internal constants to the blacklist of Mutabelle
bulwahn
parents: 42361
diff changeset
   281
 @{const_name enum_prod_inst.enum_all_prod},
46314
f6532c597300 adding some more forbidden constant names for the mutated conjecture generation
bulwahn
parents: 46310
diff changeset
   282
 @{const_name enum_prod_inst.enum_ex_prod},
46315
40522961d4b1 adding another internal constant to mutabelle's blacklust
bulwahn
parents: 46314
diff changeset
   283
 @{const_name Quickcheck.catch_match},
40522961d4b1 adding another internal constant to mutabelle's blacklust
bulwahn
parents: 46314
diff changeset
   284
 @{const_name Quickcheck_Exhaustive.unknown}
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   285
 (*@{const_name "==>"}, @{const_name "=="}*)]
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   286
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   287
val forbidden_mutant_consts =
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   288
  [
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   289
   (@{const_name "Groups.zero_class.zero"}, @{typ "prop => prop => prop"}),
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   290
   (@{const_name "Groups.one_class.one"}, @{typ "prop => prop => prop"}),
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   291
   (@{const_name "Groups.plus_class.plus"}, @{typ "prop => prop => prop"}),
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   292
   (@{const_name "Groups.minus_class.minus"}, @{typ "prop => prop => prop"}),
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   293
   (@{const_name "Groups.times_class.times"}, @{typ "prop => prop => prop"}),
44064
5bce8ff0d9ae moved division ring stuff from Rings.thy to Fields.thy
huffman
parents: 43883
diff changeset
   294
   (@{const_name "Fields.inverse_class.divide"}, @{typ "prop => prop => prop"}),
44845
5e51075cbd97 added syntactic classes for "inf" and "sup"
krauss
parents: 44064
diff changeset
   295
   (@{const_name "Lattices.inf_class.inf"}, @{typ "prop => prop => prop"}),
5e51075cbd97 added syntactic classes for "inf" and "sup"
krauss
parents: 44064
diff changeset
   296
   (@{const_name "Lattices.sup_class.sup"}, @{typ "prop => prop => prop"}),
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   297
   (@{const_name "Orderings.bot_class.bot"}, @{typ "prop => prop => prop"}),
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   298
   (@{const_name "Orderings.ord_class.min"}, @{typ "prop => prop => prop"}),
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   299
   (@{const_name "Orderings.ord_class.max"}, @{typ "prop => prop => prop"}),
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   300
   (@{const_name "Divides.div_class.mod"}, @{typ "prop => prop => prop"}),
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   301
   (@{const_name "Divides.div_class.div"}, @{typ "prop => prop => prop"}),
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   302
   (@{const_name "GCD.gcd_class.gcd"}, @{typ "prop => prop => prop"}),
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   303
   (@{const_name "GCD.gcd_class.lcm"}, @{typ "prop => prop => prop"}),
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   304
   (@{const_name "Orderings.bot_class.bot"}, @{typ "bool => prop"}),
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   305
   (@{const_name "Groups.one_class.one"}, @{typ "bool => prop"}),
46326
9a5d8e7684e5 some more constants on mutabelle's blacklist
bulwahn
parents: 46315
diff changeset
   306
   (@{const_name "Groups.zero_class.zero"},@{typ "bool => prop"}),
9a5d8e7684e5 some more constants on mutabelle's blacklist
bulwahn
parents: 46315
diff changeset
   307
   (@{const_name "equal_class.equal"},@{typ "bool => bool => bool"})]
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   308
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   309
fun is_forbidden_mutant t =
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   310
  let
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   311
    val const_names = Term.add_const_names t []
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   312
    val consts = Term.add_consts t []
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   313
  in
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   314
    exists (String.isPrefix "Nitpick") const_names orelse
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   315
    exists (String.isSubstring "_sumC") const_names orelse
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   316
    exists (member (op =) forbidden_mutant_constnames) const_names orelse
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   317
    exists (member (op =) forbidden_mutant_consts) consts
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   318
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   319
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   320
(* executable via quickcheck *)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   321
40248
2107581b404d adapting HOL-Mutabelle to changes in quickcheck
bulwahn
parents: 40132
diff changeset
   322
fun is_executable_term thy t =
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   323
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42179
diff changeset
   324
    val ctxt = Proof_Context.init_global thy
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   325
  in
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   326
    can (TimeLimit.timeLimit (seconds 2.0)
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 44845
diff changeset
   327
      (Quickcheck.test_terms
45165
f4896c792316 adding testing of quickcheck narrowing with finite types to mutabelle script; modified is_executable in mutabelle_extra
bulwahn
parents: 45159
diff changeset
   328
        ((Context.proof_map (Quickcheck.set_active_testers ["exhaustive"]) #>
f4896c792316 adding testing of quickcheck narrowing with finite types to mutabelle script; modified is_executable in mutabelle_extra
bulwahn
parents: 45159
diff changeset
   329
          Config.put Quickcheck.finite_types true #>
41106
09037a02f5ec setting finite_type_size to 1 in mutabelle_extra
bulwahn
parents: 40974
diff changeset
   330
          Config.put Quickcheck.finite_type_size 1 #>
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   331
          Config.put Quickcheck.size 1 #> Config.put Quickcheck.iterations 1) ctxt)
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 44845
diff changeset
   332
        (false, false) [])) (map (rpair [] o Object_Logic.atomize_term thy)
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 44845
diff changeset
   333
        (fst (Variable.import_terms true [t] ctxt)))
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   334
  end
40248
2107581b404d adapting HOL-Mutabelle to changes in quickcheck
bulwahn
parents: 40132
diff changeset
   335
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   336
fun is_executable_thm thy th = is_executable_term thy (prop_of th)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   337
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   338
val freezeT =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   339
  map_types (map_type_tvar (fn ((a, i), S) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   340
    TFree (if i = 0 then a else a ^ "_" ^ string_of_int i, S)))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   341
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   342
fun thms_of all thy =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   343
  filter
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   344
    (fn th => (all orelse Context.theory_name (theory_of_thm th) = Context.theory_name thy)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   345
      (* andalso is_executable_thm thy th *))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   346
    (map snd (filter_out is_forbidden_theorem (Mutabelle.all_unconcealed_thms_of thy)))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   347
41409
0bc364f772ef made SML/NJ happy;
wenzelm
parents: 41408
diff changeset
   348
fun count x = (length oo filter o equal) x
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   349
42014
75417ef605ba simplified various cpu_time clones (!): eliminated odd Exn.capture/Exn.release (no need to "stop" timing);
wenzelm
parents: 42012
diff changeset
   350
fun cpu_time description e =
75417ef605ba simplified various cpu_time clones (!): eliminated odd Exn.capture/Exn.release (no need to "stop" timing);
wenzelm
parents: 42012
diff changeset
   351
  let val ({cpu, ...}, result) = Timing.timing e ()
75417ef605ba simplified various cpu_time clones (!): eliminated odd Exn.capture/Exn.release (no need to "stop" timing);
wenzelm
parents: 42012
diff changeset
   352
  in (result, (description, Time.toMilliseconds cpu)) end
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   353
(*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   354
fun unsafe_invoke_mtd thy (mtd_name, invoke_mtd) t =
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   355
  let
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   356
    val _ = Output.urgent_message ("Invoking " ^ mtd_name)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   357
    val ((res, (timing, reports)), time) = cpu_time "total time" (fn () => invoke_mtd thy t
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   358
      handle ERROR s => (tracing s; (Error, ([], NONE))))
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   359
    val _ = Output.urgent_message (" Done")
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   360
  in (res, (time :: timing, reports)) end
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   361
*)  
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   362
fun safe_invoke_mtd thy (mtd_name, invoke_mtd) t =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   363
  let
40132
7ee65dbffa31 renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents: 39555
diff changeset
   364
    val _ = Output.urgent_message ("Invoking " ^ mtd_name)
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   365
    val (res, timing) = (*cpu_time "total time"
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   366
      (fn () => *)case try (invoke_mtd thy) t of
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   367
          SOME (res, timing) => (res, timing)
40132
7ee65dbffa31 renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents: 39555
diff changeset
   368
        | NONE => (Output.urgent_message ("**** PROBLEMS WITH " ^ Syntax.string_of_term_global thy t);
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   369
           (Error, []))
40132
7ee65dbffa31 renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents: 39555
diff changeset
   370
    val _ = Output.urgent_message (" Done")
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   371
  in (res, timing) end
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   372
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   373
(* theory -> term list -> mtd -> subentry *)
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   374
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   375
fun test_mutants_using_one_method thy mutants (mtd_name, invoke_mtd) =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   376
  let
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   377
     val res = map (fst o safe_invoke_mtd thy (mtd_name, invoke_mtd)) mutants
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   378
  in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   379
    (mtd_name, count GenuineCex res, count PotentialCex res, count NoCex res,
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   380
     count Donno res, count Timeout res, count Error res)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   381
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   382
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   383
(* creating entries *)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   384
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   385
fun create_entry thy thm exec mutants mtds =
36743
ce2297415b54 prefer Thm.get_name_hint, which is closer to a user-space idea of "theorem name";
wenzelm
parents: 36692
diff changeset
   386
  (Thm.get_name_hint thm, exec, map (test_mutants_using_one_method thy mutants) mtds)
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   387
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   388
fun create_detailed_entry thy thm exec mutants mtds =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   389
  let
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   390
    fun create_mutant_subentry mutant = (mutant,
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   391
      map (fn (mtd_name, invoke_mtd) =>
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   392
        (mtd_name, safe_invoke_mtd thy (mtd_name, invoke_mtd) mutant)) mtds)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   393
  in
36743
ce2297415b54 prefer Thm.get_name_hint, which is closer to a user-space idea of "theorem name";
wenzelm
parents: 36692
diff changeset
   394
    (Thm.get_name_hint thm, exec, prop_of thm, map create_mutant_subentry mutants)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   395
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   396
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   397
(* (theory -> thm -> bool -> term list -> mtd list -> 'a) -> theory -> mtd list -> thm -> 'a *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   398
fun mutate_theorem create_entry thy mtds thm =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   399
  let
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   400
    val exec = is_executable_thm thy thm
43277
1fd31f859fc7 pervasive Output operations;
wenzelm
parents: 43244
diff changeset
   401
    val _ = tracing (if exec then "EXEC" else "NOEXEC")
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   402
    val mutants =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   403
          (if num_mutations = 0 then
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   404
             [Thm.prop_of thm]
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   405
           else
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   406
             Mutabelle.mutate_mix (Thm.prop_of thm) thy comms forbidden
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   407
                                  num_mutations)
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   408
             |> tap (fn muts => tracing ("mutants: " ^ string_of_int (length muts)))
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   409
             |> filter_out is_forbidden_mutant
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   410
    val mutants =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   411
      if exec then
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   412
        let
40132
7ee65dbffa31 renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents: 39555
diff changeset
   413
          val _ = Output.urgent_message ("BEFORE PARTITION OF " ^
41491
a2ad5b824051 eliminated Int.toString;
wenzelm
parents: 41409
diff changeset
   414
                            string_of_int (length mutants) ^ " MUTANTS")
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   415
          val (execs, noexecs) = List.partition (is_executable_term thy) (take_random (20 * max_mutants) mutants)
41491
a2ad5b824051 eliminated Int.toString;
wenzelm
parents: 41409
diff changeset
   416
          val _ = tracing ("AFTER PARTITION (" ^ string_of_int (length execs) ^
a2ad5b824051 eliminated Int.toString;
wenzelm
parents: 41409
diff changeset
   417
                           " vs " ^ string_of_int (length noexecs) ^ ")")
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   418
        in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   419
          execs @ take_random (Int.max (0, max_mutants - length execs)) noexecs
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   420
        end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   421
      else
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   422
        mutants
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   423
    val mutants = mutants
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   424
          |> map Mutabelle.freeze |> map freezeT
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   425
(*          |> filter (not o is_forbidden_mutant) *)
42429
7691cc61720a standardized some ML aliases;
wenzelm
parents: 42361
diff changeset
   426
          |> map_filter (try (Sign.cert_term thy))
7691cc61720a standardized some ML aliases;
wenzelm
parents: 42361
diff changeset
   427
          |> filter (is_some o try (Thm.cterm_of thy))
7691cc61720a standardized some ML aliases;
wenzelm
parents: 42361
diff changeset
   428
          |> filter (is_some o try (Syntax.check_term (Proof_Context.init_global thy)))
40971
6604115019bf adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents: 40932
diff changeset
   429
          |> take_random max_mutants
40132
7ee65dbffa31 renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents: 39555
diff changeset
   430
    val _ = map (fn t => Output.urgent_message ("MUTANT: " ^ Syntax.string_of_term_global thy t)) mutants
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   431
  in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   432
    create_entry thy thm exec mutants mtds
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   433
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   434
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   435
(* theory -> mtd list -> thm list -> report *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   436
val mutate_theorems = map ooo mutate_theorem
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   437
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35092
diff changeset
   438
fun string_of_mutant_subentry thy thm_name (t, results) =
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   439
  "mutant: " ^ Syntax.string_of_term_global thy t ^ "\n" ^
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35092
diff changeset
   440
  space_implode "; "
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35092
diff changeset
   441
    (map (fn (mtd_name, (outcome, timing)) => mtd_name ^ ": " ^ string_of_outcome outcome) results) ^
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   442
  "\n"
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   443
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   444
(* string -> string *)
39555
ccb223a4d49c added XML.content_of convenience -- cover XML.body, which is the general situation;
wenzelm
parents: 39324
diff changeset
   445
val unyxml = XML.content_of o YXML.parse_body
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   446
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35092
diff changeset
   447
fun string_of_mutant_subentry' thy thm_name (t, results) =
35380
6ac5b81a763d adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents: 35325
diff changeset
   448
  let
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   449
   (* fun string_of_report (Quickcheck.Report {iterations = i, raised_match_errors = e,
35380
6ac5b81a763d adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents: 35325
diff changeset
   450
      satisfied_assms = s, positive_concl_tests = p}) =
6ac5b81a763d adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents: 35325
diff changeset
   451
      "errors: " ^ string_of_int e ^ "; conclusion tests: " ^ string_of_int p
6ac5b81a763d adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents: 35325
diff changeset
   452
    fun string_of_reports NONE = ""
6ac5b81a763d adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents: 35325
diff changeset
   453
      | string_of_reports (SOME reports) =
6ac5b81a763d adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents: 35325
diff changeset
   454
        cat_lines (map (fn (size, [report]) =>
42089
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   455
          "size " ^ string_of_int size ^ ": " ^ string_of_report report) (rev reports))*)
904897d0c5bd adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents: 42032
diff changeset
   456
    fun string_of_mtd_result (mtd_name, (outcome, timing)) =
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   457
      mtd_name ^ ": " ^ string_of_outcome outcome
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   458
      (*" with time " ^ " (" ^ space_implode "; " (map (fn (s, t) => (s ^ ": " ^ string_of_int t)) timing) ^ ")"*)
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   459
      (*^ "\n" ^ string_of_reports reports*)
35380
6ac5b81a763d adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents: 35325
diff changeset
   460
  in
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   461
    "mutant of " ^ thm_name ^ ":\n"
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   462
    ^ unyxml (Syntax.string_of_term_global thy t) ^ "\n" ^ space_implode "; " (map string_of_mtd_result results)
35380
6ac5b81a763d adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents: 35325
diff changeset
   463
  end
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35092
diff changeset
   464
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   465
fun string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) = 
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   466
   thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ": " ^
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   467
   Syntax.string_of_term_global thy t ^ "\n" ^                                    
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35092
diff changeset
   468
   cat_lines (map (string_of_mutant_subentry' thy thm_name) mutant_subentries) ^ "\n"
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   469
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   470
fun theoryfile_string_of_mutant_subentry thy thm_name (i, (t, results)) =
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   471
  "lemma " ^ thm_name ^ "_" ^ string_of_int (i + 1) ^ ":\n" ^
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   472
  "\"" ^ unyxml (Syntax.string_of_term_global thy t) ^
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   473
  "\" \nquickcheck\noops\n"
36255
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   474
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   475
fun theoryfile_string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) =
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   476
  "subsubsection {* mutants of " ^ thm_name ^ " *}\n\n" ^
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   477
  cat_lines (map_index
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   478
    (theoryfile_string_of_mutant_subentry thy thm_name) mutant_subentries) ^ "\n"
f8b3381e1437 tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents: 35625
diff changeset
   479
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   480
(* subentry -> string *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   481
fun string_for_subentry (mtd_name, genuine_cex, potential_cex, no_cex, donno,
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   482
                         timeout, error) =
41491
a2ad5b824051 eliminated Int.toString;
wenzelm
parents: 41409
diff changeset
   483
  "    " ^ mtd_name ^ ": " ^ string_of_int genuine_cex ^ "+ " ^
a2ad5b824051 eliminated Int.toString;
wenzelm
parents: 41409
diff changeset
   484
  string_of_int potential_cex ^ "= " ^ string_of_int no_cex ^ "- " ^
a2ad5b824051 eliminated Int.toString;
wenzelm
parents: 41409
diff changeset
   485
  string_of_int donno ^ "? " ^ string_of_int timeout ^ "T " ^
a2ad5b824051 eliminated Int.toString;
wenzelm
parents: 41409
diff changeset
   486
  string_of_int error ^ "!"
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   487
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   488
(* entry -> string *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   489
fun string_for_entry (thm_name, exec, subentries) =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   490
  thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ":\n" ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   491
  cat_lines (map string_for_subentry subentries) ^ "\n"
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   492
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   493
(* report -> string *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   494
fun string_for_report report = cat_lines (map string_for_entry report)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   495
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   496
(* string -> report -> unit *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   497
fun write_report file_name =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   498
  File.write (Path.explode file_name) o string_for_report
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   499
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   500
(* theory -> mtd list -> thm list -> string -> unit *)
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   501
fun mutate_theorems_and_write_report thy mtds thms file_name =
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   502
  let
40132
7ee65dbffa31 renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents: 39555
diff changeset
   503
    val _ = Output.urgent_message "Starting Mutabelle..."
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42179
diff changeset
   504
    val ctxt = Proof_Context.init_global thy
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   505
    val path = Path.explode file_name
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   506
    (* for normal report: *)
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   507
    (*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   508
    val (gen_create_entry, gen_string_for_entry) = (create_entry, string_for_entry)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   509
    *)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   510
    (* for detailled report: *)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   511
    val (gen_create_entry, gen_string_for_entry) = (create_detailed_entry, string_of_detailed_entry thy)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   512
    (* for theory creation: *)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   513
    (*val (gen_create_entry, gen_string_for_entry) = (create_detailed_entry, theoryfile_string_of_detailed_entry thy)*)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   514
  in
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   515
    File.write path (
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   516
    "Mutation options = "  ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   517
      "max_mutants: " ^ string_of_int max_mutants ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   518
      "; num_mutations: " ^ string_of_int num_mutations ^ "\n" ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   519
    "QC options = " ^
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   520
      (*"quickcheck_generator: " ^ quickcheck_generator ^ ";*)
40653
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40381
diff changeset
   521
      "size: " ^ string_of_int (Config.get ctxt Quickcheck.size) ^
43244
db041e88a805 printing environment in mutabelle's log
bulwahn
parents: 43111
diff changeset
   522
      "; iterations: " ^ string_of_int (Config.get ctxt Quickcheck.iterations) ^ "\n" ^
db041e88a805 printing environment in mutabelle's log
bulwahn
parents: 43111
diff changeset
   523
    "Isabelle environment = ISABELLE_GHC: " ^ getenv "ISABELLE_GHC" ^ "\n");
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   524
    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
   525
    ()
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   526
  end
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   527
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   528
end;