author | boehmes |
Thu, 03 Sep 2009 14:05:13 +0200 | |
changeset 32503 | 14efbc20b708 |
parent 32498 | 1132c7c13f36 |
child 32504 | 7a06bf89c038 |
permissions | -rw-r--r-- |
32381 | 1 |
(* Title: mirabelle.ML |
2 |
Author: Jasmin Blanchette and Sascha Boehme |
|
3 |
*) |
|
4 |
||
5 |
signature MIRABELLE = |
|
6 |
sig |
|
32495 | 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:
32381
diff
changeset
|
10 |
val start_line : int Config.T |
98674ac811c4
Mirabelle tool script conforming to standard Isabelle tool interface,
boehmes
parents:
32381
diff
changeset
|
11 |
val end_line : int Config.T |
32495 | 12 |
val setup : theory -> theory |
32381 | 13 |
|
32495 | 14 |
(* core *) |
15 |
type action |
|
16 |
val register : string * action -> theory -> theory |
|
17 |
val step_hook : Toplevel.transition -> Toplevel.state -> Toplevel.state -> |
|
18 |
unit |
|
19 |
||
20 |
(* utility functions *) |
|
32381 | 21 |
val goal_thm_of : Proof.state -> thm |
32469 | 22 |
val can_apply : Time.time -> (Proof.context -> int -> tactic) -> |
23 |
Proof.state -> bool |
|
32381 | 24 |
val theorems_in_proof_term : Thm.thm -> Thm.thm list |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
32383
diff
changeset
|
25 |
val theorems_of_sucessful_proof : Toplevel.state option -> Thm.thm list |
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
32383
diff
changeset
|
26 |
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:
32383
diff
changeset
|
27 |
val get_int_setting : (string * string) list -> string * int -> int |
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
28 |
val cpu_time : ('a -> 'b) -> 'a -> 'b * int |
32381 | 29 |
end |
30 |
||
31 |
||
32 |
||
32497 | 33 |
structure Mirabelle : MIRABELLE = |
32495 | 34 |
struct |
35 |
||
36 |
(* Mirabelle configuration *) |
|
37 |
||
38 |
val (logfile, setup1) = Attrib.config_string "mirabelle_logfile" "" |
|
39 |
val (timeout, setup2) = Attrib.config_int "mirabelle_timeout" 30 |
|
40 |
val (start_line, setup3) = Attrib.config_int "mirabelle_start_line" 0 |
|
41 |
val (end_line, setup4) = Attrib.config_int "mirabelle_end_line" ~1 |
|
42 |
||
43 |
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:
32383
diff
changeset
|
44 |
|
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
32383
diff
changeset
|
45 |
|
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
32383
diff
changeset
|
46 |
|
32381 | 47 |
(* Mirabelle core *) |
48 |
||
32472
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
49 |
type action = {pre: Proof.state, post: Toplevel.state option, |
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
50 |
timeout: Time.time, log: string -> unit} -> unit |
32381 | 51 |
|
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
32383
diff
changeset
|
52 |
structure Actions = TheoryDataFun |
32381 | 53 |
( |
54 |
type T = action Symtab.table |
|
55 |
val empty = Symtab.empty |
|
56 |
val copy = I |
|
57 |
val extend = I |
|
58 |
fun merge _ = Symtab.merge (K true) |
|
59 |
) |
|
60 |
||
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
32383
diff
changeset
|
61 |
val register = Actions.map o Symtab.update_new |
32381 | 62 |
|
63 |
local |
|
64 |
||
65 |
fun log thy s = |
|
66 |
let fun append_to n = if n = "" then K () else File.append (Path.explode n) |
|
67 |
in append_to (Config.get_thy thy logfile) (s ^ "\n") end |
|
68 |
(* FIXME: with multithreading and parallel proofs enabled, we might need to |
|
69 |
encapsulate this inside a critical section *) |
|
70 |
||
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
71 |
fun log_sep thy = log thy "------------------" |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
72 |
|
32503
14efbc20b708
Mirabelle: logging of exceptions (works only for PolyML)
boehmes
parents:
32498
diff
changeset
|
73 |
fun log_exn thy name (exn : exn) = log thy ("Unhandled exception in action " ^ |
14efbc20b708
Mirabelle: logging of exceptions (works only for PolyML)
boehmes
parents:
32498
diff
changeset
|
74 |
quote name ^ ":\n" ^ PolyML.makestring exn) |
32472
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
75 |
|
32503
14efbc20b708
Mirabelle: logging of exceptions (works only for PolyML)
boehmes
parents:
32498
diff
changeset
|
76 |
fun try_with f g x = SOME (g x) |
14efbc20b708
Mirabelle: logging of exceptions (works only for PolyML)
boehmes
parents:
32498
diff
changeset
|
77 |
handle (exn as Exn.Interrupt) => reraise exn | exn => (f exn; NONE) |
14efbc20b708
Mirabelle: logging of exceptions (works only for PolyML)
boehmes
parents:
32498
diff
changeset
|
78 |
|
14efbc20b708
Mirabelle: logging of exceptions (works only for PolyML)
boehmes
parents:
32498
diff
changeset
|
79 |
fun capture_exns thy name f x = (try_with (log_exn thy name) f x; log_sep thy) |
32381 | 80 |
|
32472
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
81 |
fun apply_actions thy info (pre, post, time) actions = |
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
82 |
let |
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
83 |
fun apply (name, action) = |
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
84 |
{pre=pre, post=post, timeout=time, log=log thy} |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
85 |
|> capture_exns thy name action |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
86 |
in (log thy info; log_sep thy; List.app apply actions) end |
32381 | 87 |
|
88 |
fun in_range _ _ NONE = true |
|
89 |
| in_range l r (SOME i) = (l <= i andalso (r < 0 orelse i <= r)) |
|
90 |
||
91 |
fun only_within_range thy pos f x = |
|
92 |
let val l = Config.get_thy thy start_line and r = Config.get_thy thy end_line |
|
32472
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
93 |
in if in_range l r (Position.line_of pos) then f x else () end |
32381 | 94 |
|
95 |
in |
|
96 |
||
97 |
fun basic_hook tr pre post = |
|
98 |
let |
|
99 |
val thy = Proof.theory_of pre |
|
100 |
val pos = Toplevel.pos_of tr |
|
101 |
val name = Toplevel.name_of tr |
|
32472
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
102 |
val st = (pre, post, Time.fromSeconds (Config.get_thy thy timeout)) |
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
103 |
|
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
104 |
val str0 = string_of_int o the_default 0 |
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
105 |
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:
32469
diff
changeset
|
106 |
val info = "\n\nat " ^ loc ^ " (" ^ name ^ "):" |
32381 | 107 |
in |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
32383
diff
changeset
|
108 |
Actions.get thy |
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
32383
diff
changeset
|
109 |
|> Symtab.dest |
32472
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
110 |
|> only_within_range thy pos (apply_actions thy info st) |
32381 | 111 |
end |
112 |
||
113 |
end |
|
114 |
||
32489 | 115 |
val blacklist = ["disable_pr", "enable_pr", "done", ".", "using", "txt"] |
32468
3e6f5365971e
Mirabelle: explicit command blacklist, preliminary documentation
boehmes
parents:
32396
diff
changeset
|
116 |
|
32381 | 117 |
fun step_hook tr pre post = |
118 |
(* FIXME: might require wrapping into "interruptible" *) |
|
119 |
if can (Proof.assert_backward o Toplevel.proof_of) pre andalso |
|
32468
3e6f5365971e
Mirabelle: explicit command blacklist, preliminary documentation
boehmes
parents:
32396
diff
changeset
|
120 |
not (member (op =) blacklist (Toplevel.name_of tr)) |
32381 | 121 |
then basic_hook tr (Toplevel.proof_of pre) (SOME post) |
122 |
else () (* FIXME: add theory_hook here *) |
|
123 |
||
124 |
||
125 |
||
126 |
(* Mirabelle utility functions *) |
|
127 |
||
128 |
val goal_thm_of = snd o snd o Proof.get_goal |
|
129 |
||
32472
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
130 |
fun can_apply time tac st = |
32469 | 131 |
let |
132 |
val (ctxt, (facts, goal)) = Proof.get_goal st |
|
133 |
val full_tac = HEADGOAL (Method.insert_tac facts THEN' tac ctxt) |
|
32381 | 134 |
in |
32472
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
boehmes
parents:
32469
diff
changeset
|
135 |
(case TimeLimit.timeLimit time (Seq.pull o full_tac) goal of |
32381 | 136 |
SOME (thm, _) => true |
137 |
| NONE => false) |
|
138 |
end |
|
139 |
||
140 |
local |
|
141 |
||
142 |
fun fold_body_thms f = |
|
143 |
let |
|
144 |
fun app n (PBody {thms, ...}) = thms |> fold (fn (i, (name, prop, body)) => |
|
145 |
fn (x, seen) => |
|
146 |
if Inttab.defined seen i then (x, seen) |
|
147 |
else |
|
148 |
let |
|
149 |
val body' = Future.join body |
|
150 |
val (x', seen') = app (n + (if name = "" then 0 else 1)) body' |
|
151 |
(x, Inttab.update (i, ()) seen) |
|
152 |
in (x' |> n = 0 ? f (name, prop, body'), seen') end) |
|
153 |
in fn bodies => fn x => #1 (fold (app 0) bodies (x, Inttab.empty)) end |
|
154 |
||
155 |
in |
|
156 |
||
157 |
fun theorems_in_proof_term thm = |
|
158 |
let |
|
159 |
val all_thms = PureThy.all_thms_of (Thm.theory_of_thm thm) |
|
160 |
fun collect (s, _, _) = if s <> "" then insert (op =) s else I |
|
161 |
fun member_of xs (x, y) = if member (op =) xs x then SOME y else NONE |
|
162 |
fun resolve_thms names = map_filter (member_of names) all_thms |
|
163 |
in |
|
164 |
resolve_thms (fold_body_thms collect [Thm.proof_body_of thm] []) |
|
165 |
end |
|
166 |
||
167 |
end |
|
168 |
||
169 |
fun theorems_of_sucessful_proof state = |
|
170 |
(case state of |
|
171 |
NONE => [] |
|
172 |
| SOME st => |
|
173 |
if not (Toplevel.is_proof st) then [] |
|
174 |
else theorems_in_proof_term (goal_thm_of (Toplevel.proof_of st))) |
|
175 |
||
176 |
fun get_setting settings (key, default) = |
|
177 |
the_default default (AList.lookup (op =) settings key) |
|
178 |
||
179 |
fun get_int_setting settings (key, default) = |
|
180 |
(case Option.map Int.fromString (AList.lookup (op =) settings key) of |
|
181 |
SOME (SOME i) => i |
|
182 |
| SOME NONE => error ("bad option: " ^ key) |
|
183 |
| NONE => default) |
|
184 |
||
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
185 |
fun cpu_time f x = |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
186 |
let |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
187 |
val start = start_timing () |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
188 |
val result = Exn.capture (fn () => f x) () |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
189 |
val time = Time.toMilliseconds (#cpu (end_timing start)) |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
190 |
in (Exn.release result, time) end |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32497
diff
changeset
|
191 |
|
32381 | 192 |
end |