author | blanchet |
Wed, 18 Jul 2012 08:44:04 +0200 | |
changeset 48324 | 3ee5b5589402 |
parent 48321 | c552d7f1720b |
child 48333 | 2250197977dc |
permissions | -rw-r--r-- |
48234 | 1 |
(* Title: HOL/TPTP/mash_export.ML |
2 |
Author: Jasmin Blanchette, TU Muenchen |
|
3 |
Copyright 2012 |
|
4 |
||
5 |
Export Isabelle theory information for MaSh (Machine-learning for Sledgehammer). |
|
6 |
*) |
|
7 |
||
8 |
signature MASH_EXPORT = |
|
9 |
sig |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
10 |
type params = Sledgehammer_Provers.params |
48235 | 11 |
|
48304 | 12 |
val generate_accessibility : theory -> bool -> string -> unit |
48318 | 13 |
val generate_features : |
14 |
Proof.context -> string -> theory -> bool -> string -> unit |
|
15 |
val generate_isa_dependencies : |
|
48324 | 16 |
theory -> bool -> string -> unit |
48318 | 17 |
val generate_prover_dependencies : |
48304 | 18 |
Proof.context -> params -> theory -> bool -> string -> unit |
48318 | 19 |
val generate_commands : Proof.context -> string -> theory -> string -> unit |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
20 |
val generate_iter_suggestions : |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
21 |
Proof.context -> params -> theory -> int -> string -> unit |
48234 | 22 |
end; |
23 |
||
24 |
structure MaSh_Export : MASH_EXPORT = |
|
25 |
struct |
|
26 |
||
48304 | 27 |
open Sledgehammer_Fact |
28 |
open Sledgehammer_Filter_Iter |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
29 |
open Sledgehammer_Filter_MaSh |
48245 | 30 |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
31 |
fun thy_map_from_facts ths = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
32 |
ths |> sort (thm_ord o swap o pairself snd) |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
33 |
|> map (snd #> `(theory_of_thm #> Context.theory_name)) |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
34 |
|> AList.coalesce (op =) |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
35 |
|> map (apsnd (map Thm.get_name_hint)) |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
36 |
|
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
37 |
fun has_thy thy th = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
38 |
Context.theory_name thy = Context.theory_name (theory_of_thm th) |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
39 |
|
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
40 |
fun parent_facts thy thy_map = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
41 |
let |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
42 |
fun add_last thy = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
43 |
case AList.lookup (op =) thy_map (Context.theory_name thy) of |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
44 |
SOME (last_fact :: _) => insert (op =) last_fact |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
45 |
| _ => add_parent thy |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
46 |
and add_parent thy = fold add_last (Theory.parents_of thy) |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
47 |
in add_parent thy [] end |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
48 |
|
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
49 |
val thy_name_of_fact = hd o Long_Name.explode |
48304 | 50 |
fun thy_of_fact thy = Context.this_theory thy o thy_name_of_fact |
51 |
||
48324 | 52 |
val all_names = |
53 |
filter_out (is_likely_tautology_or_too_meta) |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
54 |
#> map (rpair () o Thm.get_name_hint) #> Symtab.make |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
55 |
|
48304 | 56 |
fun generate_accessibility thy include_thy file_name = |
57 |
let |
|
58 |
val path = file_name |> Path.explode |
|
59 |
val _ = File.write path "" |
|
60 |
fun do_fact fact prevs = |
|
61 |
let |
|
62 |
val s = escape_meta fact ^ ": " ^ escape_metas prevs ^ "\n" |
|
63 |
val _ = File.append path s |
|
64 |
in [fact] end |
|
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48304
diff
changeset
|
65 |
val thy_map = |
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48304
diff
changeset
|
66 |
all_facts_of thy Termtab.empty |
48304 | 67 |
|> not include_thy ? filter_out (has_thy thy o snd) |
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48304
diff
changeset
|
68 |
|> thy_map_from_facts |
48304 | 69 |
fun do_thy facts = |
70 |
let |
|
71 |
val thy = thy_of_fact thy (hd facts) |
|
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48304
diff
changeset
|
72 |
val parents = parent_facts thy thy_map |
48304 | 73 |
in fold do_fact facts parents; () end |
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48304
diff
changeset
|
74 |
in fold_rev (fn (_, facts) => K (do_thy facts)) thy_map () end |
48304 | 75 |
|
48318 | 76 |
fun generate_features ctxt prover thy include_thy file_name = |
48304 | 77 |
let |
78 |
val path = file_name |> Path.explode |
|
79 |
val _ = File.write path "" |
|
80 |
val css_table = clasimpset_rule_table_of ctxt |
|
81 |
val facts = |
|
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48304
diff
changeset
|
82 |
all_facts_of thy css_table |
48304 | 83 |
|> not include_thy ? filter_out (has_thy thy o snd) |
84 |
fun do_fact ((_, (_, status)), th) = |
|
85 |
let |
|
86 |
val name = Thm.get_name_hint th |
|
48318 | 87 |
val feats = |
88 |
features_of ctxt prover (theory_of_thm th) status [prop_of th] |
|
48304 | 89 |
val s = escape_meta name ^ ": " ^ escape_metas feats ^ "\n" |
90 |
in File.append path s end |
|
91 |
in List.app do_fact facts end |
|
92 |
||
48324 | 93 |
fun generate_isa_dependencies thy include_thy file_name = |
48304 | 94 |
let |
95 |
val path = file_name |> Path.explode |
|
96 |
val _ = File.write path "" |
|
97 |
val ths = |
|
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48304
diff
changeset
|
98 |
all_facts_of thy Termtab.empty |
48304 | 99 |
|> not include_thy ? filter_out (has_thy thy o snd) |
100 |
|> map snd |
|
48324 | 101 |
val all_names = all_names ths |
48304 | 102 |
fun do_thm th = |
103 |
let |
|
104 |
val name = Thm.get_name_hint th |
|
105 |
val deps = isabelle_dependencies_of all_names th |
|
106 |
val s = escape_meta name ^ ": " ^ escape_metas deps ^ "\n" |
|
107 |
in File.append path s end |
|
108 |
in List.app do_thm ths end |
|
109 |
||
48318 | 110 |
fun generate_prover_dependencies ctxt (params as {provers, max_facts, ...}) thy |
111 |
include_thy file_name = |
|
48304 | 112 |
let |
113 |
val path = file_name |> Path.explode |
|
114 |
val _ = File.write path "" |
|
48318 | 115 |
val prover = hd provers |
48304 | 116 |
val facts = |
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48304
diff
changeset
|
117 |
all_facts_of thy Termtab.empty |
48304 | 118 |
|> not include_thy ? filter_out (has_thy thy o snd) |
119 |
val ths = facts |> map snd |
|
48324 | 120 |
val all_names = all_names ths |
48304 | 121 |
fun is_dep dep (_, th) = Thm.get_name_hint th = dep |
122 |
fun add_isa_dep facts dep accum = |
|
123 |
if exists (is_dep dep) accum then |
|
124 |
accum |
|
125 |
else case find_first (is_dep dep) facts of |
|
126 |
SOME ((name, status), th) => accum @ [((name, status), th)] |
|
127 |
| NONE => accum (* shouldn't happen *) |
|
128 |
fun fix_name ((_, stature), th) = |
|
129 |
((fn () => th |> Thm.get_name_hint, stature), th) |
|
130 |
fun do_thm th = |
|
131 |
let |
|
132 |
val name = Thm.get_name_hint th |
|
133 |
val goal = goal_of_thm thy th |
|
134 |
val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1 |
|
135 |
val deps = |
|
136 |
case isabelle_dependencies_of all_names th of |
|
137 |
[] => [] |
|
138 |
| isa_dep as [_] => isa_dep (* can hardly beat that *) |
|
139 |
| isa_deps => |
|
140 |
let |
|
141 |
val facts = |
|
142 |
facts |> filter (fn (_, th') => thm_ord (th', th) = LESS) |
|
143 |
val facts = |
|
48318 | 144 |
facts |> iterative_relevant_facts ctxt params prover |
48304 | 145 |
(the max_facts) NONE hyp_ts concl_t |
146 |
|> fold (add_isa_dep facts) isa_deps |
|
147 |
|> map fix_name |
|
148 |
in |
|
48321 | 149 |
case run_prover_for_mash ctxt params prover facts goal of |
48304 | 150 |
{outcome = NONE, used_facts, ...} => |
151 |
used_facts |> map fst |> sort string_ord |
|
152 |
| _ => isa_deps |
|
153 |
end |
|
154 |
val s = escape_meta name ^ ": " ^ escape_metas deps ^ "\n" |
|
155 |
in File.append path s end |
|
156 |
in List.app do_thm ths end |
|
157 |
||
48318 | 158 |
fun generate_commands ctxt prover thy file_name = |
48234 | 159 |
let |
160 |
val path = file_name |> Path.explode |
|
161 |
val _ = File.write path "" |
|
48299 | 162 |
val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt |
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48304
diff
changeset
|
163 |
val facts = all_facts_of thy css_table |
48234 | 164 |
val (new_facts, old_facts) = |
165 |
facts |> List.partition (has_thy thy o snd) |
|
166 |
|>> sort (thm_ord o pairself snd) |
|
167 |
val ths = facts |> map snd |
|
48324 | 168 |
val all_names = all_names ths |
48234 | 169 |
fun do_fact ((_, (_, status)), th) prevs = |
170 |
let |
|
171 |
val name = Thm.get_name_hint th |
|
48318 | 172 |
val feats = features_of ctxt prover thy status [prop_of th] |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
173 |
val deps = isabelle_dependencies_of all_names th |
48234 | 174 |
val kind = Thm.legacy_get_kind th |
48304 | 175 |
val core = escape_metas prevs ^ "; " ^ escape_metas feats |
48234 | 176 |
val query = if kind <> "" then "? " ^ core ^ "\n" else "" |
48300 | 177 |
val update = |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
178 |
"! " ^ escape_meta name ^ ": " ^ core ^ "; " ^ |
48304 | 179 |
escape_metas deps ^ "\n" |
48234 | 180 |
val _ = File.append path (query ^ update) |
181 |
in [name] end |
|
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48304
diff
changeset
|
182 |
val thy_map = old_facts |> thy_map_from_facts |
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48304
diff
changeset
|
183 |
val parents = parent_facts thy thy_map |
48239
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
184 |
in fold do_fact new_facts parents; () end |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
185 |
|
48292 | 186 |
fun generate_iter_suggestions ctxt (params as {provers, ...}) thy max_relevant |
187 |
file_name = |
|
48239
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
188 |
let |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
189 |
val path = file_name |> Path.explode |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
190 |
val _ = File.write path "" |
48318 | 191 |
val prover = hd provers |
48299 | 192 |
val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt |
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48304
diff
changeset
|
193 |
val facts = all_facts_of thy css_table |
48239
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
194 |
val (new_facts, old_facts) = |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
195 |
facts |> List.partition (has_thy thy o snd) |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
196 |
|>> sort (thm_ord o pairself snd) |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
197 |
fun do_fact (fact as (_, th)) old_facts = |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
198 |
let |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
199 |
val name = Thm.get_name_hint th |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
48246
diff
changeset
|
200 |
val goal = goal_of_thm thy th |
48292 | 201 |
val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1 |
48239
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
202 |
val kind = Thm.legacy_get_kind th |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
203 |
val _ = |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
204 |
if kind <> "" then |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
205 |
let |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
206 |
val suggs = |
48292 | 207 |
old_facts |
208 |
|> Sledgehammer_Filter_Iter.iterative_relevant_facts ctxt params |
|
48318 | 209 |
prover max_relevant NONE hyp_ts concl_t |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
210 |
|> map (fn ((name, _), _) => name ()) |
48292 | 211 |
|> sort string_ord |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
212 |
val s = escape_meta name ^ ": " ^ escape_metas suggs ^ "\n" |
48239
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
213 |
in File.append path s end |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
214 |
else |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
215 |
() |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
216 |
in fact :: old_facts end |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
217 |
in fold do_fact new_facts old_facts; () end |
48234 | 218 |
|
219 |
end; |