| author | hoelzl |
| Thu, 17 Jan 2013 11:59:12 +0100 | |
| changeset 50936 | b28f258ebc1a |
| parent 50907 | a86708897266 |
| child 50954 | 7bc58677860e |
| 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 |
|
| 50349 | 12 |
val generate_accessibility : |
13 |
Proof.context -> theory list -> bool -> string -> unit |
|
| 48318 | 14 |
val generate_features : |
| 50349 | 15 |
Proof.context -> string -> theory list -> bool -> string -> unit |
| 48333 | 16 |
val generate_isar_dependencies : |
| 50349 | 17 |
Proof.context -> theory list -> bool -> string -> unit |
|
50484
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50442
diff
changeset
|
18 |
val generate_prover_dependencies : |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
19 |
Proof.context -> params -> int * int option -> theory list -> bool -> string |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
20 |
-> unit |
| 50411 | 21 |
val generate_isar_commands : |
22 |
Proof.context -> string -> theory list -> string -> unit |
|
|
50484
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50442
diff
changeset
|
23 |
val generate_prover_commands : |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
24 |
Proof.context -> params -> int * int option -> theory list -> string -> unit |
|
48379
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
blanchet
parents:
48378
diff
changeset
|
25 |
val generate_mepo_suggestions : |
| 50906 | 26 |
Proof.context -> params -> int * int option -> theory list -> int -> string |
27 |
-> unit |
|
| 50814 | 28 |
val generate_mesh_suggestions : int -> string -> string -> string -> unit |
| 48234 | 29 |
end; |
30 |
||
31 |
structure MaSh_Export : MASH_EXPORT = |
|
32 |
struct |
|
33 |
||
|
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50484
diff
changeset
|
34 |
open Sledgehammer_Fact |
| 48381 | 35 |
open Sledgehammer_MePo |
36 |
open Sledgehammer_MaSh |
|
| 48245 | 37 |
|
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
38 |
fun in_range (from, to) j = |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
39 |
j >= from andalso (to = NONE orelse j <= the to) |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
40 |
|
| 50349 | 41 |
fun has_thm_thy th thy = |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
42 |
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
|
43 |
|
| 50349 | 44 |
fun has_thys thys th = exists (has_thm_thy th) thys |
45 |
||
46 |
fun all_facts ctxt = |
|
| 48531 | 47 |
let val css = Sledgehammer_Fact.clasimpset_rule_table_of ctxt in |
|
50442
4f6a4d32522c
don't blacklist "case" theorems -- this causes problems in MaSh later
blanchet
parents:
50434
diff
changeset
|
48 |
Sledgehammer_Fact.all_facts ctxt true false Symtab.empty [] [] css |
|
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50484
diff
changeset
|
49 |
|> sort (thm_ord o pairself snd) |
| 48531 | 50 |
end |
51 |
||
| 50349 | 52 |
fun generate_accessibility ctxt thys include_thys file_name = |
| 48304 | 53 |
let |
54 |
val path = file_name |> Path.explode |
|
55 |
val _ = File.write path "" |
|
56 |
fun do_fact fact prevs = |
|
57 |
let |
|
| 50826 | 58 |
val s = encode_str fact ^ ": " ^ encode_strs prevs ^ "\n" |
| 48304 | 59 |
val _ = File.append path s |
60 |
in [fact] end |
|
|
50611
99af6b652b3a
linearize eval driver, to work around horrible bug in previous implementation
blanchet
parents:
50582
diff
changeset
|
61 |
val facts = |
| 50349 | 62 |
all_facts ctxt |
63 |
|> not include_thys ? filter_out (has_thys thys o snd) |
|
|
50624
4d0997abce79
improved thm order hack, in case the default names are overridden
blanchet
parents:
50611
diff
changeset
|
64 |
|> map (snd #> nickname_of_thm) |
|
50611
99af6b652b3a
linearize eval driver, to work around horrible bug in previous implementation
blanchet
parents:
50582
diff
changeset
|
65 |
in fold do_fact facts []; () end |
| 48304 | 66 |
|
| 50349 | 67 |
fun generate_features ctxt prover thys include_thys file_name = |
| 48304 | 68 |
let |
69 |
val path = file_name |> Path.explode |
|
70 |
val _ = File.write path "" |
|
71 |
val facts = |
|
| 50349 | 72 |
all_facts ctxt |
73 |
|> not include_thys ? filter_out (has_thys thys o snd) |
|
| 48385 | 74 |
fun do_fact ((_, stature), th) = |
| 48304 | 75 |
let |
|
50624
4d0997abce79
improved thm order hack, in case the default names are overridden
blanchet
parents:
50611
diff
changeset
|
76 |
val name = nickname_of_thm th |
| 48318 | 77 |
val feats = |
| 48385 | 78 |
features_of ctxt prover (theory_of_thm th) stature [prop_of th] |
| 50582 | 79 |
val s = |
| 50826 | 80 |
encode_str name ^ ": " ^ encode_features (sort_wrt fst feats) ^ "\n" |
| 48304 | 81 |
in File.append path s end |
82 |
in List.app do_fact facts end |
|
83 |
||
|
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50624
diff
changeset
|
84 |
fun isar_or_prover_dependencies_of ctxt params_opt facts name_tabs th |
|
50515
c4a27ab89c9b
shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents:
50511
diff
changeset
|
85 |
isar_deps_opt = |
|
50754
74a6adcb96ac
also generate queries for goals with too many Isar dependencies
blanchet
parents:
50735
diff
changeset
|
86 |
case params_opt of |
|
74a6adcb96ac
also generate queries for goals with too many Isar dependencies
blanchet
parents:
50735
diff
changeset
|
87 |
SOME (params as {provers = prover :: _, ...}) =>
|
|
74a6adcb96ac
also generate queries for goals with too many Isar dependencies
blanchet
parents:
50735
diff
changeset
|
88 |
prover_dependencies_of ctxt params prover 0 facts name_tabs th |> snd |
|
74a6adcb96ac
also generate queries for goals with too many Isar dependencies
blanchet
parents:
50735
diff
changeset
|
89 |
| NONE => |
|
74a6adcb96ac
also generate queries for goals with too many Isar dependencies
blanchet
parents:
50735
diff
changeset
|
90 |
case isar_deps_opt of |
|
74a6adcb96ac
also generate queries for goals with too many Isar dependencies
blanchet
parents:
50735
diff
changeset
|
91 |
SOME deps => deps |
|
74a6adcb96ac
also generate queries for goals with too many Isar dependencies
blanchet
parents:
50735
diff
changeset
|
92 |
| NONE => isar_dependencies_of name_tabs th |
| 50411 | 93 |
|
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
94 |
fun generate_isar_or_prover_dependencies ctxt params_opt range thys include_thys |
|
50484
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50442
diff
changeset
|
95 |
file_name = |
| 48304 | 96 |
let |
97 |
val path = file_name |> Path.explode |
|
98 |
val facts = |
|
|
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50484
diff
changeset
|
99 |
all_facts ctxt |> not include_thys ? filter_out (has_thys thys o snd) |
|
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50624
diff
changeset
|
100 |
val name_tabs = build_name_tables nickname_of_thm facts |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
101 |
fun do_fact (j, (_, th)) = |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
102 |
if in_range range j then |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
103 |
let |
|
50624
4d0997abce79
improved thm order hack, in case the default names are overridden
blanchet
parents:
50611
diff
changeset
|
104 |
val name = nickname_of_thm th |
| 50561 | 105 |
val _ = tracing ("Fact " ^ string_of_int j ^ ": " ^ name)
|
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
106 |
val deps = |
|
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50624
diff
changeset
|
107 |
isar_or_prover_dependencies_of ctxt params_opt facts name_tabs th |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
108 |
NONE |
| 50826 | 109 |
in encode_str name ^ ": " ^ encode_strs deps ^ "\n" end |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
110 |
else |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
111 |
"" |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
112 |
val lines = Par_List.map do_fact (tag_list 1 facts) |
| 50519 | 113 |
in File.write_list path lines end |
| 48304 | 114 |
|
| 50411 | 115 |
fun generate_isar_dependencies ctxt = |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
116 |
generate_isar_or_prover_dependencies ctxt NONE (1, NONE) |
| 50411 | 117 |
|
|
50484
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50442
diff
changeset
|
118 |
fun generate_prover_dependencies ctxt params = |
|
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50442
diff
changeset
|
119 |
generate_isar_or_prover_dependencies ctxt (SOME params) |
| 50411 | 120 |
|
|
50515
c4a27ab89c9b
shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents:
50511
diff
changeset
|
121 |
fun is_bad_query ctxt ho_atp th isar_deps = |
|
50859
c0f38015a632
don't generate queries with empty dependency list
blanchet
parents:
50829
diff
changeset
|
122 |
Thm.legacy_get_kind th = "" orelse |
|
c0f38015a632
don't generate queries with empty dependency list
blanchet
parents:
50829
diff
changeset
|
123 |
null (these (trim_dependencies th isar_deps)) orelse |
|
50523
0799339fea0f
get rid of some junk facts in the MaSh evaluation driver
blanchet
parents:
50519
diff
changeset
|
124 |
is_blacklisted_or_something ctxt ho_atp (Thm.get_name_hint th) |
|
50515
c4a27ab89c9b
shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents:
50511
diff
changeset
|
125 |
|
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
126 |
fun generate_isar_or_prover_commands ctxt prover params_opt range thys |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
127 |
file_name = |
| 48234 | 128 |
let |
|
50511
8825c36cb1ce
don't query blacklisted theorems in evaluation driver
blanchet
parents:
50485
diff
changeset
|
129 |
val ho_atp = Sledgehammer_Provers.is_ho_atp ctxt prover |
| 48234 | 130 |
val path = file_name |> Path.explode |
| 50349 | 131 |
val facts = all_facts ctxt |
|
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50484
diff
changeset
|
132 |
val (new_facts, old_facts) = facts |> List.partition (has_thys thys o snd) |
|
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50624
diff
changeset
|
133 |
val name_tabs = build_name_tables nickname_of_thm facts |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
134 |
fun do_fact (j, ((name, ((_, stature), th)), prevs)) = |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
135 |
if in_range range j then |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
136 |
let |
| 50561 | 137 |
val _ = tracing ("Fact " ^ string_of_int j ^ ": " ^ name)
|
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
138 |
val feats = |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
139 |
features_of ctxt prover (theory_of_thm th) stature [prop_of th] |
|
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50624
diff
changeset
|
140 |
val isar_deps = isar_dependencies_of name_tabs th |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
141 |
val deps = |
|
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50624
diff
changeset
|
142 |
isar_or_prover_dependencies_of ctxt params_opt facts name_tabs th |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
143 |
(SOME isar_deps) |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
144 |
val core = |
| 50826 | 145 |
encode_str name ^ ": " ^ encode_strs prevs ^ "; " ^ |
| 50582 | 146 |
encode_features (sort_wrt fst feats) |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
147 |
val query = |
|
50754
74a6adcb96ac
also generate queries for goals with too many Isar dependencies
blanchet
parents:
50735
diff
changeset
|
148 |
if is_bad_query ctxt ho_atp th isar_deps then "" |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
149 |
else "? " ^ core ^ "\n" |
|
50754
74a6adcb96ac
also generate queries for goals with too many Isar dependencies
blanchet
parents:
50735
diff
changeset
|
150 |
val update = |
|
74a6adcb96ac
also generate queries for goals with too many Isar dependencies
blanchet
parents:
50735
diff
changeset
|
151 |
"! " ^ core ^ "; " ^ |
| 50826 | 152 |
encode_strs (these (trim_dependencies th deps)) ^ "\n" |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
153 |
in query ^ update end |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
154 |
else |
|
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
155 |
"" |
|
50624
4d0997abce79
improved thm order hack, in case the default names are overridden
blanchet
parents:
50611
diff
changeset
|
156 |
val parents = |
|
4d0997abce79
improved thm order hack, in case the default names are overridden
blanchet
parents:
50611
diff
changeset
|
157 |
map (nickname_of_thm o snd) (the_list (try List.last old_facts)) |
|
4d0997abce79
improved thm order hack, in case the default names are overridden
blanchet
parents:
50611
diff
changeset
|
158 |
val new_facts = new_facts |> map (`(nickname_of_thm o snd)) |
| 50519 | 159 |
val prevss = fst (split_last (parents :: map (single o fst) new_facts)) |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
160 |
val lines = Par_List.map do_fact (tag_list 1 (new_facts ~~ prevss)) |
| 50519 | 161 |
in File.write_list path lines end |
|
48239
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
162 |
|
| 50411 | 163 |
fun generate_isar_commands ctxt prover = |
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50523
diff
changeset
|
164 |
generate_isar_or_prover_commands ctxt prover NONE (1, NONE) |
| 50411 | 165 |
|
|
50484
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50442
diff
changeset
|
166 |
fun generate_prover_commands ctxt (params as {provers = prover :: _, ...}) =
|
|
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50442
diff
changeset
|
167 |
generate_isar_or_prover_commands ctxt prover (SOME params) |
| 50411 | 168 |
|
| 50906 | 169 |
fun generate_mepo_suggestions ctxt (params as {provers = prover :: _, ...})
|
170 |
range thys max_suggs file_name = |
|
|
48239
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
171 |
let |
|
50515
c4a27ab89c9b
shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents:
50511
diff
changeset
|
172 |
val ho_atp = Sledgehammer_Provers.is_ho_atp ctxt prover |
|
48239
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
173 |
val path = file_name |> Path.explode |
| 50349 | 174 |
val facts = all_facts ctxt |
|
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50484
diff
changeset
|
175 |
val (new_facts, old_facts) = facts |> List.partition (has_thys thys o snd) |
|
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50624
diff
changeset
|
176 |
val name_tabs = build_name_tables nickname_of_thm facts |
| 50814 | 177 |
fun do_fact (j, ((_, th), old_facts)) = |
| 50906 | 178 |
if in_range range j then |
179 |
let |
|
180 |
val name = nickname_of_thm th |
|
181 |
val _ = tracing ("Fact " ^ string_of_int j ^ ": " ^ name)
|
|
182 |
val goal = goal_of_thm (Proof_Context.theory_of ctxt) th |
|
183 |
val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1 |
|
184 |
val isar_deps = isar_dependencies_of name_tabs th |
|
185 |
in |
|
186 |
if is_bad_query ctxt ho_atp th isar_deps then |
|
187 |
"" |
|
188 |
else |
|
189 |
let |
|
190 |
val suggs = |
|
191 |
old_facts |
|
192 |
|> Sledgehammer_MePo.mepo_suggested_facts ctxt params prover |
|
193 |
max_suggs NONE hyp_ts concl_t |
|
194 |
|> map (nickname_of_thm o snd) |
|
195 |
in encode_str name ^ ": " ^ encode_strs suggs ^ "\n" end |
|
196 |
end |
|
197 |
else |
|
198 |
"" |
|
| 50519 | 199 |
fun accum x (yss as ys :: _) = (x :: ys) :: yss |
200 |
val old_factss = tl (fold accum new_facts [old_facts]) |
|
| 50814 | 201 |
val lines = Par_List.map do_fact (tag_list 1 (new_facts ~~ rev old_factss)) |
| 50519 | 202 |
in File.write_list path lines end |
| 48234 | 203 |
|
| 50814 | 204 |
fun generate_mesh_suggestions max_suggs mash_file_name mepo_file_name |
205 |
mesh_file_name = |
|
206 |
let |
|
207 |
val mesh_path = Path.explode mesh_file_name |
|
208 |
val _ = File.write mesh_path "" |
|
209 |
fun do_fact (mash_line, mepo_line) = |
|
210 |
let |
|
| 50829 | 211 |
val (name, mash_suggs) = |
| 50814 | 212 |
extract_suggestions mash_line |
213 |
||> (map fst #> weight_mash_facts) |
|
| 50829 | 214 |
val (name', mepo_suggs) = |
| 50814 | 215 |
extract_suggestions mepo_line |
216 |
||> (map fst #> weight_mash_facts) |
|
| 50829 | 217 |
val _ = if name = name' then () else error "Input files out of sync." |
| 50814 | 218 |
val mess = |
219 |
[(mepo_weight, (mepo_suggs, [])), |
|
220 |
(mash_weight, (mash_suggs, []))] |
|
221 |
val mesh_suggs = mesh_facts (op =) max_suggs mess |
|
| 50829 | 222 |
val mesh_line = encode_str name ^ ": " ^ encode_strs mesh_suggs ^ "\n" |
| 50814 | 223 |
in File.append mesh_path mesh_line end |
224 |
val mash_lines = Path.explode mash_file_name |> File.read_lines |
|
225 |
val mepo_lines = Path.explode mepo_file_name |> File.read_lines |
|
| 50907 | 226 |
in |
227 |
if length mash_lines = length mepo_lines then |
|
228 |
List.app do_fact (mash_lines ~~ mepo_lines) |
|
229 |
else |
|
230 |
warning "Skipped: MaSh file missing or out of sync with MePo file." |
|
231 |
end |
|
| 50814 | 232 |
|
| 48234 | 233 |
end; |