author | wenzelm |
Sat, 30 Nov 2024 16:01:58 +0100 | |
changeset 81513 | d11ed1bf0ad2 |
parent 81254 | d3c0734059ee |
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 |
51181
d0fa18638478
implement (more) accurate computation of parents
blanchet
parents:
51010
diff
changeset
|
11 |
val app_hd : ('a -> 'a) -> 'a list -> 'a list |
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36062
diff
changeset
|
12 |
val plural_s : int -> string |
35963 | 13 |
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
|
14 |
val simplify_spaces : string -> string |
35963 | 15 |
val parse_bool_option : bool -> string -> string -> bool option |
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54695
diff
changeset
|
16 |
val parse_time : string -> string -> Time.time |
38044 | 17 |
val subgoal_count : Proof.state -> int |
57306
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
18 |
val thms_in_proof : int -> (string Symtab.table * string Symtab.table) option -> thm -> |
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
19 |
string list option |
50267
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
20 |
val thms_of_name : Proof.context -> string -> thm list |
50557 | 21 |
val one_day : Time.time |
22 |
val one_year : Time.time |
|
53815
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
23 |
val hackish_string_of_term : Proof.context -> term -> string |
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
24 |
val spying : bool -> (unit -> Proof.state * int * string * string) -> unit |
35963 | 25 |
end; |
39318 | 26 |
|
35963 | 27 |
structure Sledgehammer_Util : SLEDGEHAMMER_UTIL = |
28 |
struct |
|
29 |
||
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43043
diff
changeset
|
30 |
open ATP_Util |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43043
diff
changeset
|
31 |
|
48392
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
blanchet
parents:
48383
diff
changeset
|
32 |
val sledgehammerN = "sledgehammer" |
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
blanchet
parents:
48383
diff
changeset
|
33 |
|
50608 | 34 |
val log10_2 = Math.log10 2.0 |
35 |
fun log2 n = Math.log10 n / log10_2 |
|
36 |
||
51181
d0fa18638478
implement (more) accurate computation of parents
blanchet
parents:
51010
diff
changeset
|
37 |
fun app_hd f (x :: xs) = f x :: xs |
d0fa18638478
implement (more) accurate computation of parents
blanchet
parents:
51010
diff
changeset
|
38 |
|
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36062
diff
changeset
|
39 |
fun plural_s n = if n = 1 then "" else "s" |
36062 | 40 |
|
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
|
41 |
val serial_commas = Try.serial_commas |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43043
diff
changeset
|
42 |
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
|
43 |
|
35963 | 44 |
fun parse_bool_option option name s = |
45 |
(case s of |
|
51930
52fd62618631
prefer explicitly qualified exceptions, which is particular important for robust handlers;
wenzelm
parents:
51181
diff
changeset
|
46 |
"smart" => if option then NONE else raise Option.Option |
35963 | 47 |
| "false" => SOME false |
48 |
| "true" => SOME true |
|
49 |
| "" => SOME true |
|
51930
52fd62618631
prefer explicitly qualified exceptions, which is particular important for robust handlers;
wenzelm
parents:
51181
diff
changeset
|
50 |
| _ => raise Option.Option) |
35963 | 51 |
handle Option.Option => |
55285 | 52 |
let val ss = map quote ((option ? cons "smart") ["true", "false"]) in |
53 |
error ("Parameter " ^ quote name ^ " must be assigned " ^ |
|
80910 | 54 |
implode_space (serial_commas "or" ss)) |
55285 | 55 |
end |
35963 | 56 |
|
40341
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents:
39890
diff
changeset
|
57 |
val has_junk = |
40627
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents:
40341
diff
changeset
|
58 |
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
|
59 |
|
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54695
diff
changeset
|
60 |
fun parse_time name s = |
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54695
diff
changeset
|
61 |
let val secs = if has_junk s then NONE else Real.fromString s in |
62826 | 62 |
if is_none secs orelse the secs < 0.0 then |
59431 | 63 |
error ("Parameter " ^ quote name ^ " must be assigned a nonnegative number of seconds \ |
63692 | 64 |
\(e.g., \"60\", \"0.5\") or \"none\"") |
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54695
diff
changeset
|
65 |
else |
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54695
diff
changeset
|
66 |
seconds (the secs) |
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54695
diff
changeset
|
67 |
end |
35963 | 68 |
|
74510 | 69 |
val subgoal_count = Logic.count_prems o Thm.prop_of o #goal o Proof.goal; |
38044 | 70 |
|
57306
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
71 |
exception TOO_MANY of unit |
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
72 |
|
57108
dc0b4f50e288
more generous max number of suggestions, for more safety
blanchet
parents:
57055
diff
changeset
|
73 |
(* FIXME: Similar yet different code in "mirabelle.ML". The code here has a few fixes that seem to |
dc0b4f50e288
more generous max number of suggestions, for more safety
blanchet
parents:
57055
diff
changeset
|
74 |
be missing over there; or maybe the two code portions are not doing the same? *) |
57306
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
75 |
fun fold_body_thm max_thms outer_name (map_plain_name, map_inclass_name) body = |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
46957
diff
changeset
|
76 |
let |
64573 | 77 |
fun app_thm map_name (_, thm_node) (accum as (num_thms, names)) = |
54116
ba709c934470
made theorem extraction code not delve too far when looking at local fact, by relying on 'strip_thm' instead of (wrongly) reverse-engineering proof terms
blanchet
parents:
54062
diff
changeset
|
78 |
let |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
78708
diff
changeset
|
79 |
val name = Thm_Name.short (Proofterm.thm_node_name thm_node) |
64573 | 80 |
val body = Proofterm.thm_node_body thm_node |
54116
ba709c934470
made theorem extraction code not delve too far when looking at local fact, by relying on 'strip_thm' instead of (wrongly) reverse-engineering proof terms
blanchet
parents:
54062
diff
changeset
|
81 |
val (anonymous, enter_class) = |
ba709c934470
made theorem extraction code not delve too far when looking at local fact, by relying on 'strip_thm' instead of (wrongly) reverse-engineering proof terms
blanchet
parents:
54062
diff
changeset
|
82 |
(* The "name = outer_name" case caters for the uncommon case where the proved theorem |
ba709c934470
made theorem extraction code not delve too far when looking at local fact, by relying on 'strip_thm' instead of (wrongly) reverse-engineering proof terms
blanchet
parents:
54062
diff
changeset
|
83 |
occurs in its own proof (e.g., "Transitive_Closure.trancl_into_trancl"). *) |
ba709c934470
made theorem extraction code not delve too far when looking at local fact, by relying on 'strip_thm' instead of (wrongly) reverse-engineering proof terms
blanchet
parents:
54062
diff
changeset
|
84 |
if name = "" orelse name = outer_name then (true, false) |
ba709c934470
made theorem extraction code not delve too far when looking at local fact, by relying on 'strip_thm' instead of (wrongly) reverse-engineering proof terms
blanchet
parents:
54062
diff
changeset
|
85 |
else if map_inclass_name name = SOME outer_name then (true, true) |
ba709c934470
made theorem extraction code not delve too far when looking at local fact, by relying on 'strip_thm' instead of (wrongly) reverse-engineering proof terms
blanchet
parents:
54062
diff
changeset
|
86 |
else (false, false) |
ba709c934470
made theorem extraction code not delve too far when looking at local fact, by relying on 'strip_thm' instead of (wrongly) reverse-engineering proof terms
blanchet
parents:
54062
diff
changeset
|
87 |
in |
ba709c934470
made theorem extraction code not delve too far when looking at local fact, by relying on 'strip_thm' instead of (wrongly) reverse-engineering proof terms
blanchet
parents:
54062
diff
changeset
|
88 |
if anonymous then |
57756
92fe49c7ab3b
peek instead of joining -- is perhaps less risky
blanchet
parents:
57727
diff
changeset
|
89 |
(case Future.peek body of |
92fe49c7ab3b
peek instead of joining -- is perhaps less risky
blanchet
parents:
57727
diff
changeset
|
90 |
SOME (Exn.Res the_body) => |
92fe49c7ab3b
peek instead of joining -- is perhaps less risky
blanchet
parents:
57727
diff
changeset
|
91 |
app_body (if enter_class then map_inclass_name else map_name) the_body accum |
92fe49c7ab3b
peek instead of joining -- is perhaps less risky
blanchet
parents:
57727
diff
changeset
|
92 |
| NONE => accum) |
54116
ba709c934470
made theorem extraction code not delve too far when looking at local fact, by relying on 'strip_thm' instead of (wrongly) reverse-engineering proof terms
blanchet
parents:
54062
diff
changeset
|
93 |
else |
57306
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
94 |
(case map_name name of |
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
95 |
SOME name' => |
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
96 |
if member (op =) names name' then accum |
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
97 |
else if num_thms = max_thms then raise TOO_MANY () |
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
98 |
else (num_thms + 1, name' :: names) |
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
99 |
| NONE => accum) |
54116
ba709c934470
made theorem extraction code not delve too far when looking at local fact, by relying on 'strip_thm' instead of (wrongly) reverse-engineering proof terms
blanchet
parents:
54062
diff
changeset
|
100 |
end |
77866 | 101 |
and app_body map_name (PBody {thms, ...}) = fold (app_thm map_name) thms |
57108
dc0b4f50e288
more generous max number of suggestions, for more safety
blanchet
parents:
57055
diff
changeset
|
102 |
in |
57306
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
103 |
snd (app_body map_plain_name body (0, [])) |
57108
dc0b4f50e288
more generous max number of suggestions, for more safety
blanchet
parents:
57055
diff
changeset
|
104 |
end |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
46957
diff
changeset
|
105 |
|
57306
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents:
57108
diff
changeset
|
106 |
fun thms_in_proof max_thms name_tabs th = |
57838
c21f2c52f54b
careful when calling 'Thm.proof_body_of' -- it can throw exceptions
blanchet
parents:
57306
diff
changeset
|
107 |
(case try Thm.proof_body_of th of |
c21f2c52f54b
careful when calling 'Thm.proof_body_of' -- it can throw exceptions
blanchet
parents:
57306
diff
changeset
|
108 |
NONE => NONE |
c21f2c52f54b
careful when calling 'Thm.proof_body_of' -- it can throw exceptions
blanchet
parents:
57306
diff
changeset
|
109 |
| SOME body => |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58928
diff
changeset
|
110 |
let val map_names = (case name_tabs of SOME p => apply2 Symtab.lookup p | NONE => `I SOME) in |
80306 | 111 |
SOME (fold_body_thm max_thms (Thm_Name.short (Thm.get_name_hint th)) map_names |
112 |
(Proofterm.strip_thm_body body)) |
|
57838
c21f2c52f54b
careful when calling 'Thm.proof_body_of' -- it can throw exceptions
blanchet
parents:
57306
diff
changeset
|
113 |
handle TOO_MANY () => NONE |
c21f2c52f54b
careful when calling 'Thm.proof_body_of' -- it can throw exceptions
blanchet
parents:
57306
diff
changeset
|
114 |
end) |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
46957
diff
changeset
|
115 |
|
50267
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
116 |
fun thms_of_name ctxt name = |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
117 |
let |
58928
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents:
58919
diff
changeset
|
118 |
val keywords = Thy_Header.get_keywords' ctxt |
50267
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
119 |
in |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80910
diff
changeset
|
120 |
Token.explode keywords Position.start name |
67498
88a02f41246a
tuned: prefer list operations over Source.source;
wenzelm
parents:
66020
diff
changeset
|
121 |
|> filter Token.is_proper |
63936 | 122 |
|> Source.of_list |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80910
diff
changeset
|
123 |
|> Source.source' (Context.Proof ctxt) Token.stopper Attrib.multi_thm |
50267
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
124 |
|> Source.exhaust |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
125 |
end |
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents:
50049
diff
changeset
|
126 |
|
50557 | 127 |
val one_day = seconds (24.0 * 60.0 * 60.0) |
128 |
val one_year = seconds (365.0 * 24.0 * 60.0 * 60.0) |
|
129 |
||
53815
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
130 |
fun hackish_string_of_term ctxt = |
61432 | 131 |
(* FIXME: map_aterms (fn Free (s, T) => Free (if Name.is_internal s then "_" else s, T) | t => t) |
80866
8c67b14fdd48
clarified internal tool output: prefer Pretty.pure_string_of over manipulation of print_mode;
wenzelm
parents:
80820
diff
changeset
|
132 |
#> *) Syntax.pretty_term ctxt |
8c67b14fdd48
clarified internal tool output: prefer Pretty.pure_string_of over manipulation of print_mode;
wenzelm
parents:
80820
diff
changeset
|
133 |
#> Pretty.pure_string_of |
61330
20af2ad9261e
further improved fine point w.r.t. replaying in the presence of chained facts and a non-empty meta-quantifier prefix + avoid printing internal names in backquotes
blanchet
parents:
59582
diff
changeset
|
134 |
#> simplify_spaces |
53815
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
135 |
|
57727
02a53c1bbb6c
generate backquotes without markup, since this confuses preplay; bump up spying version identifier;
blanchet
parents:
57306
diff
changeset
|
136 |
val spying_version = "d" |
53815
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
137 |
|
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
138 |
fun spying false _ = () |
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
139 |
| spying true f = |
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
140 |
let |
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
141 |
val (state, i, tool, message) = f () |
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
142 |
val ctxt = Proof.context_of state |
59582 | 143 |
val goal = Logic.get_goal (Thm.prop_of (#goal (Proof.goal state))) i |
53815
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
144 |
val hash = String.substring (SHA1.rep (SHA1.digest (hackish_string_of_term ctxt goal)), 0, 12) |
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
145 |
in |
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
146 |
File.append (Path.explode "$ISABELLE_HOME_USER/spy_sledgehammer") |
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
147 |
(spying_version ^ " " ^ timestamp () ^ ": " ^ hash ^ ": " ^ tool ^ ": " ^ message ^ "\n") |
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
148 |
end |
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
blanchet
parents:
53800
diff
changeset
|
149 |
|
35963 | 150 |
end; |