| author | wenzelm | 
| Sun, 22 Aug 2010 16:43:20 +0200 | |
| changeset 38576 | ce3eed2b16f7 | 
| parent 36787 | f60e4dd6d76f | 
| child 39232 | 69c6d3e87660 | 
| permissions | -rw-r--r-- | 
| 32564 | 1 | (* Title: HOL/Mirabelle/Tools/mirabelle.ML | 
| 2 | Author: Jasmin Blanchette and Sascha Boehme, TU Munich | |
| 32381 | 3 | *) | 
| 4 | ||
| 5 | signature MIRABELLE = | |
| 6 | sig | |
| 32564 | 7 | (*configuration*) | 
| 32396 | 8 | val logfile : string Config.T | 
| 32381 | 9 | val timeout : int Config.T | 
| 32382 
98674ac811c4
Mirabelle tool script conforming to standard Isabelle tool interface,
 boehmes parents: 
32381diff
changeset | 10 | val start_line : int Config.T | 
| 
98674ac811c4
Mirabelle tool script conforming to standard Isabelle tool interface,
 boehmes parents: 
32381diff
changeset | 11 | val end_line : int Config.T | 
| 32495 | 12 | val setup : theory -> theory | 
| 32381 | 13 | |
| 32564 | 14 | (*core*) | 
| 32567 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 15 | type init_action = int -> theory -> theory | 
| 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 16 |   type done_args = {last: Toplevel.state, log: string -> unit}
 | 
| 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 17 | type done_action = int -> done_args -> unit | 
| 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 18 |   type run_args = {pre: Proof.state, post: Toplevel.state option,
 | 
| 32676 | 19 | timeout: Time.time, log: string -> unit, pos: Position.T, name: string} | 
| 32567 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 20 | type run_action = int -> run_args -> unit | 
| 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 21 | type action = init_action * run_action * done_action | 
| 32521 | 22 | val catch : (int -> string) -> run_action -> run_action | 
| 34052 | 23 | val catch_result : (int -> string) -> 'a -> (int -> run_args -> 'a) -> | 
| 24 | int -> run_args -> 'a | |
| 32515 
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
 boehmes parents: 
32504diff
changeset | 25 | val register : action -> theory -> theory | 
| 32495 | 26 | val step_hook : Toplevel.transition -> Toplevel.state -> Toplevel.state -> | 
| 27 | unit | |
| 28 | ||
| 32564 | 29 | (*utility functions*) | 
| 32469 | 30 | val can_apply : Time.time -> (Proof.context -> int -> tactic) -> | 
| 31 | Proof.state -> bool | |
| 33243 | 32 | val theorems_in_proof_term : thm -> thm list | 
| 33 | val theorems_of_sucessful_proof : Toplevel.state option -> thm list | |
| 32385 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: 
32383diff
changeset | 34 | val get_setting : (string * string) list -> string * string -> string | 
| 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: 
32383diff
changeset | 35 | val get_int_setting : (string * string) list -> string * int -> int | 
| 32498 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 36 |   val cpu_time : ('a -> 'b) -> 'a -> 'b * int
 | 
| 32381 | 37 | end | 
| 38 | ||
| 39 | ||
| 40 | ||
| 32497 | 41 | structure Mirabelle : MIRABELLE = | 
| 32495 | 42 | struct | 
| 43 | ||
| 44 | (* Mirabelle configuration *) | |
| 45 | ||
| 36001 | 46 | val (logfile, setup1) = Attrib.config_string "mirabelle_logfile" (K "") | 
| 47 | val (timeout, setup2) = Attrib.config_int "mirabelle_timeout" (K 30) | |
| 48 | val (start_line, setup3) = Attrib.config_int "mirabelle_start_line" (K 0) | |
| 49 | val (end_line, setup4) = Attrib.config_int "mirabelle_end_line" (K ~1) | |
| 32495 | 50 | |
| 51 | val setup = setup1 #> setup2 #> setup3 #> setup4 | |
| 32385 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: 
32383diff
changeset | 52 | |
| 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: 
32383diff
changeset | 53 | |
| 32381 | 54 | (* Mirabelle core *) | 
| 55 | ||
| 32521 | 56 | type init_action = int -> theory -> theory | 
| 32567 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 57 | type done_args = {last: Toplevel.state, log: string -> unit}
 | 
| 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 58 | type done_action = int -> done_args -> unit | 
| 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 59 | type run_args = {pre: Proof.state, post: Toplevel.state option,
 | 
| 32676 | 60 | timeout: Time.time, log: string -> unit, pos: Position.T, name: string} | 
| 32567 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 61 | type run_action = int -> run_args -> unit | 
| 32521 | 62 | type action = init_action * run_action * done_action | 
| 32381 | 63 | |
| 33522 | 64 | structure Actions = Theory_Data | 
| 32381 | 65 | ( | 
| 32521 | 66 | type T = (int * run_action * done_action) list | 
| 32515 
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
 boehmes parents: 
32504diff
changeset | 67 | val empty = [] | 
| 32381 | 68 | val extend = I | 
| 33522 | 69 | fun merge data = Library.merge (K true) data (* FIXME ?!? *) | 
| 32381 | 70 | ) | 
| 71 | ||
| 32515 
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
 boehmes parents: 
