| author | wenzelm | 
| Fri, 17 Sep 2010 15:51:11 +0200 | |
| changeset 39439 | 1c294d150ded | 
| parent 39359 | 6f49c7fbb1b1 | 
| child 40341 | 03156257040f | 
| permissions | -rw-r--r-- | 
| 33982 | 1 | (* Title: HOL/Tools/Nitpick/nitpick_isar.ML | 
| 33192 | 2 | Author: Jasmin Blanchette, TU Muenchen | 
| 34982 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34936diff
changeset | 3 | Copyright 2008, 2009, 2010 | 
| 33192 | 4 | |
| 5 | Adds the "nitpick" and "nitpick_params" commands to Isabelle/Isar's outer | |
| 6 | syntax. | |
| 7 | *) | |
| 8 | ||
| 9 | signature NITPICK_ISAR = | |
| 10 | sig | |
| 11 | type params = Nitpick.params | |
| 12 | ||
| 39316 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 blanchet parents: 
39155diff
changeset | 13 | val auto : bool Unsynchronized.ref | 
| 33192 | 14 | val default_params : theory -> (string * string) list -> params | 
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 15 | val setup : theory -> theory | 
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35665diff
changeset | 16 | end; | 
| 33192 | 17 | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33220diff
changeset | 18 | structure Nitpick_Isar : NITPICK_ISAR = | 
| 33192 | 19 | struct | 
| 20 | ||
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33220diff
changeset | 21 | open Nitpick_Util | 
| 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33220diff
changeset | 22 | open Nitpick_HOL | 
| 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33220diff
changeset | 23 | open Nitpick_Rep | 
| 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33220diff
changeset | 24 | open Nitpick_Nut | 
| 33192 | 25 | open Nitpick | 
| 26 | ||
| 39316 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 blanchet parents: 
39155diff
changeset | 27 | val auto = Unsynchronized.ref false | 
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 28 | |
| 39344 | 29 | (* Maximum number of scopes for Auto Nitpick. Be frugal since it has to share | 
| 30 | its time slot with several other automatic tools. *) | |
| 31 | val max_auto_scopes = 6 | |
| 32 | ||
| 33601 
4608243edcfc
plain add_preference, no setmp_CRITICAL required;
 wenzelm parents: 
33599diff
changeset | 33 | val _ = | 
| 
4608243edcfc
plain add_preference, no setmp_CRITICAL required;
 wenzelm parents: 
33599diff
changeset | 34 | ProofGeneralPgip.add_preference Preferences.category_tracing | 
| 34982 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34936diff
changeset | 35 | (Preferences.bool_pref auto "auto-nitpick" | 
| 39328 | 36 | "Run Nitpick automatically.") | 
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 37 | |
| 33192 | 38 | type raw_param = string * string list | 
| 39 | ||
| 40 | val default_default_params = | |
| 38180 
7a88032f9265
bump up the max cardinalities, to use up more of the time given to us by the user
 blanchet parents: 
38127diff
changeset | 41 |   [("card", "1\<midarrow>10"),
 | 
| 
7a88032f9265
bump up the max cardinalities, to use up more of the time given to us by the user
 blanchet parents: 
38127diff
changeset | 42 |    ("iter", "0,1,2,4,8,12,16,20,24,28"),
 | 
| 
7a88032f9265
bump up the max cardinalities, to use up more of the time given to us by the user
 blanchet parents: 
38127diff
changeset | 43 |    ("bits", "1,2,3,4,6,8,10,12,14,16"),
 | 
| 
7a88032f9265
bump up the max cardinalities, to use up more of the time given to us by the user
 blanchet parents: 
38127diff
changeset | 44 |    ("bisim_depth", "9"),
 | 
| 35964 | 45 |    ("box", "smart"),
 | 
| 46 |    ("finitize", "smart"),
 | |
| 47 |    ("mono", "smart"),
 | |
| 48 |    ("std", "true"),
 | |
| 49 |    ("wf", "smart"),
 | |
| 50 |    ("sat_solver", "smart"),
 | |
| 51 |    ("batch_size", "smart"),
 | |
| 52 |    ("blocking", "true"),
 | |
| 53 |    ("falsify", "true"),
 | |
| 54 |    ("user_axioms", "smart"),
 | |
| 55 |    ("assms", "true"),
 | |
| 56 |    ("merge_type_vars", "false"),
 | |
| 57 |    ("binary_ints", "smart"),
 | |
| 58 |    ("destroy_constrs", "true"),
 | |
| 59 |    ("specialize", "true"),
 | |
| 60 |    ("star_linear_preds", "true"),
 | |
| 61 |    ("peephole_optim", "true"),
 | |
| 38127 | 62 |    ("datatype_sym_break", "5"),
 | 
| 63 |    ("kodkod_sym_break", "15"),
 | |
| 35964 | 64 |    ("timeout", "30 s"),
 | 
| 65 |    ("tac_timeout", "500 ms"),
 | |
| 66 |    ("max_threads", "0"),
 | |
| 67 |    ("debug", "false"),
 | |
| 68 |    ("verbose", "false"),
 | |
| 69 |    ("overlord", "false"),
 | |
| 70 |    ("show_datatypes", "false"),
 | |
| 71 |    ("show_consts", "false"),
 | |
| 72 |    ("format", "1"),
 | |
| 73 |    ("max_potential", "1"),
 | |
| 74 |    ("max_genuine", "1"),
 | |
| 75 |    ("check_potential", "false"),
 | |
| 76 |    ("check_genuine", "false")]
 | |
