author | wenzelm |
Mon, 12 Sep 2022 23:24:50 +0200 | |
changeset 76125 | 497e105a4618 |
parent 74870 | d54b3c96ee50 |
child 78709 | ebafb2daabb7 |
permissions | -rw-r--r-- |
43018 | 1 |
(* Title: Tools/try.ML |
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
2 |
Author: Jasmin Blanchette, TU Muenchen |
74507
a241eadc0e3f
removed unused material (left-over from fd0c85d7da38);
wenzelm
parents:
69593
diff
changeset
|
3 |
Author: Makarius |
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
4 |
|
43018 | 5 |
Manager for tools that should be tried on conjectures. |
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
6 |
*) |
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
7 |
|
43018 | 8 |
signature TRY = |
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
9 |
sig |
63690
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
10 |
val serial_commas: string -> string list -> string list |
74508 | 11 |
type body = bool -> Proof.state -> bool * (string * string list) |
12 |
type tool = {name: string, weight: int, auto_option: string, body: body} |
|
63690
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
13 |
val get_tools: theory -> tool list |
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
14 |
val try_tools: Proof.state -> (string * string) option |
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
15 |
val tool_setup: tool -> unit |
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
16 |
end; |
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
17 |
|
43018 | 18 |
structure Try : TRY = |
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
19 |
struct |
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
20 |
|
43028
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
21 |
(* helpers *) |
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
22 |
|
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
23 |
fun serial_commas _ [] = ["??"] |
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
24 |
| serial_commas _ [s] = [s] |
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
25 |
| serial_commas conj [s1, s2] = [s1, conj, s2] |
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
26 |
| serial_commas conj [s1, s2, s3] = [s1 ^ ",", s2 ^ ",", conj, s3] |
63690
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
27 |
| serial_commas conj (s :: ss) = s ^ "," :: serial_commas conj ss; |
43028
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
28 |
|
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
29 |
|
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
30 |
(* configuration *) |
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
31 |
|
74508 | 32 |
type body = bool -> Proof.state -> bool * (string * string list); |
33 |
type tool = {name: string, weight: int, auto_option: string, body: body}; |
|
34 |
||
35 |
fun tool_ord (tool1: tool, tool2: tool) = |
|
36 |
prod_ord int_ord string_ord ((#weight tool1, #name tool1), (#weight tool2, #name tool2)); |
|
43028
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
37 |
|
33600 | 38 |
structure Data = Theory_Data |
39 |
( |
|
63690
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
40 |
type T = tool list; |
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
41 |
val empty = []; |
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
42 |
fun merge data : T = Ord_List.merge tool_ord data; |
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
43 |
); |
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
44 |
|
63690
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
45 |
val get_tools = Data.get; |
43020
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43018
diff
changeset
|
46 |
|
63690
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
47 |
val register_tool = Data.map o Ord_List.insert tool_ord; |
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
48 |
|
52641
c56b6fa636e8
hybrid "auto" tool setup, for TTY (within theory) and PIDE (global print function);
wenzelm
parents:
52639
diff
changeset
|
49 |
|
43020
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43018
diff
changeset
|
50 |
(* try command *) |
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43018
diff
changeset
|
51 |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43018
diff
changeset
|
52 |
fun try_tools state = |
74510 | 53 |
if Proof.goal_finished state then |
58843 | 54 |
(writeln "No subgoal!"; NONE) |
43028
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
55 |
else |
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
56 |
get_tools (Proof.theory_of state) |
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
57 |
|> tap (fn tools => |
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
58 |
"Trying " ^ space_implode " " |
74508 | 59 |
(serial_commas "and" (map (quote o #name) tools)) ^ "..." |
58843 | 60 |
|> writeln) |
43028
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
61 |
|> Par_List.get_some |
74508 | 62 |
(fn {name, body, ...} => |
63 |
case try (body false) state of |
|
43028
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
64 |
SOME (true, (outcome_code, _)) => SOME (name, outcome_code) |
1c451bbb3ad7
repaired theory merging and defined/used helpers
blanchet
parents:
43024
diff
changeset
|
65 |
| _ => NONE) |
63690
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
66 |
|> tap (fn NONE => writeln "Tried in vain" | _ => ()); |
43020
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43018
diff
changeset
|
67 |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43018
diff
changeset
|
68 |
val _ = |
67149 | 69 |
Outer_Syntax.command \<^command_keyword>\<open>try\<close> |
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
43061
diff
changeset
|
70 |
"try a combination of automatic proving and disproving tools" |
63690
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
71 |
(Scan.succeed (Toplevel.keep_proof (ignore o try_tools o Toplevel.proof_of))); |
43020
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43018
diff
changeset
|
72 |
|
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
73 |
|
74507
a241eadc0e3f
removed unused material (left-over from fd0c85d7da38);
wenzelm
parents:
69593
diff
changeset
|
74 |
(* asynchronous print function *) |
52641
c56b6fa636e8
hybrid "auto" tool setup, for TTY (within theory) and PIDE (global print function);
wenzelm
parents:
52639
diff
changeset
|
75 |
|
74508 | 76 |
fun print_function ({name, weight, auto_option, body}: tool) = |
53839
274a892b1230
avoid clash of auto print functions with query operations, notably sledgehammer (cf. 3461985dcbc3);
wenzelm
parents:
53171
diff
changeset
|
77 |
Command.print_function ("auto_" ^ name) |
58923 | 78 |
(fn {keywords, command_name, ...} => |
74508 | 79 |
if Keyword.is_theory_goal keywords command_name andalso Options.default_bool auto_option then |
52651 | 80 |
SOME |
69593 | 81 |
{delay = SOME (seconds (Options.default_real \<^system_option>\<open>auto_time_start\<close>)), |
52651 | 82 |
pri = ~ weight, |
83 |
persistent = true, |
|
52953
2c927b7501c5
explicit "strict" flag for print functions (flipped internal meaning);
wenzelm
parents:
52850
diff
changeset
|
84 |
strict = true, |
52651 | 85 |
print_fn = fn _ => fn st => |
86 |
let |
|
87 |
val state = Toplevel.proof_of st |
|
52652
ebdbd5c79a13
avoid spurious warnings, notably of 'try0' about already declared simps etc.;
wenzelm
parents:
52651
diff
changeset
|
88 |
|> Proof.map_context (Context_Position.set_visible false) |
69593 | 89 |
val auto_time_limit = Options.default_real \<^system_option>\<open>auto_time_limit\<close> |
52651 | 90 |
in |
91 |
if auto_time_limit > 0.0 then |
|
74870
d54b3c96ee50
more robust physical timeout (despite 1bea05713dde), especially relevant for quickcheck where large unary numerals may cause excessive heap allocations and resulting GC is better included in the timing;
wenzelm
parents:
74561
diff
changeset
|
92 |
(case Timeout.apply_physical (seconds auto_time_limit) (fn () => body true state) () of |
59184
830bb7ddb3ab
explicit message channels for "state", "information";
wenzelm
parents:
58923
diff
changeset
|
93 |
(true, (_, outcome)) => List.app Output.information outcome |
52651 | 94 |
| _ => ()) |
95 |
else () |
|
62505 | 96 |
end handle exn => if Exn.is_interrupt exn then Exn.reraise exn else ()} |
63690
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
97 |
else NONE); |
52641
c56b6fa636e8
hybrid "auto" tool setup, for TTY (within theory) and PIDE (global print function);
wenzelm
parents:
52639
diff
changeset
|
98 |
|
c56b6fa636e8
hybrid "auto" tool setup, for TTY (within theory) and PIDE (global print function);
wenzelm
parents:
52639
diff
changeset
|
99 |
|
74508 | 100 |
(* tool setup *) |
52641
c56b6fa636e8
hybrid "auto" tool setup, for TTY (within theory) and PIDE (global print function);
wenzelm
parents:
52639
diff
changeset
|
101 |
|
63690
48a2c88091d7
tuning punctuation in messages output by Isabelle
blanchet
parents:
62519
diff
changeset
|
102 |
fun tool_setup tool = (Theory.setup (register_tool tool); print_function tool); |
52641
c56b6fa636e8
hybrid "auto" tool setup, for TTY (within theory) and PIDE (global print function);
wenzelm
parents:
52639
diff
changeset
|
103 |
|
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff
changeset
|
104 |
end; |