author | haftmann |
Thu, 19 Jun 2025 17:15:40 +0200 | |
changeset 82734 | 89347c0cc6a3 |
parent 82455 | eaf0b4673ab2 |
permissions | -rw-r--r-- |
82364 | 1 |
(* Title: HOL/Tools/try0.ML |
2 |
Author: Jasmin Blanchette, LMU Muenchen |
|
3 |
Author: Martin Desharnais, LMU Muenchen |
|
4 |
Author: Fabian Huch, TU Muenchen |
|
38942
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 |
Try a combination of proof methods. |
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
7 |
*) |
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
8 |
|
46641 | 9 |
signature TRY0 = |
38942
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
10 |
sig |
82360 | 11 |
val serial_commas : string -> string list -> string list |
12 |
||
82355 | 13 |
type xref = Facts.ref * Token.src list |
82364 | 14 |
|
15 |
type facts = |
|
16 |
{usings: xref list, |
|
17 |
simps : xref list, |
|
18 |
intros : xref list, |
|
19 |
elims : xref list, |
|
20 |
dests : xref list} |
|
21 |
val empty_facts : facts |
|
22 |
||
82213 | 23 |
type result = {name: string, command: string, time: Time.time, state: Proof.state} |
82364 | 24 |
type proof_method = Time.time option -> facts -> Proof.state -> result option |
82358 | 25 |
type proof_method_options = {run_if_auto_try: bool} |
26 |
||
27 |
val register_proof_method : string -> proof_method_options -> proof_method -> unit |
|
28 |
val get_proof_method : string -> proof_method option |
|
82369
eae75676ecd7
refactored tactic-based provers in sledgehammer to use Try0 module
desharna
parents:
82366
diff
changeset
|
29 |
val get_proof_method_or_noop : string -> proof_method |
82363
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82362
diff
changeset
|
30 |
val get_all_proof_method_names : unit -> string list |
82358 | 31 |
|
82396 | 32 |
val schedule : string Config.T |
82455 | 33 |
val get_schedule : Proof.context -> string list list |
82396 | 34 |
|
82360 | 35 |
datatype mode = Auto_Try | Try | Normal |
82364 | 36 |
val generic_try0 : mode -> Time.time option -> facts -> Proof.state -> |
82360 | 37 |
(bool * (string * string list)) * result list |
82364 | 38 |
val try0 : Time.time option -> facts -> Proof.state -> result list |
81365 | 39 |
end |
38942
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
40 |
|
46641 | 41 |
structure Try0 : TRY0 = |
38942
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
42 |
struct |
e10c11971fa7
"try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff
changeset
|
43 |
|
82360 | 44 |
fun serial_commas _ [] = ["??"] |
45 |
| serial_commas _ [s] = [s] |
|
46 |
| serial_commas conj [s1, s2] = [s1, conj, s2] |
|
47 |
| serial_commas conj [s1, s2, s3] = [s1 ^ ",", s2 ^ ",", conj, s3] |
|
48 |
| serial_commas conj (s :: ss) = s ^ "," :: serial_commas conj ss; |
|
49 |
||
81365 | 50 |
val noneN = "none" |
43026
0f15575a6465
handle non-auto try cases gracefully in Try Methods
blanchet
parents:
43024
diff
changeset
|
51 |
|
81365 | 52 |
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
|
53 |
|
81365 | 54 |
val default_timeout = seconds 5.0 |
38944 | 55 |
|
81369
0677712016b5
try0: add 'use' modifier for thms to insert;
Fabian Huch <huch@in.tum.de>
parents:
81368
diff
changeset
|
56 |
datatype modifier = Use | Simp | Intro | Elim | Dest |
82355 | 57 |
type xref = Facts.ref * Token.src list |
58 |
type tagged_xref = xref * modifier list |
|
82364 | 59 |
type facts = |
60 |
{usings: xref list, |
|
61 |
simps : xref list, |
|
62 |
intros : xref list, |
|
63 |
elims : xref list, |
|
64 |
dests : xref list} |
|
65 |
||
66 |
val empty_facts: facts = {usings = [], simps = [], intros = [], elims = [], dests = []} |
|
67 |
fun union_facts (left : facts) (right : facts) : facts = |
|
68 |
{usings = #usings left @ #usings right, |
|
69 |
simps = #simps left @ #simps right, |
|
70 |
intros = #intros left @ #intros right, |
|
71 |
elims = #elims left @ #elims right, |
|
72 |
dests = #dests left @ #dests right} |
|
81366 | 73 |
|
82213 | 74 |
type result = {name: string, command: string, time: Time.time, state: Proof.state} |
82364 | 75 |
type proof_method = Time.time option -> facts -> Proof.state -> result option |
82356 | 76 |
type proof_method_options = {run_if_auto_try: bool} |
77 |
||
78 |
val noop_proof_method : proof_method = fn _ => fn _ => fn _ => NONE |
|
79 |
||
80 |
local |
|
81 |
val proof_methods = |
|
82 |
Synchronized.var "Try0.proof_methods" (Symtab.empty : proof_method Symtab.table); |
|
83 |
val auto_try_proof_methods_names = |
|
84 |
Synchronized.var "Try0.auto_try_proof_methods" (Symset.empty : Symset.T); |
|
85 |
in |
|
86 |
||
87 |
fun register_proof_method name ({run_if_auto_try} : proof_method_options) proof_method = |
|
88 |
let |
|
89 |
val () = if name = "" then error "Registering unnamed proof method" else () |
|
82364 | 90 |
val () = Synchronized.change proof_methods (Symtab.update (name, proof_method)) |
82356 | 91 |
val () = |
92 |
if run_if_auto_try then |
|
93 |
Synchronized.change auto_try_proof_methods_names (Symset.insert name) |
|
94 |
else |
|
95 |
() |
|
96 |
in () end |
|
97 |
||
82358 | 98 |
fun get_proof_method (name : string) : proof_method option = |
99 |
Symtab.lookup (Synchronized.value proof_methods) name; |
|
82356 | 100 |
|
82358 | 101 |
fun get_all_proof_method_names () : string list = |
82356 | 102 |
Symtab.fold (fn (name, _) => fn names => name :: names) (Synchronized.value proof_methods) [] |
103 |
||
104 |
fun should_auto_try_proof_method (name : string) : bool = |
|
105 |
Symset.member (Synchronized.value auto_try_proof_methods_names) name |
|
106 |
||
107 |
end |
|
108 |
||
109 |
fun get_proof_method_or_noop name = |
|
110 |
(case get_proof_method name of |
|
82397 | 111 |
NONE => (warning ("Proof method \"" ^ name ^ "\" is undefined"); noop_proof_method) |
82356 | 112 |
| SOME proof_method => proof_method) |
113 |
||
114 |
fun maybe_apply_proof_method name mode : proof_method = |
|
115 |
if mode <> Auto_Try orelse should_auto_try_proof_method name then |
|
116 |
get_proof_method_or_noop name |
|
117 |
else |
|
118 |
noop_proof_method |
|
119 |
||
82396 | 120 |
val schedule = Attrib.setup_config_string \<^binding>\<open>try0_schedule\<close> (K "") |
121 |
||
82455 | 122 |
fun get_schedule (ctxt : Proof.context) : string list list = |
123 |
let |
|
124 |
fun some_nonempty_string sub = |
|
125 |
if Substring.isEmpty sub then |
|
126 |
NONE |
|
127 |
else |
|
128 |
SOME (Substring.string sub) |
|
129 |
in |
|
130 |
Config.get ctxt schedule |
|
131 |
|> Substring.full |
|
132 |
|> Substring.tokens (fn c => c = #"|") |
|
133 |
|> map (map_filter some_nonempty_string o Substring.tokens Char.isSpace) |
|
134 |
end |
|
135 |
||
82396 | 136 |
local |
137 |
||
82356 | 138 |
fun time_string time = string_of_int (Time.toMilliseconds time) ^ " ms" |
139 |
fun tool_time_string (s, time) = s ^ ": " ^ time_string time |
|
140 |
||
82396 | 141 |
fun generic_try0_step mode (timeout_opt : Time.time option) (facts : facts) (st : Proof.state) |
142 |
(proof_methods : string list) = |
|
82356 | 143 |
let |
82364 | 144 |
fun try_method (method : mode -> proof_method) = method mode timeout_opt facts st |
82356 | 145 |
fun get_message {command, time, ...} = "Found proof: " ^ Active.sendback_markup_command |
146 |
command ^ " (" ^ time_string time ^ ")" |
|
147 |
val print_step = Option.map (tap (writeln o get_message)) |
|
148 |
fun get_results methods : result list = |
|
149 |
if mode = Normal then |
|
150 |
methods |
|
151 |
|> Par_List.map (try_method #> print_step) |
|
152 |
|> map_filter I |
|
153 |
|> sort (Time.compare o apply2 #time) |
|
154 |
else |
|
155 |
methods |
|
156 |
|> Par_List.get_some try_method |
|
157 |
|> the_list |
|
82396 | 158 |
val maybe_apply_methods = map maybe_apply_proof_method proof_methods |
82356 | 159 |
in |
160 |
if mode = Normal then |
|
82396 | 161 |
let val names = map quote proof_methods in |
82360 | 162 |
writeln ("Trying " ^ implode_space (serial_commas "and" names) ^ "...") |
82356 | 163 |
end |
164 |
else |
|
165 |
(); |
|
166 |
(case get_results maybe_apply_methods of |
|
167 |
[] => (if mode = Normal then writeln "No proof found" else (); ((false, (noneN, [])), [])) |
|
168 |
| results as {name, command, ...} :: _ => |
|
169 |
let |
|
170 |
val method_times = |
|
171 |
results |
|
172 |
|> map (fn {name, time, ...} => (time, name)) |
|
173 |
|> AList.coalesce (op =) |
|
174 |
|> map (swap o apsnd commas) |
|
175 |
val message = |
|
176 |
(case mode of |
|
177 |
Auto_Try => "Auto Try0 found a proof" |
|
178 |
| Try => "Try0 found a proof" |
|
179 |
| Normal => "Try this") ^ ": " ^ |
|
180 |
Active.sendback_markup_command command ^ |
|
181 |
(case method_times of |
|
182 |
[(_, ms)] => " (" ^ time_string ms ^ ")" |
|
183 |
| method_times => "\n(" ^ space_implode "; " (map tool_time_string method_times) ^ ")") |
|
184 |
in |
|
185 |
((true, (name, if mode = Auto_Try then [message] else (writeln message; []))), results) |
|
186 |
end) |
|
187 |
end |
|
188 |
||
82396 | 189 |
in |
190 |
||
191 |
fun generic_try0 mode (timeout_opt : Time.time option) (facts : facts) (st : Proof.state) = |
|
192 |
let |
|
82455 | 193 |
val schedule = get_schedule (Proof.context_of st) |
82396 | 194 |
fun iterate [] = ((false, (noneN, [])), []) |
195 |
| iterate (proof_methods :: proof_methodss) = |
|
196 |
(case generic_try0_step mode timeout_opt facts st proof_methods of |
|
197 |
(_, []) => iterate proof_methodss |
|
198 |
| result as (_, _ :: _) => result) |
|
199 |
in |
|
200 |
iterate (if null schedule then [get_all_proof_method_names ()] else schedule) |
|
201 |
end; |
|
202 |
||
203 |
end |
|
204 |
||
82356 | 205 |
fun try0 timeout_opt = snd oo generic_try0 Normal timeout_opt |
206 |
||
82364 | 207 |
fun try0_trans (facts : facts) = |
82366 | 208 |
Toplevel.keep_proof (ignore o try0 (SOME default_timeout) facts o Toplevel.proof_of) |
82356 | 209 |
|
210 |
val parse_fact_refs = Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) Parse.thm) |
|
211 |
||
212 |
val parse_attr = |
|
82364 | 213 |
Args.$$$ "simp" |-- Args.colon |-- parse_fact_refs >> (fn xrefs => |
214 |
{usings = [], simps = xrefs, intros = [], elims = [], dests = []}) |
|
215 |
|| Args.$$$ "intro" |-- Args.colon |-- parse_fact_refs >> (fn xrefs => |
|
216 |
{usings = [], simps = [], intros = xrefs, elims = [], dests = []}) |
|
217 |
|| Args.$$$ "elim" |-- Args.colon |-- parse_fact_refs >> (fn xrefs => |
|
218 |
{usings = [], simps = [], intros = [], elims = xrefs, dests = []}) |
|
219 |
|| Args.$$$ "dest" |-- Args.colon |-- parse_fact_refs >> (fn xrefs => |
|
220 |
{usings = [], simps = [], intros = [], elims = [], dests = xrefs}) |
|
82356 | 221 |
|
222 |
fun parse_attrs x = |
|
223 |
(Args.parens parse_attrs |
|
82364 | 224 |
|| Scan.repeat parse_attr >> (fn factss => fold union_facts factss empty_facts)) x |
82356 | 225 |
|
226 |
val _ = |
|
227 |
Outer_Syntax.command \<^command_keyword>\<open>try0\<close> "try a combination of proof methods" |
|
82364 | 228 |
(Scan.optional parse_attrs empty_facts #>> try0_trans) |
82356 | 229 |
|
82216 | 230 |
end |