| 33192 | 77 | |
| 78 | val negated_params = | |
| 79 |   [("dont_box", "box"),
 | |
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35280diff
changeset | 80 |    ("dont_finitize", "finitize"),
 | 
| 33192 | 81 |    ("non_mono", "mono"),
 | 
| 34982 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34936diff
changeset | 82 |    ("non_std", "std"),
 | 
| 33192 | 83 |    ("non_wf", "wf"),
 | 
| 84 |    ("non_blocking", "blocking"),
 | |
| 85 |    ("satisfy", "falsify"),
 | |
| 86 |    ("no_user_axioms", "user_axioms"),
 | |
| 87 |    ("no_assms", "assms"),
 | |
| 33556 
cba22e2999d5
renamed Nitpick option "coalesce_type_vars" to "merge_type_vars" (shorter) and cleaned up old hacks that are no longer necessary
 blanchet parents: 
33232diff
changeset | 88 |    ("dont_merge_type_vars", "merge_type_vars"),
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 89 |    ("unary_ints", "binary_ints"),
 | 
| 33192 | 90 |    ("dont_destroy_constrs", "destroy_constrs"),
 | 
| 91 |    ("dont_specialize", "specialize"),
 | |
| 92 |    ("dont_star_linear_preds", "star_linear_preds"),
 | |
| 93 |    ("no_peephole_optim", "peephole_optim"),
 | |
| 35964 | 94 |    ("no_debug", "debug"),
 | 
| 33192 | 95 |    ("quiet", "verbose"),
 | 
| 96 |    ("no_overlord", "overlord"),
 | |
| 97 |    ("hide_datatypes", "show_datatypes"),
 | |
| 98 |    ("hide_consts", "show_consts"),
 | |
| 99 |    ("trust_potential", "check_potential"),
 | |
| 100 |    ("trust_genuine", "check_genuine")]
 | |
| 101 | ||
| 102 | fun is_known_raw_param s = | |
| 34936 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34124diff
changeset | 103 | AList.defined (op =) default_default_params s orelse | 
| 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34124diff
changeset | 104 | AList.defined (op =) negated_params s orelse | 
| 39346 
d837998f1e60
remove "atoms" from the list of options with default values
 blanchet parents: 
39344diff
changeset | 105 | member (op =) ["max", "show_all", "whack", "atoms", "eval", "expect"] s orelse | 
| 34936 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34124diff
changeset | 106 | exists (fn p => String.isPrefix (p ^ " ") s) | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35280diff
changeset | 107 | ["card", "max", "iter", "box", "dont_box", "finitize", "dont_finitize", | 
| 37260 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 108 | "mono", "non_mono", "std", "non_std", "wf", "non_wf", "format", | 
| 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 109 | "atoms"] | 
| 33192 | 110 | |
| 111 | fun check_raw_param (s, _) = | |
| 112 | if is_known_raw_param s then () | |
| 39316 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 blanchet parents: 
39155diff
changeset | 113 |   else error ("Unknown parameter: " ^ quote s ^ ".")
 | 
