src/HOL/Mutabelle/MutabelleExtra.thy
author wenzelm
Thu, 18 Apr 2013 17:07:01 +0200
changeset 51717 9e7d1c139569
parent 49992 1e68f4701906
child 52639 df830310e550
permissions -rw-r--r--
simplifier uses proper Proof.context instead of historic type simpset;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
     1
theory MutabelleExtra
49992
1e68f4701906 repaired "Mutabelle" after Refute move
blanchet
parents: 48891
diff changeset
     2
imports Complex_Main "~~/src/HOL/Library/Refute"
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: 40341
diff changeset
     3
(*  "~/repos/afp/thys/AVL-Trees/AVL"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
     4
  "~/repos/afp/thys/BinarySearchTree/BinaryTree"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
     5
  "~/repos/afp/thys/Huffman/Huffman"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
     6
  "~/repos/afp/thys/List-Index/List_Index"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
     7
  "~/repos/afp/thys/Matrix/Matrix"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
     8
  "~/repos/afp/thys/NormByEval/NBE"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
     9
  "~/repos/afp/thys/Polynomials/Polynomial"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    10
  "~/repos/afp/thys/Presburger-Automata/Presburger_Automata"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    11
  "~/repos/afp/thys/Abstract-Rewriting/Abstract_Rewriting"*)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    12
begin
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    13
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 46454
diff changeset
    14
ML_file "mutabelle.ML"
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 46454
diff changeset
    15
ML_file "mutabelle_extra.ML"
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 46454
diff changeset
    16
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: 40341
diff changeset
    17
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    18
section {* configuration *}
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    19
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: 40341
diff changeset
    20
ML {* val log_directory = "" *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    21
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    22
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    23
quickcheck_params [quiet, finite_types = false, report = false, size = 5, iterations = 1000]
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    24
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    25
(*
40341
03156257040f standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents: 40133
diff changeset
    26
nitpick_params [timeout = 5, sat_solver = MiniSat, no_overlord, verbose, card = 1-5, iter = 1,2,4,8,12]
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    27
refute_params [maxtime = 10, minsize = 1, maxsize = 5, satsolver = jerusat]
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    28
*)
43018
121aa59b4d17 renamed "Auto_Tools" "Try"
blanchet
parents: 40931
diff changeset
    29
ML {* Try.auto_time_limit := 10.0 *}
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    30
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: 40341
diff changeset
    31
ML {* val mtds = [
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    32
  MutabelleExtra.quickcheck_mtd (Config.put Quickcheck.finite_types false) "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: 40341
diff changeset
    33
  MutabelleExtra.quickcheck_mtd (Config.put Quickcheck.finite_types true) "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: 40341
diff changeset
    34
  MutabelleExtra.quickcheck_mtd (Config.put Quickcheck.finite_types false) "small",
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    35
  MutabelleExtra.quickcheck_mtd (Config.put Quickcheck.finite_types true) "small"
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    36
]
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    37
*}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    38
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    39
ML {*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    40
fun mutation_testing_of (name, 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: 40341
diff changeset
    41
  (MutabelleExtra.random_seed := 1.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: 40341
diff changeset
    42
  MutabelleExtra.thms_of false 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: 40341
diff changeset
    43
  |> MutabelleExtra.take_random 200
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    44
  |> (fn thms => MutabelleExtra.mutate_theorems_and_write_report
46454
d72ab6bf6e6d making num_mutations a configuration that can be changed with the mutabelle bash command
bulwahn
parents: 46453
diff changeset
    45
         @{theory} (1, 50) mtds thms (log_directory ^ "/" ^ name)))
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: 40341
diff changeset
    46
*}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    47
  
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    48
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    49
text {* Uncomment the following ML code to check the counterexample generation with all theorems of Complex_Main. *}
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    50
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: 40341
diff changeset
    51
(*
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    52
ML {*
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: 40341
diff changeset
    53
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    54
      MutabelleExtra.random_seed := 1.0;
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    55
      MutabelleExtra.thms_of true @{theory Complex_Main}
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    56
      |> MutabelleExtra.take_random 1000000
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    57
      |> (fn thms => List.drop (thms, 1000))
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    58
      |> (fn thms => MutabelleExtra.mutate_theorems_and_write_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: 40341
diff changeset
    59
             @{theory} [
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    60
                        MutabelleExtra.quickcheck_mtd "code",
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    61
                        MutabelleExtra.smallcheck_mtd
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    62
                        (*MutabelleExtra.refute_mtd,
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    63
                        MutabelleExtra.nitpick_mtd*)] thms "/tmp/muta.out")
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: 40341
diff changeset
    64
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    65
 *}
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    66
*)
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: 40341
diff changeset
    67
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    68
section {* Mutation testing Isabelle theories *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    69
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    70
subsection {* List theory *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    71
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    72
(*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    73
ML {*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    74
mutation_testing_of ("List", @{theory List})
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
    75
 *}
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: 40341
diff changeset
    76
*)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    77
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    78
section {* Mutation testing AFP theories *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    79
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    80
subsection {* AVL-Trees *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    81
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    82
(*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    83
ML {*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    84
mutation_testing_of ("AVL-Trees", @{theory AVL})
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    85
 *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    86
*)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    87
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    88
subsection {* BinaryTree *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    89
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
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: 40341
diff changeset
    91
ML {*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    92
mutation_testing_of ("BinaryTree", @{theory BinaryTree})
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    93
 *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
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: 40341
diff changeset
    95
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    96
subsection {* Huffman *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    97
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    98
(*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
    99
ML {*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   100
mutation_testing_of ("Huffman", @{theory Huffman})
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   101
 *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   102
*)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   103
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: 40341
diff changeset
   104
subsection {* List_Index *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   105
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   106
(*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   107
ML {*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   108
mutation_testing_of ("List_Index", @{theory List_Index})
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   109
 *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   110
*)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   111
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   112
subsection {* Matrix *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   113
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   114
(*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   115
ML {*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   116
mutation_testing_of ("Matrix", @{theory Matrix})
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   117
 *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   118
*)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   119
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   120
subsection {* NBE *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   121
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   122
(*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   123
ML {*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   124
mutation_testing_of ("NBE", @{theory NBE})
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   125
 *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   126
*)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   127
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   128
subsection {* Polynomial *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   129
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   130
(*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   131
ML {*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   132
mutation_testing_of ("Polynomial", @{theory Polynomial})
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   133
 *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   134
*)
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   135
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   136
subsection {* Presburger Automata *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   137
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
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: 40341
diff changeset
   139
ML {*
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   140
mutation_testing_of ("Presburger_Automata", @{theory Presburger_Automata})
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   141
 *}
d921c97bdbd8 adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents: 40341
diff changeset
   142
*)
34965
3b4762c1052c adding Mutabelle to repository
bulwahn
parents:
diff changeset
   143
39324
05452dd66b2b finished renaming "Auto_Counterexample" to "Auto_Tools"
blanchet
parents: 34965
diff changeset
   144
end