32504diff
changeset | 72 | |
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33522diff
changeset | 73 | fun log_exn log tag id e = log (tag id ^ "exception:\n" ^ General.exnMessage e) | 
| 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33522diff
changeset | 74 | |
| 34052 | 75 | fun catch tag f id (st as {log, ...}: run_args) = (f id st; ())
 | 
| 76 | handle (exn as Exn.Interrupt) => reraise exn | |
| 77 | | exn => (log_exn log tag id exn; ()) | |
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33522diff
changeset | 78 | |
| 34052 | 79 | fun catch_result tag d f id (st as {log, ...}: run_args) = f id st
 | 
| 80 | handle (exn as Exn.Interrupt) => reraise exn | |
| 81 | | exn => (log_exn log tag id exn; d) | |
| 32515 
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
 boehmes parents: 
32504diff
changeset | 82 | |
| 32521 | 83 | fun register (init, run, done) thy = | 
| 84 | let val id = length (Actions.get thy) + 1 | |
| 85 | in | |
| 86 | thy | |
| 87 | |> init id | |
| 88 | |> Actions.map (cons (id, run, done)) | |
| 89 | end | |
| 32381 | 90 | |
| 91 | local | |
| 92 | ||
| 93 | fun log thy s = | |
| 94 | let fun append_to n = if n = "" then K () else File.append (Path.explode n) | |
| 36787 
f60e4dd6d76f
renamed Config.get_thy to Config.get_global etc. to indicate that this is not the real thing;
 wenzelm parents: 
36001diff
changeset | 95 | in append_to (Config.get_global thy logfile) (s ^ "\n") end | 
| 32381 | 96 | (* FIXME: with multithreading and parallel proofs enabled, we might need to | 
| 97 | encapsulate this inside a critical section *) | |
| 98 | ||
| 32498 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 99 | fun log_sep thy = log thy "------------------" | 
| 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 100 | |
| 32676 | 101 | fun apply_actions thy pos name info (pre, post, time) actions = | 
| 32472 
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
 boehmes parents: 
32469diff
changeset | 102 | let | 
| 32676 | 103 |     fun apply f = f {pre=pre, post=post, timeout=time, log=log thy, pos=pos, name=name}
 | 
| 32521 | 104 | fun run (id, run, _) = (apply (run id); log_sep thy) | 
| 32515 
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
 boehmes parents: 
32504diff
changeset | 105 | in (log thy info; log_sep thy; List.app run actions) end | 
| 32381 | 106 | |
| 107 | fun in_range _ _ NONE = true | |
| 108 | | in_range l r (SOME i) = (l <= i andalso (r < 0 orelse i <= r)) | |
| 109 | ||
| 110 | fun only_within_range thy pos f x = | |
| 36787 
f60e4dd6d76f
renamed Config.get_thy to Config.get_global etc. to indicate that this is not the real thing;
 wenzelm parents: 
36001diff
changeset | 111 | let val l = Config.get_global thy start_line and r = Config.get_global thy end_line | 
| 32472 
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
 boehmes parents: 
32469diff
changeset | 112 | in if in_range l r (Position.line_of pos) then f x else () end | 
| 32381 | 113 | |
| 114 | in | |
| 115 | ||
| 32521 | 116 | fun run_actions tr pre post = | 
| 32381 | 117 | let | 
| 118 | val thy = Proof.theory_of pre | |
| 119 | val pos = Toplevel.pos_of tr | |
| 120 | val name = Toplevel.name_of tr | |
| 36787 
f60e4dd6d76f
renamed Config.get_thy to Config.get_global etc. to indicate that this is not the real thing;
 wenzelm parents: 
36001diff
changeset | 121 | val st = (pre, post, Time.fromSeconds (Config.get_global thy timeout)) | 
| 32472 
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
 boehmes parents: 
32469diff
changeset | 122 | |
| 
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
 boehmes parents: 
32469diff
changeset | 123 | val str0 = string_of_int o the_default 0 | 
| 
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
 boehmes parents: 
32469diff
changeset | 124 | val loc = str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos) | 
| 
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
 boehmes parents: 
32469diff
changeset | 125 |     val info = "\n\nat " ^ loc ^ " (" ^ name ^ "):"
 | 
| 32381 | 126 | in | 
| 32676 | 127 | only_within_range thy pos (apply_actions thy pos name info st) (Actions.get thy) | 
| 32381 | 128 | end | 
| 129 | ||
| 32521 | 130 | fun done_actions st = | 
| 131 | let | |
| 132 | val thy = Toplevel.theory_of st | |
| 133 | val _ = log thy "\n\n"; | |
| 134 | in | |
| 135 | thy | |
| 136 | |> Actions.get | |
| 137 |     |> List.app (fn (id, _, done) => done id {last=st, log=log thy})
 | |