| 33192 | 114 | |
| 115 | fun unnegate_param_name name = | |
| 116 | case AList.lookup (op =) negated_params name of | |
| 117 | NONE => if String.isPrefix "dont_" name then SOME (unprefix "dont_" name) | |
| 118 | else if String.isPrefix "non_" name then SOME (unprefix "non_" name) | |
| 119 | else NONE | |
| 120 | | some_name => some_name | |
| 37169 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 121 | fun normalize_raw_param (name, value) = | 
| 33192 | 122 | case unnegate_param_name name of | 
| 37169 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 123 | SOME name' => [(name', case value of | 
| 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 124 | ["false"] => ["true"] | 
| 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 125 | | ["true"] => ["false"] | 
| 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 126 | | [] => ["false"] | 
| 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 127 | | _ => value)] | 
| 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 128 | | NONE => if name = "show_all" then | 
| 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 129 |               [("show_datatypes", value), ("show_consts", value)]
 | 
| 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 130 | else | 
| 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 131 | [(name, value)] | 
| 33192 | 132 | |
| 33583 
b5e0909cd5ea
merged, and renamed local "TheoryData" to "Data" (following common Isabelle conventions)
 blanchetdiff
changeset | 133 | structure Data = Theory_Data( | 
| 35964 | 134 | type T = raw_param list | 
| 36391 | 135 | val empty = map (apsnd single) default_default_params | 
| 33192 | 136 | val extend = I | 
| 36607 
e5f7235f39c5
made sml/nj happy about Sledgehammer and Nitpick (cf. 6f11c9b1fb3e, 3c2438efe224)
 krauss parents: 
36391diff
changeset | 137 | fun merge (x, y) = AList.merge (op =) (K true) (x, y)) | 
| 33192 | 138 | |
| 37169 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 139 | val set_default_raw_param = | 
| 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 140 | Data.map o fold (AList.update (op =)) o normalize_raw_param | 
| 35964 | 141 | val default_raw_params = Data.get | 
| 33192 | 142 | |
| 143 | fun is_punctuation s = (s = "," orelse s = "-" orelse s = "\<midarrow>") | |
| 144 | ||
| 145 | fun stringify_raw_param_value [] = "" | |
| 146 | | stringify_raw_param_value [s] = s | |
| 147 | | stringify_raw_param_value (s1 :: s2 :: ss) = | |
| 148 | s1 ^ (if is_punctuation s1 orelse is_punctuation s2 then "" else " ") ^ | |
| 149 | stringify_raw_param_value (s2 :: ss) | |
| 150 | ||
| 151 | fun maxed_int_from_string min_int s = Int.max (min_int, the (Int.fromString s)) | |
| 152 | ||
| 153 | fun extract_params ctxt auto default_params override_params = | |
| 154 | let | |
| 37169 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 155 | val override_params = maps normalize_raw_param override_params | 
| 33192 | 156 | val raw_params = rev override_params @ rev default_params | 
| 37260 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 157 | val raw_lookup = AList.lookup (op =) raw_params | 
| 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 158 | val lookup = Option.map stringify_raw_param_value o raw_lookup | 
| 35964 | 159 | val lookup_string = the_default "" o lookup | 
| 160 | fun general_lookup_bool option default_value name = | |
| 161 | case lookup name of | |
| 36380 
1e8fcaccb3e8
stop referring to Sledgehammer_Util stuff all over Nitpick code; instead, redeclare any needed function in Nitpick_Util as synonym for the Sledgehammer_Util function of the same name
 blanchet parents: 
