author | desharna |
Thu, 27 Mar 2025 13:30:16 +0100 | |
changeset 82360 | 6a09257afd06 |
parent 82359 | d2960b321468 |
child 82362 | 396676efbd6d |
permissions | -rw-r--r-- |
46641 | 1 |
(* Title: HOL/Tools/try0.ML |
38942
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
2 |
Author: Jasmin Blanchette, TU Muenchen |
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
3 |
|
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
4 |
Try a combination of proof methods. |
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
5 |
*) |
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
6 |
|
46641 | 7 |
signature TRY0 = |
38942
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
8 |
sig |
82360 | 9 |
val serial_commas : string -> string list -> string list |
10 |
||
81369
0677712016b5
try0: add 'use' modifier for thms to insert;
Fabian Huch <huch@in.tum.de>
parents:
81368
diff
changeset
|
11 |
datatype modifier = Use | Simp | Intro | Elim | Dest |
82355 | 12 |
type xref = Facts.ref * Token.src list |
13 |
type tagged_xref = xref * modifier list |
|
82213 | 14 |
type result = {name: string, command: string, time: Time.time, state: Proof.state} |
82358 | 15 |
type proof_method = Time.time option -> tagged_xref list -> Proof.state -> result option |
16 |
type proof_method_options = {run_if_auto_try: bool} |
|
17 |
||
18 |
val register_proof_method : string -> proof_method_options -> proof_method -> unit |
|
19 |
val get_proof_method : string -> proof_method option |
|
20 |
||
82360 | 21 |
datatype mode = Auto_Try | Try | Normal |
22 |
val generic_try0 : mode -> Time.time option -> tagged_xref list -> Proof.state -> |
|
23 |
(bool * (string * string list)) * result list |
|
82355 | 24 |
val try0 : Time.time option -> tagged_xref list -> Proof.state -> result list |
81365 | 25 |
end |
38942
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
26 |
|
46641 | 27 |
structure Try0 : TRY0 = |
38942
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
28 |
struct |
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
29 |
|
82360 | 30 |
fun serial_commas _ [] = ["??"] |
31 |
| serial_commas _ [s] = [s] |
|
32 |
| serial_commas conj [s1, s2] = [s1, conj, s2] |
|
33 |
| serial_commas conj [s1, s2, s3] = [s1 ^ ",", s2 ^ ",", conj, s3] |
|
34 |
| serial_commas conj (s :: ss) = s ^ "," :: serial_commas conj ss; |
|
35 |
||
81365 | 36 |
val noneN = "none" |
43026
0f15575a6465
handle non-auto try cases gracefully in Try Methods
blanchet
parents:
43024
diff
changeset
|
37 |
|
81365 | 38 |
datatype mode = Auto_Try | Try | Normal |
43020
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43018
diff
changeset
|
39 |
|
81365 | 40 |
val default_timeout = seconds 5.0 |
38944 | 41 |
|
81369
0677712016b5
try0: add 'use' modifier for thms to insert;
Fabian Huch <huch@in.tum.de>
parents:
81368
diff
changeset
|
42 |
datatype modifier = Use | Simp | Intro | Elim | Dest |
82355 | 43 |
type xref = Facts.ref * Token.src list |
44 |
type tagged_xref = xref * modifier list |
|
81366 | 45 |
|
82213 | 46 |
type result = {name: string, command: string, time: Time.time, state: Proof.state} |
82356 | 47 |
type proof_method = Time.time option -> tagged_xref list -> Proof.state -> result option |
48 |
type proof_method_options = {run_if_auto_try: bool} |
|
49 |
||
50 |
val noop_proof_method : proof_method = fn _ => fn _ => fn _ => NONE |
|
51 |
||
52 |
local |
|
53 |
val proof_methods = |
|
54 |
Synchronized.var "Try0.proof_methods" (Symtab.empty : proof_method Symtab.table); |
|
55 |
val auto_try_proof_methods_names = |
|
56 |
Synchronized.var "Try0.auto_try_proof_methods" (Symset.empty : Symset.T); |
|
57 |
in |
|
58 |
||
59 |
fun register_proof_method name ({run_if_auto_try} : proof_method_options) proof_method = |
|
60 |
let |
|
61 |
val () = if name = "" then error "Registering unnamed proof method" else () |
|
62 |
val () = Synchronized.change proof_methods (Symtab.update_new (name, proof_method)) |
|
63 |
val () = |
|
64 |
if run_if_auto_try then |
|
65 |
Synchronized.change auto_try_proof_methods_names (Symset.insert name) |
|
66 |
else |
|
67 |
() |
|
68 |
in () end |
|
69 |
||
82358 | 70 |
fun get_proof_method (name : string) : proof_method option = |
71 |
Symtab.lookup (Synchronized.value proof_methods) name; |
|
82356 | 72 |
|
82358 | 73 |
fun get_all_proof_methods () : (string * proof_method) list = |
82356 | 74 |
Symtab.fold (fn x => fn xs => x :: xs) (Synchronized.value proof_methods) [] |
75 |
||
82358 | 76 |
fun get_all_proof_method_names () : string list = |
82356 | 77 |
Symtab.fold (fn (name, _) => fn names => name :: names) (Synchronized.value proof_methods) [] |
78 |
||
79 |
fun get_all_auto_try_proof_method_names () : string list = |
|
80 |
Symset.dest (Synchronized.value auto_try_proof_methods_names) |
|
81 |
||
82 |
fun should_auto_try_proof_method (name : string) : bool = |
|
83 |
Symset.member (Synchronized.value auto_try_proof_methods_names) name |
|
84 |
||
85 |
end |
|
86 |
||
87 |
fun get_proof_method_or_noop name = |
|
88 |
(case get_proof_method name of |
|
89 |
NONE => noop_proof_method |
|
90 |
| SOME proof_method => proof_method) |
|
91 |
||
92 |
val apply_proof_method = get_proof_method_or_noop |
|
93 |
||
94 |
fun maybe_apply_proof_method name mode : proof_method = |
|
95 |
if mode <> Auto_Try orelse should_auto_try_proof_method name then |
|
96 |
get_proof_method_or_noop name |
|
97 |
else |
|
98 |
noop_proof_method |
|
99 |
||
100 |
fun time_string time = string_of_int (Time.toMilliseconds time) ^ " ms" |
|
101 |
fun tool_time_string (s, time) = s ^ ": " ^ time_string time |
|
102 |
||
103 |
fun generic_try0 mode timeout_opt (tagged : tagged_xref list) st = |
|
104 |
let |
|
105 |
fun try_method method = method mode timeout_opt tagged st |
|
106 |
fun get_message {command, time, ...} = "Found proof: " ^ Active.sendback_markup_command |
|
107 |
command ^ " (" ^ time_string time ^ ")" |
|
108 |
val print_step = Option.map (tap (writeln o get_message)) |
|
109 |
fun get_results methods : result list = |
|
110 |
if mode = Normal then |
|
111 |
methods |
|
112 |
|> Par_List.map (try_method #> print_step) |
|
113 |
|> map_filter I |
|
114 |
|> sort (Time.compare o apply2 #time) |
|
115 |
else |
|
116 |
methods |
|
117 |
|> Par_List.get_some try_method |
|
118 |
|> the_list |
|
119 |
val proof_method_names = get_all_proof_method_names () |
|
120 |
val maybe_apply_methods = map maybe_apply_proof_method proof_method_names |
|
121 |
in |
|
122 |
if mode = Normal then |
|
123 |
let val names = map quote proof_method_names in |
|
82360 | 124 |
writeln ("Trying " ^ implode_space (serial_commas "and" names) ^ "...") |
82356 | 125 |
end |
126 |
else |
|
127 |
(); |
|
128 |
(case get_results maybe_apply_methods of |
|
129 |
[] => (if mode = Normal then writeln "No proof found" else (); ((false, (noneN, [])), [])) |
|
130 |
| results as {name, command, ...} :: _ => |
|
131 |
let |
|
132 |
val method_times = |
|
133 |
results |
|
134 |
|> map (fn {name, time, ...} => (time, name)) |
|
135 |
|> AList.coalesce (op =) |
|
136 |
|> map (swap o apsnd commas) |
|
137 |
val message = |
|
138 |
(case mode of |
|
139 |
Auto_Try => "Auto Try0 found a proof" |
|
140 |
| Try => "Try0 found a proof" |
|
141 |
| Normal => "Try this") ^ ": " ^ |
|
142 |
Active.sendback_markup_command command ^ |
|
143 |
(case method_times of |
|
144 |
[(_, ms)] => " (" ^ time_string ms ^ ")" |
|
145 |
| method_times => "\n(" ^ space_implode "; " (map tool_time_string method_times) ^ ")") |
|
146 |
in |
|
147 |
((true, (name, if mode = Auto_Try then [message] else (writeln message; []))), results) |
|
148 |
end) |
|
149 |
end |
|
150 |
||
151 |
fun try0 timeout_opt = snd oo generic_try0 Normal timeout_opt |
|
152 |
||
153 |
fun try0_trans tagged = |
|
154 |
Toplevel.keep_proof |
|
155 |
(ignore o generic_try0 Normal (SOME default_timeout) tagged o Toplevel.proof_of) |
|
156 |
||
157 |
val parse_fact_refs = Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) Parse.thm) |
|
158 |
||
159 |
val parse_attr = |
|
160 |
Args.$$$ "simp" |-- Args.colon |-- parse_fact_refs >> (map (rpair [Simp])) |
|
161 |
|| Args.$$$ "intro" |-- Args.colon |-- parse_fact_refs >> (map (rpair [Intro])) |
|
162 |
|| Args.$$$ "elim" |-- Args.colon |-- parse_fact_refs >> (map (rpair [Elim])) |
|
163 |
|| Args.$$$ "dest" |-- Args.colon |-- parse_fact_refs >> (map (rpair [Dest])) |
|
164 |
||
165 |
fun parse_attrs x = |
|
166 |
(Args.parens parse_attrs |
|
167 |
|| Scan.repeat parse_attr >> (fn tagged => fold (curry (op @)) tagged [])) x |
|
168 |
||
169 |
val _ = |
|
170 |
Outer_Syntax.command \<^command_keyword>\<open>try0\<close> "try a combination of proof methods" |
|
171 |
(Scan.optional parse_attrs [] #>> try0_trans) |
|
172 |
||
82216 | 173 |
end |