author | boehmes |
Thu, 03 Sep 2009 22:47:31 +0200 | |
changeset 32515 | e7c0d3c0494a |
parent 32511 | 43d2ac4aa2de |
child 32518 | e3c4e337196c |
permissions | -rw-r--r-- |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
1 |
(* Title: mirabelle_sledgehammer.ML |
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
2 |
Author: Jasmin Blanchette and Sascha Boehme |
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
3 |
*) |
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
4 |
|
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
5 |
structure Mirabelle_Sledgehammer : MIRABELLE_ACTION = |
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
6 |
struct |
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
7 |
|
32454
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
8 |
fun thms_of_name ctxt name = |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
9 |
let |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
10 |
val lex = OuterKeyword.get_lexicons |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
11 |
val get = maps (ProofContext.get_fact ctxt o fst) |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
12 |
in |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
13 |
Source.of_string name |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
14 |
|> Symbol.source {do_recover=false} |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
15 |
|> OuterLex.source {do_recover=SOME false} lex Position.start |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
16 |
|> OuterLex.source_proper |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
17 |
|> Source.source OuterLex.stopper (SpecParse.xthms1 >> get) NONE |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
18 |
|> Source.exhaust |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
19 |
end |
32452
d84edd022efe
apply metis with found theorems in case sledgehammer was successful
boehmes
parents:
32434
diff
changeset
|
20 |
|
32455
c71414177792
Mirabelle sledgehammer: added option to keep problem files, enabled "metis" switch again (was accidentally removed)
boehmes
parents:
32454
diff
changeset
|
21 |
fun safe init done f x = |
c71414177792
Mirabelle sledgehammer: added option to keep problem files, enabled "metis" switch again (was accidentally removed)
boehmes
parents:
32454
diff
changeset
|
22 |
let |
c71414177792
Mirabelle sledgehammer: added option to keep problem files, enabled "metis" switch again (was accidentally removed)
boehmes
parents:
32454
diff
changeset
|
23 |
val y = init x |
c71414177792
Mirabelle sledgehammer: added option to keep problem files, enabled "metis" switch again (was accidentally removed)
boehmes
parents:
32454
diff
changeset
|
24 |
val z = Exn.capture f y |
c71414177792
Mirabelle sledgehammer: added option to keep problem files, enabled "metis" switch again (was accidentally removed)
boehmes
parents:
32454
diff
changeset
|
25 |
val _ = done y |
c71414177792
Mirabelle sledgehammer: added option to keep problem files, enabled "metis" switch again (was accidentally removed)
boehmes
parents:
32454
diff
changeset
|
26 |
in Exn.release z end |
c71414177792
Mirabelle sledgehammer: added option to keep problem files, enabled "metis" switch again (was accidentally removed)
boehmes
parents:
32454
diff
changeset
|
27 |
|
c71414177792
Mirabelle sledgehammer: added option to keep problem files, enabled "metis" switch again (was accidentally removed)
boehmes
parents:
32454
diff
changeset
|
28 |
val proverK = "prover" |
c71414177792
Mirabelle sledgehammer: added option to keep problem files, enabled "metis" switch again (was accidentally removed)
boehmes
parents:
32454
diff
changeset
|
29 |
val keepK = "keep" |
c71414177792
Mirabelle sledgehammer: added option to keep problem files, enabled "metis" switch again (was accidentally removed)
boehmes
parents:
32454
diff
changeset
|
30 |
val metisK = "metis" |
32511 | 31 |
val full_typesK = "full_types" |
32455
c71414177792
Mirabelle sledgehammer: added option to keep problem files, enabled "metis" switch again (was accidentally removed)
boehmes
parents:
32454
diff
changeset
|
32 |
|
32515
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
33 |
val sh_tag = "sledgehammer: " |
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
34 |
val metis_tag = "metis (sledgehammer): " |
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
35 |
|
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
36 |
|
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
37 |
local |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
38 |
|
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
39 |
fun init NONE = !AtpWrapper.destdir |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
40 |
| init (SOME path) = |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
41 |
let |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
42 |
(* Warning: we implicitly assume single-threaded execution here! *) |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
43 |
val old = !AtpWrapper.destdir |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
44 |
val _ = AtpWrapper.destdir := path |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
45 |
in old end |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
46 |
|
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
47 |
fun done path = AtpWrapper.destdir := path |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
48 |
|
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
49 |
fun run prover_name timeout st _ = |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
50 |
let |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
51 |
val prover = the (AtpManager.get_prover prover_name (Proof.theory_of st)) |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
52 |
val atp_timeout = AtpManager.get_timeout () |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
53 |
val atp = prover atp_timeout NONE NONE prover_name 1 |
32510 | 54 |
val (success, (message, thm_names), time, _, _, _) = |
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
55 |
TimeLimit.timeLimit timeout atp (Proof.get_goal st) |
32510 | 56 |
in if success then (message, SOME (time, thm_names)) else (message, NONE) end |
57 |
handle ResHolClause.TOO_TRIVIAL => ("trivial", SOME (0, [])) |
|
58 |
| TimeLimit.TimeOut => ("time out", NONE) |
|
59 |
| ERROR msg => ("error: " ^ msg, NONE) |
|
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
60 |
|
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
61 |
in |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
62 |
|
32515
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
63 |
fun run_sledgehammer args thm_names {pre=st, timeout, log, ...} = |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
64 |
let |
32434
6b93b73a712b
selectable prover for sledgehammer, proper environment lookup
boehmes
parents:
32385
diff
changeset
|
65 |
val prover_name = |
32455
c71414177792
Mirabelle sledgehammer: added option to keep problem files, enabled "metis" switch again (was accidentally removed)
boehmes
parents:
32454
diff
changeset
|
66 |
AList.lookup (op =) args proverK |
32434
6b93b73a712b
selectable prover for sledgehammer, proper environment lookup
boehmes
parents:
32385
diff
changeset
|
67 |
|> the_default (hd (space_explode " " (AtpManager.get_atps ()))) |
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
68 |
val dir = AList.lookup (op =) args keepK |
32510 | 69 |
val (msg, result) = safe init done (run prover_name timeout st) dir |
70 |
val _ = |
|
71 |
if is_some result |
|
32515
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
72 |
then log (sh_tag ^ "succeeded (" ^ string_of_int (fst (the result)) ^ |
32510 | 73 |
") [" ^ prover_name ^ "]:\n" ^ msg) |
32515
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
74 |
else log (sh_tag ^ "failed: " ^ msg) |
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
75 |
in change thm_names (K (Option.map snd result)) end |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
76 |
|
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
77 |
end |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
78 |
|
32452
d84edd022efe
apply metis with found theorems in case sledgehammer was successful
boehmes
parents:
32434
diff
changeset
|
79 |
|
32510 | 80 |
fun with_time true t = "succeeded (" ^ string_of_int t ^ ")" |
81 |
| with_time false t = "failed (" ^ string_of_int t ^ ")" |
|
82 |
||
32515
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
83 |
fun run_metis args thm_names {pre=st, timeout, log, ...} = |
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
84 |
let |
32454
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
85 |
fun get_thms ctxt = maps (thms_of_name ctxt) |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
86 |
fun metis thms ctxt = MetisTools.metis_tac ctxt thms |
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
87 |
fun apply_metis thms = Mirabelle.can_apply timeout (metis thms) st |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
88 |
fun timed_metis thms = |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
89 |
uncurry with_time (Mirabelle.cpu_time apply_metis thms) |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
90 |
handle TimeLimit.TimeOut => "time out" |
32503
14efbc20b708
Mirabelle: logging of exceptions (works only for PolyML)
boehmes
parents:
32498
diff
changeset
|
91 |
| ERROR msg => "error: " ^ msg |
32515
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
92 |
fun log_metis s = log (metis_tag ^ s) |
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
93 |
in |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
94 |
if not (AList.defined (op =) args metisK) then () |
32515
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
95 |
else if is_none (!thm_names) then () |
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
96 |
else |
32503
14efbc20b708
Mirabelle: logging of exceptions (works only for PolyML)
boehmes
parents:
32498
diff
changeset
|
97 |
log "-----" |
32515
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
98 |
|> K (these (!thm_names)) |
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
99 |
|> get_thms (Proof.context_of st) |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
100 |
|> timed_metis |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
101 |
|> log_metis |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
102 |
end |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
103 |
|
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
104 |
|
32515
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
105 |
fun sledgehammer_action args (st as {log, ...}) = |
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
106 |
let |
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
107 |
val thm_names = ref (NONE : string list option) |
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
108 |
val _ = Mirabelle.catch sh_tag (run_sledgehammer args thm_names) st |
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
109 |
val _ = Mirabelle.catch metis_tag (run_metis args thm_names) st |
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
110 |
in () end |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
111 |
|
32511 | 112 |
fun invoke args = |
32515
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
113 |
let |
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
114 |
val _ = AtpManager.set_full_types (AList.defined (op =) args full_typesK) |
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
115 |
in Mirabelle.register (sledgehammer_action args) end |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
116 |
|
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
117 |
end |