35968diff
changeset | 162 | SOME s => parse_bool_option option name s | 
| 35964 | 163 | | NONE => default_value | 
| 164 | val lookup_bool = the o general_lookup_bool false (SOME false) | |
| 165 | val lookup_bool_option = general_lookup_bool true NONE | |
| 33192 | 166 | fun do_int name value = | 
| 167 | case value of | |
| 168 | SOME s => (case Int.fromString s of | |
| 169 | SOME i => i | |
| 170 |                    | NONE => error ("Parameter " ^ quote name ^
 | |
| 171 | " must be assigned an integer value.")) | |
| 172 | | NONE => 0 | |
| 173 | fun lookup_int name = do_int name (lookup name) | |
| 174 | fun lookup_int_option name = | |
| 175 | case lookup name of | |
| 176 | SOME "smart" => NONE | |
| 177 | | value => SOME (do_int name value) | |
| 178 | fun int_range_from_string name min_int s = | |
| 179 | let | |
| 180 | val (k1, k2) = | |
| 181 | (case space_explode "-" s of | |
| 182 | [s] => the_default (s, s) (first_field "\<midarrow>" s) | |
| 183 |            | ["", s2] => ("-" ^ s2, "-" ^ s2)
 | |
| 184 | | [s1, s2] => (s1, s2) | |
| 185 | | _ => raise Option) | |
| 186 | |> pairself (maxed_int_from_string min_int) | |
| 187 | in if k1 <= k2 then k1 upto k2 else k1 downto k2 end | |
| 188 | handle Option.Option => | |
| 189 |              error ("Parameter " ^ quote name ^
 | |
| 190 | " must be assigned a sequence of integers.") | |
| 191 | fun int_seq_from_string name min_int s = | |
| 192 | maps (int_range_from_string name min_int) (space_explode "," s) | |
| 193 | fun lookup_int_seq name min_int = | |
| 194 | case lookup name of | |
| 195 | SOME s => (case int_seq_from_string name min_int s of | |
| 196 | [] => [min_int] | |
| 197 | | value => value) | |
| 198 | | NONE => [min_int] | |
| 37260 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 199 | fun lookup_assigns read prefix default convert = | 
| 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 200 | (NONE, convert (the_default default (lookup prefix))) | 
| 34982 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34936diff
changeset | 201 | :: map (fn (name, value) => | 
| 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34936diff
changeset | 202 | (SOME (read (String.extract (name, size prefix + 1, NONE))), | 
| 37260 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 203 | convert (stringify_raw_param_value value))) | 
| 34982 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34936diff
changeset | 204 | (filter (String.isPrefix (prefix ^ " ") o fst) raw_params) | 
| 37260 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 205 | fun lookup_ints_assigns read prefix min_int = | 
| 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 206 | lookup_assigns read prefix (signed_string_of_int min_int) | 
| 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 207 | (int_seq_from_string prefix min_int) | 
| 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 208 | fun lookup_bool_assigns read prefix = | 
| 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 209 | lookup_assigns read prefix "" (the o parse_bool_option false prefix) | 
| 33192 | 210 | fun lookup_bool_option_assigns read prefix = | 
| 37260 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 211 | lookup_assigns read prefix "" (parse_bool_option true prefix) | 
| 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 212 | fun lookup_strings_assigns read prefix = | 
| 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 213 | lookup_assigns read prefix "" (space_explode " ") | 
| 33192 | 214 | fun lookup_time name = | 
| 215 | case lookup name of | |
| 39316 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 blanchet parents: 
39155diff
changeset | 216 | SOME s => parse_time_option name s | 
| 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 blanchet parents: 
39155diff
changeset | 217 | | NONE => NONE | 
| 33192 | 218 | val read_type_polymorphic = | 
| 219 | Syntax.read_typ ctxt #> Logic.mk_type | |
| 220 | #> singleton (Variable.polymorphic ctxt) #> Logic.dest_type | |
| 221 | val read_term_polymorphic = | |
| 222 | Syntax.read_term ctxt #> singleton (Variable.polymorphic ctxt) | |
| 38209 | 223 | val lookup_term_list_polymorphic = | 
| 224 | AList.lookup (op =) raw_params #> these #> map read_term_polymorphic | |
| 33192 | 225 | val read_const_polymorphic = read_term_polymorphic #> dest_Const | 
| 226 | val cards_assigns = lookup_ints_assigns read_type_polymorphic "card" 1 | |
| 39344 | 227 | |> auto ? map (apsnd (take max_auto_scopes)) | 
| 33192 | 228 | val maxes_assigns = lookup_ints_assigns read_const_polymorphic "max" ~1 | 
| 229 | val iters_assigns = lookup_ints_assigns read_const_polymorphic "iter" 0 | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 230 | val bitss = lookup_int_seq "bits" 1 | 
| 33192 | 231 | val bisim_depths = lookup_int_seq "bisim_depth" ~1 | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35280diff
changeset | 232 | val boxes = lookup_bool_option_assigns read_type_polymorphic "box" | 
| 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35280diff
changeset | 233 | val finitizes = lookup_bool_option_assigns read_type_polymorphic "finitize" | 
| 39316 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 blanchet parents: 
39155diff
changeset | 234 | val monos = if auto then [(NONE, SOME true)] | 
| 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 blanchet parents: 
39155diff
changeset | 235 | else lookup_bool_option_assigns read_type_polymorphic "mono" | 
| 34982 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34936diff
changeset | 236 | val stds = lookup_bool_assigns read_type_polymorphic "std" | 
| 33192 | 237 | val wfs = lookup_bool_option_assigns read_const_polymorphic "wf" | 
| 238 | val sat_solver = lookup_string "sat_solver" | |
| 39316 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 blanchet parents: 
39155diff
changeset | 239 | val blocking = auto orelse lookup_bool "blocking" | 
| 33192 | 240 | val falsify = lookup_bool "falsify" | 
| 241 | val debug = not auto andalso lookup_bool "debug" | |
| 242 | val verbose = debug orelse (not auto andalso lookup_bool "verbose") | |
| 243 | val overlord = lookup_bool "overlord" | |
| 244 | val user_axioms = lookup_bool_option "user_axioms" | |
| 245 | val assms = lookup_bool "assms" | |
| 38209 | 246 | val whacks = lookup_term_list_polymorphic "whack" | 
| 33556 
cba22e2999d5
renamed Nitpick option "coalesce_type_vars" to "merge_type_vars" (shorter) and cleaned up old hacks that are no longer necessary
 blanchet parents: 
