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