author | blanchet |
Sun, 13 Nov 2011 20:28:22 +0100 | |
changeset 45478 | 8e299034eab4 |
parent 43085 | 0a2f5b86bdd7 |
child 46957 | 0c15caf47040 |
permissions | -rw-r--r-- |
36062 | 1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_util.ML |
35963 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
3 |
||
4 |
General-purpose functions used by the Sledgehammer modules. |
|
5 |
*) |
|
6 |
||
7 |
signature SLEDGEHAMMER_UTIL = |
|
8 |
sig |
|
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36062
diff
changeset
|
9 |
val plural_s : int -> string |
35963 | 10 |
val serial_commas : string -> string list -> string list |
38738
0ce517c1970f
make sure that "undo_ascii_of" is the inverse of "ascii_of", also for non-printable characters -- and avoid those in ``-style facts
blanchet
parents:
38698
diff
changeset
|
11 |
val simplify_spaces : string -> string |
35963 | 12 |
val parse_bool_option : bool -> string -> string -> bool option |
13 |
val parse_time_option : string -> string -> Time.time option |
|
38044 | 14 |
val subgoal_count : Proof.state -> int |
43043
1406f6fc5dc3
normalize indices in chained facts to make sure that backtick facts (which often result in different names) are recognized + changed definition of urgent messages
blanchet
parents:
43034
diff
changeset
|
15 |
val normalize_chained_theorems : thm list -> thm list |
38696
4c6b65d6a135
quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents:
38652
diff
changeset
|
16 |
val reserved_isar_keyword_table : unit -> unit Symtab.table |
35963 | 17 |
end; |
39318 | 18 |
|
35963 | 19 |
structure Sledgehammer_Util : SLEDGEHAMMER_UTIL = |
20 |
struct |
|
21 |
||
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43043
diff
changeset
|
22 |
open ATP_Util |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43043
diff
changeset
|
23 |
|
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36062
diff
changeset
|
24 |
fun plural_s n = if n = 1 then "" else "s" |
36062 | 25 |
|
43029
3e060b1c844b
use helpers and tweak Quickcheck's priority to it comes second (to give Solve Direct slightly more time before another prover runs)
blanchet
parents:
43005
diff
changeset
|
26 |
val serial_commas = Try.serial_commas |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43043
diff
changeset
|
27 |
val simplify_spaces = strip_spaces false (K true) |
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37575
diff
changeset
|
28 |
|
35963 | 29 |
fun parse_bool_option option name s = |
30 |
(case s of |
|
31 |
"smart" => if option then NONE else raise Option |
|
32 |
| "false" => SOME false |
|
33 |
| "true" => SOME true |
|
34 |
| "" => SOME true |
|
35 |
| _ => raise Option) |
|
36 |
handle Option.Option => |
|
37 |
let val ss = map quote ((option ? cons "smart") ["true", "false"]) in |
|
38 |
error ("Parameter " ^ quote name ^ " must be assigned " ^ |
|
39 |
space_implode " " (serial_commas "or" ss) ^ ".") |
|
40 |
end |
|
41 |
||
40341
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents:
39890
diff
changeset
|
42 |
val has_junk = |
40627
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents:
40341
diff
changeset
|
43 |
exists (fn s => not (Symbol.is_digit s) andalso s <> ".") o raw_explode (* FIXME Symbol.explode (?) *) |
40341
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents:
39890
diff
changeset
|
44 |
|
35963 | 45 |
fun parse_time_option _ "none" = NONE |
46 |
| parse_time_option name s = |
|
40341
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents:
39890
diff
changeset
|
47 |
let val secs = if has_junk s then NONE else Real.fromString s in |
43034
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
48 |
if is_none secs orelse Real.< (the secs, 0.0) then |
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
49 |
error ("Parameter " ^ quote name ^ " must be assigned a nonnegative \ |
40341
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents:
39890
diff
changeset
|
50 |
\number of seconds (e.g., \"60\", \"0.5\") or \"none\".") |
35963 | 51 |
else |
40341
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents:
39890
diff
changeset
|
52 |
SOME (seconds (the secs)) |
35963 | 53 |
end |
54 |
||
43029
3e060b1c844b
use helpers and tweak Quickcheck's priority to it comes second (to give Solve Direct slightly more time before another prover runs)
blanchet
parents:
43005
diff
changeset
|
55 |
val subgoal_count = Try.subgoal_count |
38044 | 56 |
|
43043
1406f6fc5dc3
normalize indices in chained facts to make sure that backtick facts (which often result in different names) are recognized + changed definition of urgent messages
blanchet
parents:
43034
diff
changeset
|
57 |
val normalize_chained_theorems = |
1406f6fc5dc3
normalize indices in chained facts to make sure that backtick facts (which often result in different names) are recognized + changed definition of urgent messages
blanchet
parents:
43034
diff
changeset
|
58 |
maps (fn th => insert Thm.eq_thm_prop (zero_var_indexes th) [th]) |
38696
4c6b65d6a135
quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents:
38652
diff
changeset
|
59 |
fun reserved_isar_keyword_table () = |
4c6b65d6a135
quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents:
38652
diff
changeset
|
60 |
union (op =) (Keyword.dest_keywords ()) (Keyword.dest_commands ()) |
4c6b65d6a135
quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents:
38652
diff
changeset
|
61 |
|> map (rpair ()) |> Symtab.make |
4c6b65d6a135
quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents:
38652
diff
changeset
|
62 |
|
35963 | 63 |
end; |