33232diff
changeset | 247 | val merge_type_vars = lookup_bool "merge_type_vars" | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 248 | val binary_ints = lookup_bool_option "binary_ints" | 
| 33192 | 249 | val destroy_constrs = lookup_bool "destroy_constrs" | 
| 250 | val specialize = lookup_bool "specialize" | |
| 251 | val star_linear_preds = lookup_bool "star_linear_preds" | |
| 252 | val peephole_optim = lookup_bool "peephole_optim" | |
| 38124 | 253 | val datatype_sym_break = lookup_int "datatype_sym_break" | 
| 254 | val kodkod_sym_break = lookup_int "kodkod_sym_break" | |
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 255 | val timeout = if auto then NONE else lookup_time "timeout" | 
| 33192 | 256 | val tac_timeout = lookup_time "tac_timeout" | 
| 39316 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 blanchet parents: 
39155diff
changeset | 257 | val max_threads = if auto then 1 else Int.max (0, lookup_int "max_threads") | 
| 37169 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 258 | val show_datatypes = debug orelse lookup_bool "show_datatypes" | 
| 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
36960diff
changeset | 259 | val show_consts = debug orelse lookup_bool "show_consts" | 
| 33192 | 260 | val formats = lookup_ints_assigns read_term_polymorphic "format" 0 | 
| 37260 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 261 | val atomss = lookup_strings_assigns read_type_polymorphic "atoms" | 
| 38209 | 262 | val evals = lookup_term_list_polymorphic "eval" | 
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 263 | val max_potential = | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 264 | if auto then 0 else Int.max (0, lookup_int "max_potential") | 
| 33192 | 265 | val max_genuine = Int.max (0, lookup_int "max_genuine") | 
| 266 | val check_potential = lookup_bool "check_potential" | |
| 267 | val check_genuine = lookup_bool "check_genuine" | |
| 36390 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36389diff
changeset | 268 | val batch_size = | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36389diff
changeset | 269 | case lookup_int_option "batch_size" of | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36389diff
changeset | 270 | SOME n => Int.max (1, n) | 
| 38180 
7a88032f9265
bump up the max cardinalities, to use up more of the time given to us by the user
 blanchet parents: 