| 138 | end | |
| 139 | ||
| 32381 | 140 | end | 
| 141 | ||
| 32504 | 142 | val whitelist = ["apply", "by", "proof"] | 
| 32468 
3e6f5365971e
Mirabelle: explicit command blacklist, preliminary documentation
 boehmes parents: 
32396diff
changeset | 143 | |
| 32381 | 144 | fun step_hook tr pre post = | 
| 145 | (* FIXME: might require wrapping into "interruptible" *) | |
| 146 | if can (Proof.assert_backward o Toplevel.proof_of) pre andalso | |
| 32504 | 147 | member (op =) whitelist (Toplevel.name_of tr) | 
| 32521 | 148 | then run_actions tr (Toplevel.proof_of pre) (SOME post) | 
| 149 | else if not (Toplevel.is_toplevel pre) andalso Toplevel.is_toplevel post | |
| 150 | then done_actions pre | |
| 32381 | 151 | else () (* FIXME: add theory_hook here *) | 
| 152 | ||
| 153 | ||
| 154 | ||
| 155 | (* Mirabelle utility functions *) | |
| 156 | ||
| 32472 
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
 boehmes parents: 
32469diff
changeset | 157 | fun can_apply time tac st = | 
| 32469 | 158 | let | 
| 35592 
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
 wenzelm parents: 
34052diff
changeset | 159 |     val {context = ctxt, facts, goal} = Proof.goal st
 | 
| 32469 | 160 | val full_tac = HEADGOAL (Method.insert_tac facts THEN' tac ctxt) | 
| 32381 | 161 | in | 
| 32472 
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
 boehmes parents: 
32469diff
changeset | 162 | (case TimeLimit.timeLimit time (Seq.pull o full_tac) goal of | 
| 32381 | 163 | SOME (thm, _) => true | 
| 164 | | NONE => false) | |
| 165 | end | |
| 166 | ||
| 167 | local | |
| 168 | ||
| 169 | fun fold_body_thms f = | |
| 170 | let | |
| 171 |     fun app n (PBody {thms, ...}) = thms |> fold (fn (i, (name, prop, body)) =>
 | |
| 172 | fn (x, seen) => | |
| 173 | if Inttab.defined seen i then (x, seen) | |
| 174 | else | |
| 175 | let | |
| 176 | val body' = Future.join body | |
| 177 | val (x', seen') = app (n + (if name = "" then 0 else 1)) body' | |
| 178 | (x, Inttab.update (i, ()) seen) | |
| 179 | in (x' |> n = 0 ? f (name, prop, body'), seen') end) | |
| 180 | in fn bodies => fn x => #1 (fold (app 0) bodies (x, Inttab.empty)) end | |
| 181 | ||
| 182 | in | |
| 183 | ||
| 184 | fun theorems_in_proof_term thm = | |
| 185 | let | |
| 186 | val all_thms = PureThy.all_thms_of (Thm.theory_of_thm thm) | |
| 187 | fun collect (s, _, _) = if s <> "" then insert (op =) s else I | |
| 188 | fun member_of xs (x, y) = if member (op =) xs x then SOME y else NONE | |
| 189 | fun resolve_thms names = map_filter (member_of names) all_thms | |
| 190 | in | |
| 191 | resolve_thms (fold_body_thms collect [Thm.proof_body_of thm] []) | |
| 192 | end | |
| 193 | ||
| 194 | end | |
| 195 | ||
| 196 | fun theorems_of_sucessful_proof state = | |
| 197 | (case state of | |
| 198 | NONE => [] | |
| 199 | | SOME st => | |
| 200 | if not (Toplevel.is_proof st) then [] | |
| 35592 
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
 wenzelm parents: 
34052diff
changeset | 201 | else theorems_in_proof_term (#goal (Proof.goal (Toplevel.proof_of st)))) | 
| 32381 | 202 | |
| 203 | fun get_setting settings (key, default) = | |
| 204 | the_default default (AList.lookup (op =) settings key) | |
| 205 | ||
| 206 | fun get_int_setting settings (key, default) = | |
| 207 | (case Option.map Int.fromString (AList.lookup (op =) settings key) of | |
| 208 | SOME (SOME i) => i | |
| 209 |   | SOME NONE => error ("bad option: " ^ key)
 | |
| 210 | | NONE => default) | |
| 211 | ||
| 32498 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 212 | fun cpu_time f x = | 
| 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 213 | let | 
| 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 214 | val start = start_timing () | 
| 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 215 | val result = Exn.capture (fn () => f x) () | 
| 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 216 | val time = Time.toMilliseconds (#cpu (end_timing start)) | 
| 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 217 | in (Exn.release result, time) end | 
| 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 218 | |
| 32381 | 219 | end |