author | desharna |
Thu, 27 Mar 2025 13:30:16 +0100 | |
changeset 82360 | 6a09257afd06 |
parent 82024 | bbda3b4f3c99 |
child 82363 | 3a7fc54b50ca |
permissions | -rw-r--r-- |
66646 | 1 |
(* Title: HOL/Tools/Nunchaku/nunchaku_commands.ML |
66614
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
blanchet
parents:
66163
diff
changeset
|
2 |
Author: Jasmin Blanchette, VU Amsterdam |
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
blanchet
parents:
66163
diff
changeset
|
3 |
Copyright 2015, 2016, 2017 |
64389 | 4 |
|
5 |
Adds the "nunchaku" and "nunchaku_params" commands to Isabelle/Isar's outer syntax. |
|
6 |
*) |
|
7 |
||
8 |
signature NUNCHAKU_COMMANDS = |
|
9 |
sig |
|
10 |
type params = Nunchaku.params |
|
11 |
||
12 |
val default_params: theory -> (string * string) list -> params |
|
13 |
end; |
|
14 |
||
15 |
structure Nunchaku_Commands : NUNCHAKU_COMMANDS = |
|
16 |
struct |
|
17 |
||
18 |
open Nunchaku_Util; |
|
19 |
open Nunchaku; |
|
20 |
||
21 |
type raw_param = string * string list; |
|
22 |
||
23 |
val default_default_params = |
|
24 |
[("assms", "true"), |
|
66631 | 25 |
("bound_increment", "4"), |
64389 | 26 |
("debug", "false"), |
27 |
("falsify", "true"), |
|
66631 | 28 |
("min_bound", "4"), |
66627 | 29 |
("max_bound", "none"), |
64389 | 30 |
("max_genuine", "1"), |
31 |
("max_potential", "1"), |
|
32 |
("overlord", "false"), |
|
82024
bbda3b4f3c99
switch from CVC5 to cvc5, including updates of internal tool references;
wenzelm
parents:
80910
diff
changeset
|
33 |
("solvers", "cvc5 kodkod paradox smbc"), |
64389 | 34 |
("specialize", "true"), |
35 |
("spy", "false"), |
|
36 |
("timeout", "30"), |
|
37 |
("verbose", "false"), |
|
38 |
("wf_timeout", "0.5")]; |
|
39 |
||
66619 | 40 |
val alias_params = |
41 |
[("solver", "solvers")]; |
|
42 |
||
43 |
val negated_alias_params = |
|
64389 | 44 |
[("dont_whack", "whack"), |
45 |
("dont_specialize", "specialize"), |
|
46 |
("dont_spy", "spy"), |
|
47 |
("no_assms", "assms"), |
|
48 |
("no_debug", "debug"), |
|
49 |
("no_overlord", "overlord"), |
|
50 |
("non_mono", "mono"), |
|
51 |
("non_wf", "wf"), |
|
52 |
("quiet", "verbose"), |
|
53 |
("satisfy", "falsify")]; |
|
54 |
||
55 |
fun is_known_raw_param s = |
|
56 |
AList.defined (op =) default_default_params s orelse |
|
66619 | 57 |
AList.defined (op =) alias_params s orelse |
58 |
AList.defined (op =) negated_alias_params s orelse |
|
67405
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents:
67399
diff
changeset
|
59 |
member (op =) ["atoms", "card", "eval", "expect"] s orelse |
64389 | 60 |
exists (fn p => String.isPrefix (p ^ " ") s) |
61 |
["atoms", "card", "dont_whack", "mono", "non_mono", "non_wf", "wf", "whack"]; |
|
62 |
||
63 |
fun check_raw_param (s, _) = |
|
64 |
if is_known_raw_param s then () else error ("Unknown parameter: " ^ quote s); |
|
65 |
||
66 |
fun unnegate_param_name name = |
|
66619 | 67 |
(case AList.lookup (op =) negated_alias_params name of |
64389 | 68 |
NONE => |
69 |
if String.isPrefix "dont_" name then SOME (unprefix "dont_" name) |
|
70 |
else if String.isPrefix "non_" name then SOME (unprefix "non_" name) |
|
71 |
else NONE |
|
72 |
| some_name => some_name); |
|
73 |
||
74 |
fun normalize_raw_param (name, value) = |
|
66619 | 75 |
(case AList.lookup (op =) alias_params name of |
76 |
SOME alias => [(alias, value)] |
|
77 |
| NONE => |
|
78 |
(case unnegate_param_name name of |
|
79 |
SOME name' => |
|
80 |
[(name', |
|
81 |
(case value of |
|
82 |
["false"] => ["true"] |
|
83 |
| ["true"] => ["false"] |
|
84 |
| [] => ["false"] |
|
85 |
| _ => value))] |
|
86 |
| NONE => [(name, value)])); |
|
64389 | 87 |
|
88 |
structure Data = Theory_Data |
|
89 |
( |
|
90 |
type T = raw_param list |
|
91 |
val empty = default_default_params |> map (apsnd single) |
|
92 |
fun merge data = AList.merge (op =) (K true) data |
|
93 |
); |
|
94 |
||
95 |
val set_default_raw_param = Data.map o fold (AList.update (op =)) o normalize_raw_param; |
|
96 |
val default_raw_params = Data.get; |
|
97 |
||
98 |
fun is_punctuation s = (s = "," orelse s = "-"); |
|
99 |
||
100 |
fun stringify_raw_param_value [] = "" |
|
101 |
| stringify_raw_param_value [s] = s |
|
102 |
| stringify_raw_param_value (s1 :: s2 :: ss) = |
|
103 |
s1 ^ (if is_punctuation s1 orelse is_punctuation s2 then "" else " ") ^ |
|
104 |
stringify_raw_param_value (s2 :: ss); |
|
105 |
||
106 |
fun extract_params ctxt mode default_params override_params = |
|
107 |
let |
|
108 |
val override_params = maps normalize_raw_param override_params; |
|
109 |
val raw_params = rev override_params @ rev default_params; |
|
110 |
val raw_lookup = AList.lookup (op =) raw_params; |
|
111 |
val lookup = Option.map stringify_raw_param_value o raw_lookup; |
|
112 |
val lookup_string = the_default "" o lookup; |
|
66163 | 113 |
val lookup_strings = these o Option.map (space_explode " ") o lookup; |
64389 | 114 |
|
115 |
fun general_lookup_bool option default_value name = |
|
116 |
(case lookup name of |
|
117 |
SOME s => parse_bool_option option name s |
|
118 |
| NONE => default_value); |
|
119 |
||
120 |
val lookup_bool = the o general_lookup_bool false (SOME false); |
|
121 |
||
122 |
fun lookup_int name = |
|
123 |
(case lookup name of |
|
124 |
SOME s => |
|
125 |
(case Int.fromString s of |
|
126 |
SOME i => i |
|
127 |
| NONE => error ("Parameter " ^ quote name ^ " must be assigned an integer value")) |
|
128 |
| NONE => 0); |
|
129 |
||
130 |
fun int_range_from_string name s = |
|
131 |
(case space_explode "-" s of |
|
132 |
[s] => (s, s) |
|
133 |
| [s1, s2] => (s1, s2) |
|
134 |
| _ => error ("Parameter " ^ quote name ^ " must be assigned a range of integers")) |
|
135 |
|> apply2 Int.fromString; |
|
136 |
||
137 |
fun lookup_assigns read pre of_str = |
|
138 |
(case lookup pre of |
|
139 |
SOME s => [(NONE, of_str s)] |
|
140 |
| NONE => []) @ |
|
141 |
map (fn (name, value) => (SOME (read (String.extract (name, size pre + 1, NONE))), |
|
142 |
of_str (stringify_raw_param_value value))) |
|
143 |
(filter (String.isPrefix (pre ^ " ") o fst) raw_params); |
|
144 |
||
145 |
fun lookup_int_range_assigns read pre = |
|
146 |
lookup_assigns read pre (int_range_from_string pre); |
|
147 |
||
148 |
fun lookup_bool_assigns read pre = |
|
149 |
lookup_assigns read pre (the o parse_bool_option false pre); |
|
150 |
||
151 |
fun lookup_bool_option_assigns read pre = |
|
152 |
lookup_assigns read pre (parse_bool_option true pre); |
|
153 |
||
154 |
fun lookup_strings_assigns read pre = |
|
155 |
lookup_assigns read pre (space_explode " "); |
|
156 |
||
157 |
fun lookup_time name = |
|
158 |
(case lookup name of |
|
159 |
SOME s => parse_time name s |
|
160 |
| NONE => Time.zeroTime); |
|
161 |
||
162 |
val read_type_polymorphic = |
|
163 |
Syntax.read_typ ctxt #> Logic.mk_type |
|
164 |
#> singleton (Variable.polymorphic ctxt) #> Logic.dest_type; |
|
165 |
val read_term_polymorphic = |
|
166 |
Syntax.read_term ctxt #> singleton (Variable.polymorphic ctxt); |
|
66627 | 167 |
|
64389 | 168 |
val lookup_term_list_option_polymorphic = |
169 |
AList.lookup (op =) raw_params #> Option.map (map read_term_polymorphic); |
|
170 |
||
171 |
fun read_const_polymorphic s = |
|
172 |
(case read_term_polymorphic s of |
|
173 |
Const x => x |
|
174 |
| t => error ("Not a constant: " ^ Syntax.string_of_term ctxt t)); |
|
175 |
||
66627 | 176 |
fun lookup_none_option lookup' name = |
177 |
(case lookup name of |
|
178 |
SOME "none" => NONE |
|
179 |
| _ => SOME (lookup' name)) |
|
180 |
||
66163 | 181 |
val solvers = lookup_strings "solvers"; |
182 |
val falsify = lookup_bool "falsify"; |
|
183 |
val assms = lookup_bool "assms"; |
|
184 |
val spy = getenv "NUNCHAKU_SPY" = "yes" orelse lookup_bool "spy"; |
|
185 |
val overlord = lookup_bool "overlord"; |
|
186 |
val expect = lookup_string "expect"; |
|
187 |
||
64389 | 188 |
val wfs = lookup_bool_option_assigns read_const_polymorphic "wf"; |
66627 | 189 |
val min_bound = lookup_int "min_bound"; |
190 |
val max_bound = lookup_none_option lookup_int "max_bound"; |
|
66625 | 191 |
val bound_increment = lookup_int "bound_increment"; |
64389 | 192 |
val whacks = lookup_bool_assigns read_term_polymorphic "whack"; |
193 |
val cards = lookup_int_range_assigns read_type_polymorphic "card"; |
|
194 |
val monos = lookup_bool_option_assigns read_type_polymorphic "mono"; |
|
66163 | 195 |
|
64389 | 196 |
val debug = (mode <> Auto_Try andalso lookup_bool "debug"); |
197 |
val verbose = debug orelse (mode <> Auto_Try andalso lookup_bool "verbose"); |
|
66163 | 198 |
val max_potential = if mode = Normal then Int.max (0, lookup_int "max_potential") else 0; |
199 |
val max_genuine = Int.max (0, lookup_int "max_genuine"); |
|
200 |
val evals = these (lookup_term_list_option_polymorphic "eval"); |
|
201 |
val atomss = lookup_strings_assigns read_type_polymorphic "atoms"; |
|
202 |
||
64389 | 203 |
val specialize = lookup_bool "specialize"; |
66163 | 204 |
val multithread = mode = Normal andalso lookup_bool "multithread"; |
205 |
||
64389 | 206 |
val timeout = lookup_time "timeout"; |
207 |
val wf_timeout = lookup_time "wf_timeout"; |
|
208 |
||
209 |
val mode_of_operation_params = |
|
66163 | 210 |
{solvers = solvers, falsify = falsify, assms = assms, spy = spy, overlord = overlord, |
64389 | 211 |
expect = expect}; |
212 |
||
213 |
val scope_of_search_params = |
|
66627 | 214 |
{wfs = wfs, whacks = whacks, min_bound = min_bound, max_bound = max_bound, |
215 |
bound_increment = bound_increment, cards = cards, monos = monos}; |
|
64389 | 216 |
|
217 |
val output_format_params = |
|
218 |
{verbose = verbose, debug = debug, max_potential = max_potential, max_genuine = max_genuine, |
|
219 |
evals = evals, atomss = atomss}; |
|
220 |
||
221 |
val optimization_params = |
|
222 |
{specialize = specialize, multithread = multithread}; |
|
223 |
||
224 |
val timeout_params = |
|
225 |
{timeout = timeout, wf_timeout = wf_timeout}; |
|
226 |
in |
|
227 |
{mode_of_operation_params = mode_of_operation_params, |
|
228 |
scope_of_search_params = scope_of_search_params, |
|
229 |
output_format_params = output_format_params, |
|
230 |
optimization_params = optimization_params, |
|
231 |
timeout_params = timeout_params} |
|
232 |
end; |
|
233 |
||
234 |
fun default_params thy = |
|
235 |
extract_params (Proof_Context.init_global thy) Normal (default_raw_params thy) |
|
236 |
o map (apsnd single); |
|
237 |
||
80910 | 238 |
val parse_key = Scan.repeat1 Parse.embedded >> implode_space; |
64389 | 239 |
val parse_value = |
240 |
Scan.repeat1 (Parse.minus >> single |
|
241 |
|| Scan.repeat1 (Scan.unless Parse.minus (Parse.name || Parse.float_number)) |
|
69593 | 242 |
|| \<^keyword>\<open>,\<close> |-- Parse.number >> prefix "," >> single) |
64389 | 243 |
>> flat; |
69593 | 244 |
val parse_param = parse_key -- Scan.optional (\<^keyword>\<open>=\<close> |-- parse_value) []; |
245 |
val parse_params = Scan.optional (\<^keyword>\<open>[\<close> |-- Parse.list parse_param --| \<^keyword>\<open>]\<close>) []; |
|
64389 | 246 |
|
247 |
fun run_chaku override_params mode i state0 = |
|
248 |
let |
|
82360 | 249 |
val state = Proof.map_contexts (Try0_Util.silence_methods false) state0; |
64389 | 250 |
val thy = Proof.theory_of state; |
251 |
val ctxt = Proof.context_of state; |
|
252 |
val _ = List.app check_raw_param override_params; |
|
253 |
val params = extract_params ctxt mode (default_raw_params thy) override_params; |
|
254 |
in |
|
255 |
(if mode = Auto_Try then perhaps o try else fn f => fn x => f x) |
|
256 |
(fn _ => run_chaku_on_subgoal state params mode i) (unknownN, NONE) |
|
257 |
|> `(fn (outcome_code, _) => outcome_code = genuineN) |
|
258 |
end; |
|
259 |
||
260 |
fun string_for_raw_param (name, value) = |
|
261 |
name ^ " = " ^ stringify_raw_param_value value; |
|
262 |
||
263 |
fun nunchaku_params_trans params = |
|
264 |
Toplevel.theory (fold set_default_raw_param params |
|
265 |
#> tap (fn thy => |
|
266 |
let val params = rev (default_raw_params thy) in |
|
267 |
List.app check_raw_param params; |
|
268 |
writeln ("Default parameters for Nunchaku:\n" ^ |
|
269 |
(params |> map string_for_raw_param |> sort_strings |> cat_lines)) |
|
270 |
end)); |
|
271 |
||
272 |
val _ = |
|
69593 | 273 |
Outer_Syntax.command \<^command_keyword>\<open>nunchaku\<close> |
64389 | 274 |
"try to find a countermodel using Nunchaku" |
275 |
(parse_params -- Scan.optional Parse.nat 1 >> (fn (params, i) => |
|
276 |
Toplevel.keep_proof (fn state => |
|
277 |
ignore (run_chaku params Normal i (Toplevel.proof_of state))))); |
|
278 |
||
279 |
val _ = |
|
69593 | 280 |
Outer_Syntax.command \<^command_keyword>\<open>nunchaku_params\<close> |
64389 | 281 |
"set and display the default parameters for Nunchaku" |
282 |
(parse_params #>> nunchaku_params_trans); |
|
283 |
||
284 |
end; |