38127diff
changeset | 271 | | NONE => if debug then 1 else 50 | 
| 33192 | 272 | val expect = lookup_string "expect" | 
| 273 | in | |
| 274 |     {cards_assigns = cards_assigns, maxes_assigns = maxes_assigns,
 | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 275 | iters_assigns = iters_assigns, bitss = bitss, bisim_depths = bisim_depths, | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35280diff
changeset | 276 | boxes = boxes, finitizes = finitizes, monos = monos, stds = stds, | 
| 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35280diff
changeset | 277 | wfs = wfs, sat_solver = sat_solver, blocking = blocking, falsify = falsify, | 
| 34982 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34936diff
changeset | 278 | debug = debug, verbose = verbose, overlord = overlord, | 
| 38209 | 279 | user_axioms = user_axioms, assms = assms, whacks = whacks, | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 280 | merge_type_vars = merge_type_vars, binary_ints = binary_ints, | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 281 | destroy_constrs = destroy_constrs, specialize = specialize, | 
| 39359 | 282 | star_linear_preds = star_linear_preds, peephole_optim = peephole_optim, | 
| 283 | datatype_sym_break = datatype_sym_break, | |
| 38124 | 284 | kodkod_sym_break = kodkod_sym_break, timeout = timeout, | 
| 36389 
8228b3a4a2ba
remove "skolemize" option from Nitpick, since Skolemization is always useful
 blanchet parents: 
36388diff
changeset | 285 | tac_timeout = tac_timeout, max_threads = max_threads, | 
| 36390 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36389diff
changeset | 286 | show_datatypes = show_datatypes, show_consts = show_consts, | 
| 37260 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 287 | formats = formats, atomss = atomss, evals = evals, | 
| 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 288 | max_potential = max_potential, max_genuine = max_genuine, | 
| 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 289 | check_potential = check_potential, check_genuine = check_genuine, | 
| 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37213diff
changeset | 290 | batch_size = batch_size, expect = expect} | 
| 33192 | 291 | end | 
| 292 | ||
| 293 | fun default_params thy = | |
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36607diff
changeset | 294 | extract_params (ProofContext.init_global thy) false (default_raw_params thy) | 
| 33192 | 295 | o map (apsnd single) | 
| 296 | ||
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 297 | val parse_key = Scan.repeat1 Parse.typ_group >> space_implode " " | 
| 35964 | 298 | val parse_value = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 299 | Scan.repeat1 (Parse.minus >> single | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 300 | || Scan.repeat1 (Scan.unless Parse.minus Parse.name) | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 301 | || Parse.$$$ "," |-- Parse.number >> prefix "," >> single) >> flat | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 302 | val parse_param = parse_key -- Scan.optional (Parse.$$$ "=" |-- parse_value) [] | 
| 35968 
b7f98ff9c7d9
simplify Nitpick parameter parsing code a little bit + make compile
 blanchet parents: 
35964diff
changeset | 303 | val parse_params = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 304 | Scan.optional (Parse.$$$ "[" |-- Parse.list parse_param --| Parse.$$$ "]") [] | 
| 33192 | 305 | |
| 306 | fun handle_exceptions ctxt f x = | |
| 307 | f x | |
| 308 | handle ARG (loc, details) => | |
| 309 |          error ("Bad argument(s) to " ^ quote loc ^ ": " ^ details ^ ".")
 | |
| 310 | | BAD (loc, details) => | |
| 311 |          error ("Internal error (" ^ quote loc ^ "): " ^ details ^ ".")
 | |
| 312 | | NOT_SUPPORTED details => | |
| 313 |          (warning ("Unsupported case: " ^ details ^ "."); x)
 | |
| 314 | | NUT (loc, us) => | |
| 315 |          error ("Invalid intermediate term" ^ plural_s_for_list us ^
 | |
| 316 |                 " (" ^ quote loc ^ "): " ^
 | |
| 317 | commas (map (string_for_nut ctxt) us) ^ ".") | |
| 318 | | REP (loc, Rs) => | |
| 319 |          error ("Invalid representation" ^ plural_s_for_list Rs ^
 | |
| 320 |                 " (" ^ quote loc ^ "): " ^ commas (map string_for_rep Rs) ^ ".")
 | |
| 321 | | TERM (loc, ts) => | |
| 322 |          error ("Invalid term" ^ plural_s_for_list ts ^
 | |
| 323 |                 " (" ^ quote loc ^ "): " ^
 | |
| 324 | commas (map (Syntax.string_of_term ctxt) ts) ^ ".") | |
| 325 | | TYPE (loc, Ts, ts) => | |
| 326 |          error ("Invalid type" ^ plural_s_for_list Ts ^
 | |
| 327 | (if null ts then | |
| 328 | "" | |
| 329 | else | |
| 330 | " for term" ^ plural_s_for_list ts ^ " " ^ | |
| 331 | commas (map (quote o Syntax.string_of_term ctxt) ts)) ^ | |
| 332 |                 " (" ^ quote loc ^ "): " ^
 | |
| 333 | commas (map (Syntax.string_of_typ ctxt) Ts) ^ ".") | |
| 334 | | Refute.REFUTE (loc, details) => | |
| 335 |          error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^ ".")
 | |
