author | blanchet |
Sun, 13 Jan 2013 20:57:48 +0100 | |
changeset 50861 | fa4253914e98 |
parent 50735 | 6b232d76cbc9 |
child 51007 | 4f694d52bf62 |
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 |
|
48392
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
blanchet
parents:
48383
diff
changeset
|
9 |
val sledgehammerN : string |
50608 | 10 |
val log2 : real -> real |
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36062
diff
changeset
|
11 |
val plural_s : int -> string |
35963 | 12 |
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
|
13 |
val simplify_spaces : string -> string |
48656
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48392
diff
changeset
|
14 |
val with_cleanup : ('a -> unit) -> ('a -> 'b) -> 'a -> 'b |
48383 | 15 |
val infinite_timeout : Time.time |
48318 | 16 |
val time_mult : real -> Time.time -> Time.time |
35963 | 17 |
val parse_bool_option : bool -> string -> string -> bool option |
18 |
val parse_time_option : string -> string -> Time.time option |
|
38044 | 19 |
val subgoal_count : Proof.state -> int |
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
|
20 |
val reserved_isar_keyword_table : unit -> unit Symtab.table |
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
21 |
val thms_in_proof : |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
22 |
(string Symtab.table * string Symtab.table) option -> thm -> string list |
50267
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
23 |
val thms_of_name : Proof.context -> string -> thm list |
50557 | 24 |
val one_day : Time.time |
25 |
val one_year : Time.time |
|
26 |
val time_limit : Time.time option -> ('a -> 'b) -> 'a -> 'b |
|
50048 | 27 |
val with_vanilla_print_mode : ('a -> 'b) -> 'a -> 'b |
35963 | 28 |
end; |
39318 | 29 |
|
35963 | 30 |
structure Sledgehammer_Util : SLEDGEHAMMER_UTIL = |
31 |
struct |
|
32 |
||
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43043
diff
changeset
|
33 |
open ATP_Util |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43043
diff
changeset
|
34 |
|
48392
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
blanchet
parents:
48383
diff
changeset
|
35 |
val sledgehammerN = "sledgehammer" |
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
blanchet
parents:
48383
diff
changeset
|
36 |
|
50608 | 37 |
val log10_2 = Math.log10 2.0 |
38 |
||
39 |
fun log2 n = Math.log10 n / log10_2 |
|
40 |
||
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36062
diff
changeset
|
41 |
fun plural_s n = if n = 1 then "" else "s" |
36062 | 42 |
|
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
|
43 |
val serial_commas = Try.serial_commas |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43043
diff
changeset
|
44 |
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
|
45 |
|
48656
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48392
diff
changeset
|
46 |
fun with_cleanup clean_up f x = |
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48392
diff
changeset
|
47 |
Exn.capture f x |
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48392
diff
changeset
|
48 |
|> tap (fn _ => clean_up x) |
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48392
diff
changeset
|
49 |
|> Exn.release |
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48392
diff
changeset
|
50 |
|
48383 | 51 |
val infinite_timeout = seconds 31536000.0 (* one year *) |
52 |
||
48318 | 53 |
fun time_mult k t = |
54 |
Time.fromMilliseconds (Real.ceil (k * Real.fromInt (Time.toMilliseconds t))) |
|
55 |
||
35963 | 56 |
fun parse_bool_option option name s = |
57 |
(case s of |
|
58 |
"smart" => if option then NONE else raise Option |
|
59 |
| "false" => SOME false |
|
60 |
| "true" => SOME true |
|
61 |
| "" => SOME true |
|
62 |
| _ => raise Option) |
|
63 |
handle Option.Option => |
|
64 |
let val ss = map quote ((option ? cons "smart") ["true", "false"]) in |
|
65 |
error ("Parameter " ^ quote name ^ " must be assigned " ^ |
|
66 |
space_implode " " (serial_commas "or" ss) ^ ".") |
|
67 |
end |
|
68 |
||
40341
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents:
39890
diff
changeset
|
69 |
val has_junk = |
40627
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents:
40341
diff
changeset
|
70 |
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
|
71 |
|
35963 | 72 |
fun parse_time_option _ "none" = NONE |
73 |
| parse_time_option name s = |
|
40341
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents:
39890
diff
changeset
|
74 |
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
|
75 |
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
|
76 |
error ("Parameter " ^ quote name ^ " must be assigned a nonnegative \ |
40341
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents:
39890
diff
changeset
|
77 |
\number of seconds (e.g., \"60\", \"0.5\") or \"none\".") |
35963 | 78 |
else |
40341
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents:
39890
diff
changeset
|
79 |
SOME (seconds (the secs)) |
35963 | 80 |
end |
81 |
||
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
|
82 |
val subgoal_count = Try.subgoal_count |
38044 | 83 |
|
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
|
84 |
fun reserved_isar_keyword_table () = |
48292 | 85 |
Keyword.dest () |-> union (op =) |> map (rpair ()) |> Symtab.make |
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
|
86 |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
46957
diff
changeset
|
87 |
(* FIXME: Similar yet different code in "mirabelle.ML". The code here has a few |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
46957
diff
changeset
|
88 |
fixes that seem to be missing over there; or maybe the two code portions are |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
46957
diff
changeset
|
89 |
not doing the same? *) |
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
90 |
fun fold_body_thms outer_name (map_plain_name, map_inclass_name) = |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
46957
diff
changeset
|
91 |
let |
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
92 |
fun app map_name n (PBody {thms, ...}) = |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48292
diff
changeset
|
93 |
thms |> fold (fn (_, (name, _, body)) => fn accum => |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48292
diff
changeset
|
94 |
let |
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
95 |
val collect = union (op =) o the_list o map_name |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
96 |
(* The "name = outer_name" case caters for the uncommon case where |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
97 |
the proved theorem occurs in its own proof (e.g., |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
98 |
"Transitive_Closure.trancl_into_trancl"). *) |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
99 |
val (anonymous, enter_class) = |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
100 |
if name = "" orelse (n = 1 andalso name = outer_name) then |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
101 |
(true, false) |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
102 |
else if n = 1 andalso map_inclass_name name = SOME outer_name then |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
103 |
(true, true) |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
104 |
else |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
105 |
(false, false) |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48292
diff
changeset
|
106 |
val accum = |
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
107 |
accum |> (if n = 1 andalso not anonymous then collect name else I) |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
108 |
val n = n + (if anonymous then 0 else 1) |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
109 |
in |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
110 |
accum |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
111 |
|> (if n <= 1 then |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
112 |
app (if enter_class then map_inclass_name else map_name) n |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
113 |
(Future.join body) |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
114 |
else |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
115 |
I) |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
116 |
end) |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
117 |
in fold (app map_plain_name 0) end |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
46957
diff
changeset
|
118 |
|
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
119 |
fun thms_in_proof name_tabs th = |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
46957
diff
changeset
|
120 |
let |
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
121 |
val map_names = |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
122 |
case name_tabs of |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
123 |
SOME p => pairself Symtab.lookup p |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
124 |
| NONE => `I SOME |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
46957
diff
changeset
|
125 |
val names = |
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
126 |
fold_body_thms (Thm.get_name_hint th) map_names [Thm.proof_body_of th] [] |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
46957
diff
changeset
|
127 |
in names end |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
46957
diff
changeset
|
128 |
|
50267
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
129 |
fun thms_of_name ctxt name = |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
130 |
let |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
131 |
val lex = Keyword.get_lexicons |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
132 |
val get = maps (Proof_Context.get_fact ctxt o fst) |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
133 |
in |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
134 |
Source.of_string name |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
135 |
|> Symbol.source |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
136 |
|> Token.source {do_recover = SOME false} lex Position.start |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
137 |
|> Token.source_proper |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
138 |
|> Source.source Token.stopper (Parse_Spec.xthms1 >> get) NONE |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
139 |
|> Source.exhaust |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
140 |
end |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
141 |
|
50557 | 142 |
val one_day = seconds (24.0 * 60.0 * 60.0) |
143 |
val one_year = seconds (365.0 * 24.0 * 60.0 * 60.0) |
|
144 |
||
145 |
fun time_limit NONE = I |
|
146 |
| time_limit (SOME delay) = TimeLimit.timeLimit delay |
|
147 |
||
50048 | 148 |
fun with_vanilla_print_mode f x = |
149 |
Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN) |
|
150 |
(print_mode_value ())) f x |
|
151 |
||
35963 | 152 |
end; |