| 336 | ||
| 34982 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34936diff
changeset | 337 | fun pick_nits override_params auto i step state = | 
| 33192 | 338 | let | 
| 339 | val thy = Proof.theory_of state | |
| 340 | val ctxt = Proof.context_of state | |
| 341 | val _ = List.app check_raw_param override_params | |
| 342 |     val params as {blocking, debug, ...} =
 | |
| 343 | extract_params ctxt auto (default_raw_params thy) override_params | |
| 344 | fun go () = | |
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 345 | (false, state) | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 346 | |> (if auto then perhaps o try | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 347 | else if debug then fn f => fn x => f x | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 348 | else handle_exceptions ctxt) | 
| 34982 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34936diff
changeset | 349 | (fn (_, state) => pick_nits_in_subgoal state params auto i step | 
| 34121 
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
 blanchet parents: 
33982diff
changeset | 350 | |>> curry (op =) "genuine") | 
| 39316 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 blanchet parents: 
39155diff
changeset | 351 | in if blocking then go () else Future.fork (tap go) |> K (false, state) end | 
| 33192 | 352 | |
| 35968 
b7f98ff9c7d9
simplify Nitpick parameter parsing code a little bit + make compile
 blanchet parents: 
35964diff
changeset | 353 | fun nitpick_trans (params, i) = | 
| 34982 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34936diff
changeset | 354 | Toplevel.keep (fn st => | 
| 35968 
b7f98ff9c7d9
simplify Nitpick parameter parsing code a little bit + make compile
 blanchet parents: 
35964diff
changeset | 355 | (pick_nits params false i (Toplevel.proof_position_of st) | 
| 
b7f98ff9c7d9
simplify Nitpick parameter parsing code a little bit + make compile
 blanchet parents: 
35964diff
changeset | 356 | (Toplevel.proof_of st); ())) | 
| 33192 | 357 | |
| 358 | fun string_for_raw_param (name, value) = | |
| 359 | name ^ " = " ^ stringify_raw_param_value value | |
| 360 | ||
| 35968 
b7f98ff9c7d9
simplify Nitpick parameter parsing code a little bit + make compile
 blanchet parents: 
35964diff
changeset | 361 | fun nitpick_params_trans params = | 
| 33192 | 362 | Toplevel.theory | 
| 35968 
b7f98ff9c7d9
simplify Nitpick parameter parsing code a little bit + make compile
 blanchet parents: 
35964diff
changeset | 363 | (fold set_default_raw_param params | 
| 39316 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 blanchet parents: 
39155diff
changeset | 364 | #> tap (fn thy => | 
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 365 |                   writeln ("Default parameters for Nitpick:\n" ^
 | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 366 | (case rev (default_raw_params thy) of | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 367 | [] => "none" | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 368 | | params => | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 369 | (map check_raw_param params; | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 370 | params |> map string_for_raw_param | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 371 | |> sort_strings |> cat_lines))))) | 
| 33192 | 372 | |
| 35964 | 373 | val parse_nitpick_command = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 374 | (parse_params -- Scan.optional Parse.nat 1) #>> nitpick_trans | 
| 35964 | 375 | val parse_nitpick_params_command = parse_params #>> nitpick_params_trans | 
| 33192 | 376 | |
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 377 | val _ = Outer_Syntax.improper_command "nitpick" | 
| 36390 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36389diff
changeset | 378 | "try to find a counterexample for a given subgoal using Nitpick" | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 379 | Keyword.diag parse_nitpick_command | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 380 | val _ = Outer_Syntax.command "nitpick_params" | 
| 33192 | 381 | "set and display the default parameters for Nitpick" | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 382 | Keyword.thy_decl parse_nitpick_params_command | 
| 33192 | 383 | |
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 384 | fun auto_nitpick state = | 
| 34982 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34936diff
changeset | 385 | if not (!auto) then (false, state) else pick_nits [] true 1 0 state | 
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 386 | |
| 39324 
05452dd66b2b
finished renaming "Auto_Counterexample" to "Auto_Tools"
 blanchet parents: 
39316diff
changeset | 387 | val setup = Auto_Tools.register_tool ("nitpick", auto_nitpick)
 | 
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33556diff
changeset | 388 | |
| 33192 | 389 | end; |