author | sultana |
Wed, 04 Apr 2012 21:57:39 +0100 | |
changeset 47366 | f5a304ca037e |
parent 47360 | d1ecc9cec531 |
child 47411 | 7df9a4f320a5 |
permissions | -rw-r--r-- |
46879 | 1 |
(* Title: HOL/TPTP/TPTP_Parser/tptp_interpret.ML |
46844 | 2 |
Author: Nik Sultana, Cambridge University Computer Laboratory |
3 |
||
4 |
Interprets TPTP problems in Isabelle/HOL. |
|
5 |
*) |
|
6 |
||
7 |
signature TPTP_INTERPRET = |
|
8 |
sig |
|
9 |
(*signature extension: typing information for variables and constants |
|
10 |
(note that the former isn't really part of the signature, but it's bundled |
|
11 |
along for more configurability -- though it's probably useless and could be |
|
12 |
dropped in the future.*) |
|
13 |
type const_map = (string * term) list |
|
14 |
type var_types = (string * typ option) list |
|
15 |
||
47360 | 16 |
(*mapping from TPTP types to Isabelle/HOL types. This map works over all |
17 |
base types (i.e. TPTP functions must be interpreted as Isabelle/HOL functions). |
|
18 |
The map must be total wrt TPTP types. Later on there could be a configuration |
|
46844 | 19 |
option to make a map extensible.*) |
20 |
type type_map = (TPTP_Syntax.tptp_type * typ) list |
|
21 |
||
22 |
(*inference info, when available, consists of the name of the inference rule |
|
23 |
and the names of the inferences involved in the reasoning step.*) |
|
24 |
type inference_info = (string * string list) option |
|
25 |
||
26 |
(*a parsed annotated formula is represented as a triple consisting of |
|
27 |
the formula's label, its role, and a constant function to its Isabelle/HOL |
|
28 |
term meaning*) |
|
29 |
type tptp_formula_meaning = |
|
30 |
string * TPTP_Syntax.role * TPTP_Syntax.tptp_formula * term * inference_info |
|
31 |
||
32 |
(*In general, the meaning of a TPTP statement (which, through the Include |
|
33 |
directive, may include the meaning of an entire TPTP file, is an extended |
|
34 |
Isabelle/HOL theory, an explicit map from constants to their Isabelle/HOL |
|
35 |
counterparts and their types, the meaning of any TPTP formulas encountered |
|
47360 | 36 |
while interpreting that statement, and a map from TPTP to Isabelle/HOL types |
46844 | 37 |
(these types would have been added to the theory returned in the first position |
38 |
of the tuple). The last value is NONE when the function which produced the |
|
39 |
whole tptp_general_meaning value was given a type_map argument -- since |
|
40 |
this means that the type_map is already known, and it has not been changed.*) |
|
41 |
type tptp_general_meaning = |
|
42 |
(type_map * const_map * tptp_formula_meaning list) |
|
43 |
||
44 |
(*if problem_name is given then it is used to qualify types & constants*) |
|
45 |
type config = |
|
46 |
{(*init_type_map : type_map with_origin, |
|
47 |
init_const_map : type_map with_origin,*) |
|
48 |
cautious : bool, |
|
49 |
problem_name : TPTP_Problem_Name.problem_name option |
|
50 |
(*dont_build_maps : bool, |
|
51 |
extend_given_type_map : bool, |
|
52 |
extend_given_const_map : bool, |
|
53 |
generative_type_interpretation : bool, |
|
54 |
generative_const_interpretation : bool*)} |
|
55 |
||
47360 | 56 |
(*map types form TPTP to Isabelle/HOL*) |
46844 | 57 |
val interpret_type : config -> theory -> type_map -> |
58 |
TPTP_Syntax.tptp_type -> typ |
|
59 |
||
60 |
(*interpret a TPTP line: return an updated theory including the |
|
61 |
types & constants which were specified in that formula, a map from |
|
62 |
constant names to their types, and a map from constant names to Isabelle/HOL |
|
63 |
constants; and a list possible formulas resulting from that line. |
|
64 |
Note that type/constant declarations do not result in any formulas being |
|
65 |
returned. A typical TPTP line might update the theory, or return a single |
|
66 |
formula. The sole exception is the "include" directive which may update the |
|
67 |
theory and also return a list of formulas from the included file. |
|
68 |
Arguments: |
|
69 |
cautious = indicates whether additional checks are made to check |
|
70 |
that all used types have been declared. |
|
47360 | 71 |
type_map = mapping of TPTP-types to Isabelle/HOL types. This argument may be |
46844 | 72 |
given to force a specific mapping: this is usually done for using an |
47360 | 73 |
imported TPTP problem in Isar. |
46844 | 74 |
const_map = as previous, but for constants. |
47360 | 75 |
path_prefix = path where TPTP problems etc are located (to help the "include" |
46844 | 76 |
directive find the files. |
77 |
thy = theory where interpreted info will be stored. |
|
78 |
*) |
|
79 |
||
80 |
(*map terms form TPTP to Isabelle/HOL*) |
|
81 |
val interpret_term : bool -> config -> TPTP_Syntax.language -> theory -> |
|
82 |
const_map -> var_types -> type_map -> TPTP_Syntax.tptp_term -> |
|
83 |
term * theory |
|
84 |
||
85 |
val interpret_formula : config -> TPTP_Syntax.language -> const_map -> |
|
86 |
var_types -> type_map -> TPTP_Syntax.tptp_formula -> theory -> |
|
87 |
term * theory |
|
88 |
||
89 |
val interpret_line : config -> type_map -> const_map -> |
|
90 |
Path.T -> TPTP_Syntax.tptp_line -> theory -> tptp_general_meaning * theory |
|
91 |
||
92 |
(*Like "interpret_line" above, but works over a whole parsed problem. |
|
93 |
Arguments: |
|
94 |
new_basic_types = indicates whether interpretations of $i and $o |
|
95 |
types are to be added to the type map (unless it is Given). |
|
47360 | 96 |
This is "true" if we are running this over a fresh TPTP problem, but |
97 |
"false" if we are calling this _during_ the interpretation of a TPTP file |
|
46844 | 98 |
(i.e. when interpreting an "include" directive. |
99 |
config = config |
|
100 |
path_prefix = " " |
|
101 |
contents = parsed TPTP syntax |
|
102 |
type_map = see "interpret_line" |
|
103 |
const_map = " " |
|
104 |
thy = " " |
|
105 |
*) |
|
106 |
val interpret_problem : bool -> config -> Path.T -> |
|
107 |
TPTP_Syntax.tptp_line list -> type_map -> const_map -> theory -> |
|
108 |
tptp_general_meaning * theory |
|
109 |
||
110 |
(*Like "interpret_problem" above, but it's given a filename rather than |
|
111 |
a parsed problem.*) |
|
112 |
val interpret_file : bool -> Path.T -> Path.T -> type_map -> const_map -> |
|
113 |
theory -> tptp_general_meaning * theory |
|
114 |
||
115 |
(*General exception for this signature.*) |
|
116 |
exception MISINTERPRET_FORMULA of string * TPTP_Syntax.tptp_formula |
|
117 |
exception MISINTERPRET_SYMBOL of string * TPTP_Syntax.symbol |
|
118 |
exception MISINTERPRET_TERM of string * TPTP_Syntax.tptp_term |
|
119 |
exception MISINTERPRET_TYPE of string * TPTP_Syntax.tptp_type |
|
120 |
||
47360 | 121 |
(*Generates a fresh Isabelle/HOL type for interpreting a TPTP type in a theory.*) |
46844 | 122 |
val declare_type : config -> (TPTP_Syntax.tptp_type * string) -> type_map -> |
123 |
theory -> type_map * theory |
|
124 |
||
125 |
(*Returns the list of all files included in a directory and its |
|
126 |
subdirectories. This is only used for testing the parser/interpreter against |
|
47360 | 127 |
all TPTP problems.*) |
46844 | 128 |
val get_file_list : Path.T -> Path.T list |
129 |
||
130 |
type manifest = TPTP_Problem_Name.problem_name * tptp_general_meaning |
|
131 |
||
132 |
val get_manifests : theory -> manifest list |
|
133 |
||
134 |
val import_file : bool -> Path.T -> Path.T -> type_map -> const_map -> |
|
135 |
theory -> theory |
|
136 |
||
137 |
val extract_inference_info : (TPTP_Syntax.general_term * 'a list) option -> |
|
138 |
(string * string list) option |
|
139 |
end |
|
140 |
||
141 |
structure TPTP_Interpret : TPTP_INTERPRET = |
|
142 |
struct |
|
143 |
||
144 |
open TPTP_Syntax |
|
145 |
exception MISINTERPRET_FORMULA of string * TPTP_Syntax.tptp_formula |
|
146 |
exception MISINTERPRET_SYMBOL of string * TPTP_Syntax.symbol |
|
147 |
exception MISINTERPRET_TERM of string * TPTP_Syntax.tptp_term |
|
148 |
exception MISINTERPRET_TYPE of string * TPTP_Syntax.tptp_type |
|
149 |
||
150 |
(* General stuff *) |
|
151 |
||
152 |
type config = |
|
153 |
{(*init_type_map : type_map with_origin, |
|
154 |
init_const_map : type_map with_origin,*) |
|
155 |
cautious : bool, |
|
156 |
problem_name : TPTP_Problem_Name.problem_name option |
|
157 |
(*dont_build_maps : bool, |
|
158 |
extend_given_type_map : bool, |
|
159 |
extend_given_const_map : bool, |
|
160 |
generative_type_interpretation : bool, |
|
161 |
generative_const_interpretation : bool*)} |
|
162 |
||
163 |
(* Interpretation *) |
|
164 |
||
165 |
(** Signatures and other type abbrevations **) |
|
166 |
||
167 |
type const_map = (string * term) list |
|
168 |
type var_types = (string * typ option) list |
|
169 |
type type_map = (TPTP_Syntax.tptp_type * typ) list |
|
170 |
type inference_info = (string * string list) option |
|
171 |
type tptp_formula_meaning = |
|
172 |
string * TPTP_Syntax.role * TPTP_Syntax.tptp_formula * term * inference_info |
|
173 |
type tptp_general_meaning = |
|
174 |
(type_map * const_map * tptp_formula_meaning list) |
|
175 |
||
176 |
type manifest = TPTP_Problem_Name.problem_name * tptp_general_meaning |
|
177 |
structure TPTP_Data = Theory_Data |
|
178 |
(type T = manifest list |
|
179 |
val empty = [] |
|
180 |
val extend = I |
|
47332
360e080fd13e
dealing with SMLNJ errors 'value type in structure doesn't match signature spec' (which originated from lack of type generalisation) and 'unresolved flex record' wrt tptp_interpret
sultana
parents:
46961
diff
changeset
|
181 |
(*SMLNJ complains if non-expanded form used below*) |
360e080fd13e
dealing with SMLNJ errors 'value type in structure doesn't match signature spec' (which originated from lack of type generalisation) and 'unresolved flex record' wrt tptp_interpret
sultana
parents:
46961
diff
changeset
|
182 |
fun merge v = Library.merge (op =) v) |
46844 | 183 |
val get_manifests = TPTP_Data.get |
184 |
||
185 |
||
186 |
(** Adding types to a signature **) |
|
187 |
||
188 |
fun type_exists thy ty_name = |
|
189 |
Sign.declared_tyname thy (Sign.full_bname thy ty_name) |
|
190 |
||
191 |
val IND_TYPE = "ind" |
|
192 |
||
47332
360e080fd13e
dealing with SMLNJ errors 'value type in structure doesn't match signature spec' (which originated from lack of type generalisation) and 'unresolved flex record' wrt tptp_interpret
sultana
parents:
46961
diff
changeset
|
193 |
(*SMLNJ complains if type annotation not used below*) |
360e080fd13e
dealing with SMLNJ errors 'value type in structure doesn't match signature spec' (which originated from lack of type generalisation) and 'unresolved flex record' wrt tptp_interpret
sultana
parents:
46961
diff
changeset
|
194 |
fun mk_binding (config : config) ident = |
46844 | 195 |
let |
196 |
val pre_binding = Binding.name ident |
|
197 |
in |
|
198 |
case #problem_name config of |
|
199 |
NONE => pre_binding |
|
200 |
| SOME prob => |
|
201 |
Binding.qualify |
|
202 |
false |
|
203 |
(TPTP_Problem_Name.mangle_problem_name prob) |
|
204 |
pre_binding |
|
205 |
end |
|
206 |
||
207 |
(*Returns updated theory and the name of the final type's name -- i.e. if the |
|
208 |
original name is already taken then the function looks for an available |
|
209 |
alternative. It also returns an updated type_map if one has been passed to it.*) |
|
210 |
fun declare_type (config : config) (thf_type, type_name) ty_map thy = |
|
211 |
if type_exists thy type_name then |
|
212 |
raise MISINTERPRET_TYPE ("Type already exists", Atom_type type_name) |
|
213 |
else |
|
214 |
let |
|
215 |
val binding = mk_binding config type_name |
|
216 |
val final_fqn = Sign.full_name thy binding |
|
217 |
val thy' = Typedecl.typedecl_global (mk_binding config type_name, [], NoSyn) thy |> snd |
|
218 |
val typ_map_entry = (thf_type, (Type (final_fqn, []))) |
|
219 |
val ty_map' = typ_map_entry :: ty_map |
|
220 |
in (ty_map', thy') end |
|
221 |
||
222 |
||
223 |
(** Adding constants to signature **) |
|
224 |
||
225 |
fun full_name thy config const_name = |
|
226 |
Sign.full_name thy (mk_binding config const_name) |
|
227 |
||
228 |
fun const_exists config thy const_name = |
|
229 |
Sign.declared_const thy (full_name thy config const_name) |
|
230 |
||
231 |
fun declare_constant config const_name ty thy = |
|
232 |
if const_exists config thy const_name then |
|
233 |
raise MISINTERPRET_TERM |
|
234 |
("Const already declared", Term_Func (Uninterpreted const_name, [])) |
|
235 |
else |
|
236 |
Theory.specify_const |
|
237 |
((mk_binding config const_name, ty), NoSyn) thy |
|
238 |
||
239 |
||
240 |
(** Interpretation **) |
|
241 |
||
242 |
(*Types in THF are encoded as formulas. This function translates them to type form.*) |
|
243 |
(*FIXME possibly incomplete hack*) |
|
244 |
fun fmlatype_to_type (Atom (THF_Atom_term (Term_Func (TypeSymbol typ, [])))) = |
|
245 |
Defined_type typ |
|
246 |
| fmlatype_to_type (Atom (THF_Atom_term (Term_Func (Uninterpreted str, [])))) = |
|
247 |
Atom_type str |
|
47360 | 248 |
| fmlatype_to_type (Type_fmla (Fn_type (ty1, ty2))) = |
46844 | 249 |
let |
250 |
val ty1' = case ty1 of |
|
47360 | 251 |
Fn_type _ => fmlatype_to_type (Type_fmla ty1) |
46844 | 252 |
| Fmla_type ty1 => fmlatype_to_type ty1 |
253 |
val ty2' = case ty2 of |
|
47360 | 254 |
Fn_type _ => fmlatype_to_type (Type_fmla ty2) |
46844 | 255 |
| Fmla_type ty2 => fmlatype_to_type ty2 |
256 |
in Fn_type (ty1', ty2') end |
|
257 |
||
258 |
fun interpret_type config thy type_map thf_ty = |
|
259 |
let |
|
260 |
fun lookup_type ty_map thf_ty = |
|
261 |
(case AList.lookup (op =) ty_map thf_ty of |
|
262 |
NONE => |
|
263 |
raise MISINTERPRET_TYPE |
|
264 |
("Could not find the interpretation of this TPTP type in the " ^ |
|
265 |
"mapping to Isabelle/HOL", thf_ty) |
|
266 |
| SOME ty => ty) |
|
267 |
in |
|
268 |
case thf_ty of (*FIXME make tail*) |
|
269 |
Prod_type (thf_ty1, thf_ty2) => |
|
270 |
Type ("Product_Type.prod", |
|
271 |
[interpret_type config thy type_map thf_ty1, |
|
272 |
interpret_type config thy type_map thf_ty2]) |
|
273 |
| Fn_type (thf_ty1, thf_ty2) => |
|
274 |
Type ("fun", |
|
275 |
[interpret_type config thy type_map thf_ty1, |
|
276 |
interpret_type config thy type_map thf_ty2]) |
|
277 |
| Atom_type _ => |
|
278 |
lookup_type type_map thf_ty |
|
279 |
| Defined_type tptp_base_type => |
|
280 |
(case tptp_base_type of |
|
281 |
Type_Ind => |
|
282 |
lookup_type type_map thf_ty |
|
283 |
| Type_Bool => HOLogic.boolT |
|
284 |
| Type_Type => |
|
285 |
raise MISINTERPRET_TYPE |
|
286 |
("No type interpretation", thf_ty) |
|
287 |
| Type_Int => Type ("Int.int", []) |
|
288 |
(*FIXME these types are currently unsupported, so they're treated as |
|
289 |
individuals*) |
|
290 |
| Type_Rat => |
|
291 |
interpret_type config thy type_map (Defined_type Type_Ind) |
|
292 |
| Type_Real => |
|
293 |
interpret_type config thy type_map (Defined_type Type_Ind) |
|
294 |
) |
|
295 |
| Sum_type _ => |
|
296 |
raise MISINTERPRET_TYPE (*FIXME*) |
|
297 |
("No type interpretation (sum type)", thf_ty) |
|
298 |
| Fmla_type tptp_ty => |
|
299 |
fmlatype_to_type tptp_ty |
|
300 |
|> interpret_type config thy type_map |
|
301 |
| Subtype _ => |
|
302 |
raise MISINTERPRET_TYPE (*FIXME*) |
|
303 |
("No type interpretation (subtype)", thf_ty) |
|
304 |
end |
|
305 |
||
306 |
fun the_const config thy language const_map_payload str = |
|
307 |
if language = THF then |
|
308 |
(case AList.lookup (op =) const_map_payload str of |
|
309 |
NONE => raise MISINTERPRET_TERM |
|
310 |
("Could not find the interpretation of this constant in the " ^ |
|
311 |
"mapping to Isabelle/HOL", Term_Func (Uninterpreted str, [])) |
|
312 |
| SOME t => t) |
|
313 |
else |
|
314 |
if const_exists config thy str then |
|
315 |
Sign.mk_const thy ((Sign.full_name thy (mk_binding config str)), []) |
|
316 |
else raise MISINTERPRET_TERM |
|
317 |
("Could not find the interpretation of this constant in the " ^ |
|
318 |
"mapping to Isabelle/HOL: ", Term_Func (Uninterpreted str, [])) |
|
319 |
||
320 |
(*Eta-expands Isabelle/HOL binary function. |
|
321 |
"str" is the name of a polymorphic constant in Isabelle/HOL*) |
|
322 |
fun mk_fun str = |
|
323 |
(Const (str, Term.dummyT) $ Bound 1 $ Bound 0) |
|
324 |
|> (Term.absdummy Term.dummyT #> Term.absdummy Term.dummyT) |
|
325 |
(*As above, but swaps the function's arguments*) |
|
326 |
fun mk_swapped_fun str = |
|
327 |
(Const (str, Term.dummyT) $ Bound 0 $ Bound 1) |
|
328 |
|> (Term.absdummy Term.dummyT #> Term.absdummy Term.dummyT) |
|
329 |
||
47360 | 330 |
fun dummy_term () = |
46844 | 331 |
HOLogic.choice_const Term.dummyT $ Term.absdummy Term.dummyT @{term "True"} |
332 |
||
333 |
fun interpret_symbol config thy language const_map symb = |
|
334 |
case symb of |
|
335 |
Uninterpreted n => |
|
336 |
(*Constant would have been added to the map at the time of |
|
337 |
declaration*) |
|
338 |
the_const config thy language const_map n |
|
339 |
| Interpreted_ExtraLogic interpreted_symbol => |
|
340 |
(case interpreted_symbol of (*FIXME incomplete, |
|
341 |
and doesn't work with $i*) |
|
342 |
Sum => mk_fun @{const_name Groups.plus} |
|
343 |
| UMinus => |
|
344 |
(Const (@{const_name Groups.uminus}, Term.dummyT) $ Bound 0) |
|
345 |
|> Term.absdummy Term.dummyT |
|
346 |
| Difference => mk_fun @{const_name Groups.minus} |
|
347 |
| Product => mk_fun @{const_name Groups.times} |
|
348 |
| Quotient => mk_fun @{const_name Fields.divide} |
|
349 |
| Less => mk_fun @{const_name Orderings.less} |
|
350 |
| LessEq => mk_fun @{const_name Orderings.less_eq} |
|
351 |
| Greater => mk_swapped_fun @{const_name Orderings.less} |
|
352 |
(*FIXME would be nicer to expand "greater" |
|
353 |
and "greater_eq" abbreviations*) |
|
354 |
| GreaterEq => mk_swapped_fun @{const_name Orderings.less_eq} |
|
355 |
(* FIXME todo |
|
356 |
| Quotient_E | Quotient_T | Quotient_F | Remainder_E | Remainder_T |
|
357 |
| Remainder_F | Floor | Ceiling | Truncate | Round | To_Int | To_Rat |
|
358 |
| To_Real | EvalEq | Is_Int | Is_Rat*) |
|
359 |
| Apply => raise MISINTERPRET_SYMBOL ("Malformed term?", symb) |
|
47360 | 360 |
| _ => dummy_term () |
46844 | 361 |
) |
362 |
| Interpreted_Logic logic_symbol => |
|
363 |
(case logic_symbol of |
|
364 |
Equals => HOLogic.eq_const Term.dummyT |
|
365 |
| NEquals => |
|
366 |
HOLogic.mk_not (HOLogic.eq_const Term.dummyT $ Bound 1 $ Bound 0) |
|
367 |
|> (Term.absdummy Term.dummyT #> Term.absdummy Term.dummyT) |
|
368 |
| Or => HOLogic.disj |
|
369 |
| And => HOLogic.conj |
|
370 |
| Iff => HOLogic.eq_const HOLogic.boolT |
|
371 |
| If => HOLogic.imp |
|
372 |
| Fi => @{term "\<lambda> x. \<lambda> y. y \<longrightarrow> x"} |
|
373 |
| Xor => |
|
374 |
@{term |
|
375 |
"\<lambda> x. \<lambda> y. \<not> (x \<longleftrightarrow> y)"} |
|
376 |
| Nor => @{term "\<lambda> x. \<lambda> y. \<not> (x \<or> y)"} |
|
377 |
| Nand => @{term "\<lambda> x. \<lambda> y. \<not> (x \<and> y)"} |
|
378 |
| Not => HOLogic.Not |
|
379 |
| Op_Forall => HOLogic.all_const Term.dummyT |
|
380 |
| Op_Exists => HOLogic.exists_const Term.dummyT |
|
381 |
| True => @{term "True"} |
|
382 |
| False => @{term "False"} |
|
383 |
) |
|
384 |
| TypeSymbol _ => |
|
385 |
raise MISINTERPRET_SYMBOL |
|
386 |
("Cannot have TypeSymbol in term", symb) |
|
387 |
| System str => |
|
388 |
raise MISINTERPRET_SYMBOL |
|
389 |
("How to interpret system terms?", symb) |
|
390 |
||
391 |
(*Apply a function to a list of arguments*) |
|
392 |
val mapply = fold (fn x => fn y => y $ x) |
|
393 |
(*As above, but for products*) |
|
394 |
fun mtimes thy = |
|
395 |
fold (fn x => fn y => |
|
396 |
Sign.mk_const thy |
|
397 |
("Product_Type.Pair",[Term.dummyT, Term.dummyT]) $ y $ x) |
|
398 |
||
399 |
(*Adds constants to signature in FOL. "formula_level" means that the constants |
|
400 |
are to be interpreted as predicates, otherwise as functions over individuals.*) |
|
401 |
fun FO_const_types config formula_level type_map tptp_t thy = |
|
402 |
let |
|
403 |
val ind_type = interpret_type config thy type_map (Defined_type Type_Ind) |
|
404 |
val value_type = |
|
405 |
if formula_level then |
|
406 |
interpret_type config thy type_map (Defined_type Type_Bool) |
|
407 |
else ind_type |
|
408 |
in case tptp_t of |
|
409 |
Term_Func (symb, tptp_ts) => |
|
410 |
let |
|
411 |
val thy' = fold (FO_const_types config false type_map) tptp_ts thy |
|
412 |
val ty = fold (fn ty1 => fn ty2 => Type ("fun", [ty1, ty2])) |
|
413 |
(map (fn _ => ind_type) tptp_ts) value_type |
|
414 |
in |
|
415 |
case symb of |
|
416 |
Uninterpreted const_name => |
|
417 |
if const_exists config thy' const_name then thy' |
|
418 |
else snd (declare_constant config const_name ty thy') |
|
419 |
| _ => thy' |
|
420 |
end |
|
421 |
| _ => thy |
|
422 |
end |
|
423 |
||
424 |
(*like FO_const_types but works over symbols*) |
|
425 |
fun symb_const_types config type_map formula_level const_name n_args thy = |
|
426 |
let |
|
427 |
val ind_type = interpret_type config thy type_map (Defined_type Type_Ind) |
|
428 |
val value_type = |
|
429 |
if formula_level then |
|
430 |
interpret_type config thy type_map (Defined_type Type_Bool) |
|
431 |
else interpret_type config thy type_map (Defined_type Type_Ind) |
|
432 |
fun mk_i_fun b n acc = |
|
433 |
if n = 0 then acc b |
|
434 |
else |
|
435 |
mk_i_fun b (n - 1) (fn ty => Type ("fun", [ind_type, acc ty])) |
|
436 |
in |
|
437 |
if const_exists config thy const_name then thy |
|
438 |
else |
|
439 |
snd (declare_constant config const_name |
|
440 |
(mk_i_fun value_type n_args I) thy) |
|
441 |
end |
|
442 |
||
443 |
(*Next batch of functions give info about Isabelle/HOL types*) |
|
444 |
fun is_fun (Type ("fun", _)) = true |
|
445 |
| is_fun _ = false |
|
446 |
fun is_prod (Type ("Product_Type.prod", _)) = true |
|
447 |
| is_prod _ = false |
|
448 |
fun dom_type (Type ("fun", [ty1, _])) = ty1 |
|
449 |
fun is_prod_typed thy config symb = |
|
450 |
let |
|
451 |
fun symb_type const_name = |
|
452 |
Sign.the_const_type thy (full_name thy config const_name) |
|
453 |
in |
|
454 |
case symb of |
|
455 |
Uninterpreted const_name => |
|
456 |
if is_fun (symb_type const_name) then |
|
457 |
symb_type const_name |> dom_type |> is_prod |
|
458 |
else false |
|
459 |
| _ => false |
|
460 |
end |
|
461 |
||
462 |
||
463 |
(* |
|
464 |
Information needed to be carried around: |
|
465 |
- constant mapping: maps constant names to Isabelle terms with fully-qualified |
|
466 |
names. This is fixed, and it is determined by declaration lines earlier |
|
467 |
in the script (i.e. constants must be declared before appearing in terms. |
|
468 |
- type context: maps bound variables to their Isabelle type. This is discarded |
|
469 |
after each call of interpret_term since variables' can't be bound across |
|
470 |
terms. |
|
471 |
*) |
|
472 |
fun interpret_term formula_level config language thy const_map var_types type_map tptp_t = |
|
473 |
case tptp_t of |
|
474 |
Term_Func (symb, tptp_ts) => |
|
475 |
let |
|
476 |
val thy' = FO_const_types config formula_level type_map tptp_t thy |
|
477 |
fun typeof_const const_name = Sign.the_const_type thy' |
|
478 |
(Sign.full_name thy' (mk_binding config const_name)) |
|
479 |
in |
|
480 |
case symb of |
|
481 |
Interpreted_ExtraLogic Apply => |
|
482 |
(*apply the head of the argument list to the tail*) |
|
483 |
(mapply |
|
484 |
(map (interpret_term false config language thy' const_map |
|
485 |
var_types type_map #> fst) (tl tptp_ts)) |
|
486 |
(interpret_term formula_level config language thy' const_map |
|
487 |
var_types type_map (hd tptp_ts) |> fst), |
|
488 |
thy') |
|
489 |
| _ => |
|
490 |
(*apply symb to tptp_ts*) |
|
491 |
if is_prod_typed thy' config symb then |
|
492 |
(interpret_symbol config thy' language const_map symb $ |
|
493 |
mtimes thy' |
|
494 |
(map (interpret_term false config language thy' const_map |
|
495 |
var_types type_map #> fst) (tl tptp_ts)) |
|
496 |
((interpret_term false config language thy' const_map |
|
497 |
var_types type_map #> fst) (hd tptp_ts)), |
|
498 |
thy') |
|
499 |
else |
|
500 |
(mapply |
|
501 |
(map (interpret_term false config language thy' const_map |
|
502 |
var_types type_map #> fst) tptp_ts) |
|
503 |
(interpret_symbol config thy' language const_map symb), |
|
504 |
thy') |
|
505 |
end |
|
506 |
| Term_Var n => |
|
507 |
(if language = THF orelse language = TFF then |
|
508 |
(case AList.lookup (op =) var_types n of |
|
509 |
SOME ty => |
|
510 |
(case ty of |
|
511 |
SOME ty => Free (n, ty) |
|
512 |
| NONE => Free (n, Term.dummyT) (*to infer the variable's type*) |
|
513 |
) |
|
514 |
| NONE => |
|
515 |
raise MISINTERPRET_TERM ("Could not type variable", tptp_t)) |
|
516 |
else (*variables range over individuals*) |
|
517 |
Free (n, interpret_type config thy type_map (Defined_type Type_Ind)), |
|
518 |
thy) |
|
519 |
| Term_Conditional (tptp_formula, tptp_t1, tptp_t2) => |
|
520 |
(mk_fun @{const_name If}, thy) |
|
521 |
| Term_Num (number_kind, num) => |
|
522 |
let |
|
523 |
val num_term = |
|
524 |
if number_kind = Int_num then |
|
525 |
HOLogic.mk_number @{typ "int"} (the (Int.fromString num)) |
|
47360 | 526 |
else dummy_term () (*FIXME: not yet supporting rationals and reals*) |
46844 | 527 |
in (num_term, thy) end |
528 |
| Term_Distinct_Object str => |
|
529 |
let |
|
530 |
fun alphanumerate c = |
|
531 |
let |
|
532 |
val c' = ord c |
|
533 |
val out_of_range = |
|
534 |
((c' > 64) andalso (c' < 91)) orelse ((c' > 96) |
|
535 |
andalso (c' < 123)) orelse ((c' > 47) andalso (c' < 58)) |
|
536 |
in |
|
537 |
if (not out_of_range) andalso (not (c = "_")) then |
|
538 |
"_nom_" ^ Int.toString (ord c) ^ "_" |
|
539 |
else c |
|
540 |
end |
|
541 |
val mangle_name = raw_explode #> map alphanumerate #> implode |
|
542 |
in declare_constant config (mangle_name str) |
|
543 |
(interpret_type config thy type_map (Defined_type Type_Ind)) thy |
|
544 |
end |
|
545 |
||
546 |
(*Given a list of "'a option", then applies a function to each element if that |
|
547 |
element is SOME value, otherwise it leaves it as NONE.*) |
|
548 |
fun map_opt f xs = |
|
549 |
fold |
|
550 |
(fn x => fn y => |
|
551 |
(if Option.isSome x then |
|
552 |
SOME (f (Option.valOf x)) |
|
553 |
else NONE) :: y |
|
554 |
) xs [] |
|
555 |
||
556 |
fun interpret_formula config language const_map var_types type_map tptp_fmla thy = |
|
557 |
case tptp_fmla of |
|
558 |
Pred app => |
|
559 |
interpret_term true config language thy const_map |
|
560 |
var_types type_map (Term_Func app) |
|
561 |
| Fmla (symbol, tptp_formulas) => |
|
562 |
(case symbol of |
|
563 |
Interpreted_ExtraLogic Apply => |
|
564 |
let |
|
565 |
val (args, thy') = (fold_map (interpret_formula config language |
|
566 |
const_map var_types type_map) (tl tptp_formulas) thy) |
|
567 |
val (func, thy') = interpret_formula config language const_map |
|
568 |
var_types type_map (hd tptp_formulas) thy' |
|
569 |
in (mapply args func, thy') end |
|
570 |
| Uninterpreted const_name => |
|
571 |
let |
|
572 |
val (args, thy') = (fold_map (interpret_formula config language |
|
573 |
const_map var_types type_map) tptp_formulas thy) |
|
574 |
val thy' = symb_const_types config type_map true const_name |
|
575 |
(length tptp_formulas) thy' |
|
576 |
in |
|
577 |
(if is_prod_typed thy' config symbol then |
|
578 |
mtimes thy' args (interpret_symbol config thy' language |
|
579 |
const_map symbol) |
|
580 |
else |
|
581 |
mapply args |
|
582 |
(interpret_symbol config thy' language const_map symbol), |
|
583 |
thy') |
|
584 |
end |
|
585 |
| _ => |
|
586 |
let |
|
587 |
val (args, thy') = |
|
588 |
fold_map |
|
589 |
(interpret_formula config language |
|
590 |
const_map var_types type_map) |
|
591 |
tptp_formulas thy |
|
592 |
in |
|
593 |
(if is_prod_typed thy' config symbol then |
|
594 |
mtimes thy' args (interpret_symbol config thy' language |
|
595 |
const_map symbol) |
|
596 |
else |
|
597 |
mapply args |
|
598 |
(interpret_symbol config thy' language const_map symbol), |
|
599 |
thy') |
|
600 |
end) |
|
601 |
| Sequent _ => |
|
602 |
raise MISINTERPRET_FORMULA ("Sequent", tptp_fmla) (*FIXME unsupported*) |
|
603 |
| Quant (quantifier, bindings, tptp_formula) => |
|
604 |
let |
|
605 |
val (bound_vars, bound_var_types) = ListPair.unzip bindings |
|
606 |
val bound_var_types' : typ option list = |
|
607 |
(map_opt : (tptp_type -> typ) -> (tptp_type option list) -> typ option list) |
|
608 |
(interpret_type config thy type_map) bound_var_types |
|
609 |
val var_types' = |
|
610 |
ListPair.zip (bound_vars, List.rev bound_var_types') |
|
611 |
|> List.rev |
|
612 |
fun fold_bind f = |
|
613 |
fold |
|
614 |
(fn (n, ty) => fn t => |
|
615 |
case ty of |
|
616 |
NONE => |
|
617 |
f (n, |
|
618 |
if language = THF then Term.dummyT |
|
619 |
else |
|
620 |
interpret_type config thy type_map |
|
621 |
(Defined_type Type_Ind), |
|
622 |
t) |
|
623 |
| SOME ty => f (n, ty, t) |
|
624 |
) |
|
625 |
var_types' |
|
626 |
in case quantifier of |
|
627 |
Forall => |
|
628 |
interpret_formula config language const_map (var_types' @ var_types) |
|
629 |
type_map tptp_formula thy |
|
630 |
|>> fold_bind HOLogic.mk_all |
|
631 |
| Exists => |
|
632 |
interpret_formula config language const_map (var_types' @ var_types) |
|
633 |
type_map tptp_formula thy |
|
634 |
|>> fold_bind HOLogic.mk_exists |
|
635 |
| Lambda => |
|
636 |
interpret_formula config language const_map (var_types' @ var_types) |
|
637 |
type_map tptp_formula thy |
|
638 |
|>> fold_bind (fn (n, ty, rest) => lambda (Free (n, ty)) rest) |
|
639 |
| Epsilon => |
|
640 |
let val (t, thy') = |
|
641 |
interpret_formula config language const_map var_types type_map |
|
642 |
(Quant (Lambda, bindings, tptp_formula)) thy |
|
643 |
in ((HOLogic.choice_const Term.dummyT) $ t, thy') end |
|
644 |
| Iota => |
|
645 |
let val (t, thy') = |
|
646 |
interpret_formula config language const_map var_types type_map |
|
647 |
(Quant (Lambda, bindings, tptp_formula)) thy |
|
648 |
in (Const (@{const_name The}, Term.dummyT) $ t, thy') end |
|
649 |
| Dep_Prod => |
|
650 |
raise MISINTERPRET_FORMULA ("Unsupported", tptp_fmla) |
|
651 |
| Dep_Sum => |
|
652 |
raise MISINTERPRET_FORMULA ("Unsupported", tptp_fmla) |
|
653 |
end |
|
47359 | 654 |
| Conditional (fmla1, fmla2, fmla3) => |
655 |
let |
|
656 |
val interp = interpret_formula config language const_map var_types type_map |
|
657 |
val (((fmla1', fmla2'), fmla3'), thy') = |
|
658 |
interp fmla1 thy |
|
659 |
||>> interp fmla2 |
|
660 |
||>> interp fmla3 |
|
661 |
in (HOLogic.mk_conj (HOLogic.mk_imp (fmla1', fmla2'), |
|
662 |
HOLogic.mk_imp (HOLogic.mk_not fmla1', fmla3')), |
|
663 |
thy') |
|
664 |
end |
|
665 |
| Let (tptp_let_list, tptp_formula) => (*FIXME not yet implemented*) |
|
666 |
raise MISINTERPRET_FORMULA ("'Let' not yet implemented", tptp_fmla) |
|
46844 | 667 |
| Atom tptp_atom => |
668 |
(case tptp_atom of |
|
669 |
TFF_Typed_Atom (symbol, tptp_type_opt) => |
|
670 |
(*FIXME ignoring type info*) |
|
671 |
(interpret_symbol config thy language const_map symbol, thy) |
|
672 |
| THF_Atom_term tptp_term => |
|
673 |
interpret_term true config language thy const_map var_types |
|
674 |
type_map tptp_term |
|
675 |
| THF_Atom_conn_term symbol => |
|
676 |
(interpret_symbol config thy language const_map symbol, thy)) |
|
47360 | 677 |
| Type_fmla _ => (*FIXME unsupported*) |
46844 | 678 |
raise MISINTERPRET_FORMULA |
679 |
("Cannot interpret types as formulas", tptp_fmla) |
|
680 |
| THF_typing (tptp_formula, _) => (*FIXME ignoring type info*) |
|
681 |
interpret_formula config language |
|
682 |
const_map var_types type_map tptp_formula thy |
|
683 |
||
684 |
||
685 |
(*Extract name of inference rule, and the inferences it relies on*) |
|
686 |
(*This is tuned to work with LEO-II, and is brittle to upstream |
|
687 |
changes of the proof protocol.*) |
|
688 |
fun extract_inference_info annot = |
|
689 |
let |
|
690 |
fun get_line_id (General_Data (Number (Int_num, num))) = [num] |
|
691 |
| get_line_id (General_Data (Atomic_Word name)) = [name] |
|
692 |
| get_line_id (General_Term (Number (Int_num, num), _ (*e.g. a bind*))) = [num] |
|
693 |
| get_line_id _ = [] |
|
694 |
(*e.g. General_Data (Application ("theory", [General_Data |
|
695 |
(Atomic_Word "equality")])) -- which would come from E through LEO-II*) |
|
696 |
in |
|
697 |
case annot of |
|
698 |
NONE => NONE |
|
699 |
| SOME annot => |
|
700 |
(case annot of |
|
701 |
(General_Data (Application ("inference", |
|
702 |
[General_Data (Atomic_Word inference_name), |
|
703 |
_, |
|
704 |
General_List related_lines])), _) => |
|
705 |
(SOME (inference_name, map get_line_id related_lines |> List.concat)) |
|
706 |
| _ => NONE) |
|
707 |
end |
|
708 |
||
709 |
||
710 |
(*Extract the type from a typing*) |
|
711 |
local |
|
712 |
fun extract_type tptp_type = |
|
713 |
case tptp_type of |
|
714 |
Fmla_type typ => fmlatype_to_type typ |
|
715 |
| _ => tptp_type |
|
716 |
in |
|
717 |
fun typeof_typing (THF_typing (_, tptp_type)) = extract_type tptp_type |
|
718 |
fun typeof_tff_typing (Atom (TFF_Typed_Atom (_, SOME tptp_type))) = |
|
719 |
extract_type tptp_type |
|
720 |
end |
|
721 |
fun nameof_typing |
|
722 |
(THF_typing |
|
723 |
((Atom (THF_Atom_term (Term_Func (Uninterpreted str, [])))), _)) = str |
|
724 |
fun nameof_tff_typing (Atom (TFF_Typed_Atom (Uninterpreted str, _))) = str |
|
725 |
||
726 |
fun interpret_line config type_map const_map path_prefix line thy = |
|
727 |
case line of |
|
728 |
Include (quoted_path, _) => (*FIXME currently everything is imported*) |
|
729 |
interpret_file' |
|
730 |
false |
|
731 |
config |
|
732 |
path_prefix |
|
733 |
(Path.append |
|
734 |
path_prefix |
|
735 |
(quoted_path |
|
736 |
|> unenclose |
|
737 |
|> Path.explode)) |
|
738 |
type_map |
|
739 |
const_map |
|
740 |
thy |
|
741 |
| Annotated_Formula (pos, lang, label, role, tptp_formula, annotation_opt) => |
|
742 |
case role of |
|
743 |
Role_Type => |
|
744 |
let |
|
745 |
val thy_name = Context.theory_name thy |
|
746 |
val (tptp_ty, name) = |
|
747 |
if lang = THF then |
|
748 |
(typeof_typing tptp_formula (*assuming tptp_formula is a typing*), |
|
749 |
nameof_typing tptp_formula (*and that the LHS is a (atom) name*)) |
|
750 |
else |
|
751 |
(typeof_tff_typing tptp_formula, |
|
752 |
nameof_tff_typing tptp_formula) |
|
753 |
in |
|
754 |
case tptp_ty of |
|
755 |
Defined_type Type_Type => (*add new type*) |
|
47360 | 756 |
(*generate an Isabelle/HOL type to interpret this TPTP type, |
46844 | 757 |
and add it to both the Isabelle/HOL theory and to the type_map*) |
758 |
let |
|
759 |
val (type_map', thy') = |
|
760 |
declare_type |
|
761 |
config |
|
762 |
(Atom_type name, name) |
|
763 |
type_map |
|
764 |
thy |
|
765 |
in ((type_map', |
|
766 |
const_map, |
|
767 |
[]), |
|
768 |
thy') |
|
769 |
end |
|
770 |
||
771 |
| _ => (*declaration of constant*) |
|
772 |
(*Here we populate the map from constants to the Isabelle/HOL |
|
773 |
terms they map to (which in turn contain their types).*) |
|
774 |
let |
|
775 |
val ty = interpret_type config thy type_map tptp_ty |
|
776 |
(*make sure that the theory contains all the types appearing |
|
777 |
in this constant's signature. Exception is thrown if encounter |
|
778 |
an undeclared types.*) |
|
779 |
val (t, thy') = |
|
780 |
let |
|
781 |
fun analyse_type thy thf_ty = |
|
782 |
if #cautious config then |
|
783 |
case thf_ty of |
|
784 |
Fn_type (thf_ty1, thf_ty2) => |
|
785 |
(analyse_type thy thf_ty1; |
|
786 |
analyse_type thy thf_ty2) |
|
787 |
| Atom_type ty_name => |
|
788 |
if type_exists thy ty_name then () |
|
789 |
else |
|
790 |
raise MISINTERPRET_TYPE |
|
791 |
("This type, in signature of " ^ |
|
792 |
name ^ " has not been declared.", |
|
793 |
Atom_type ty_name) |
|
794 |
| _ => () |
|
795 |
else () (*skip test if we're not being cautious.*) |
|
796 |
in |
|
797 |
analyse_type thy tptp_ty; |
|
798 |
declare_constant config name ty thy |
|
799 |
end |
|
800 |
(*register a mapping from name to t. Constants' type |
|
801 |
declarations should occur at most once, so it's justified to |
|
802 |
use "::". Note that here we use a constant's name rather |
|
803 |
than the possibly- new one, since all references in the |
|
804 |
theory will be to this name.*) |
|
805 |
val const_map' = ((name, t) :: const_map) |
|
806 |
in ((type_map,(*type_map unchanged, since a constant's |
|
807 |
declaration shouldn't include any new types.*) |
|
808 |
const_map',(*typing table of constant was extended*) |
|
809 |
[]),(*no formulas were interpreted*) |
|
810 |
thy')(*theory was extended with new constant*) |
|
811 |
end |
|
812 |
end |
|
813 |
||
814 |
| _ => (*i.e. the AF is not a type declaration*) |
|
815 |
let |
|
816 |
(*gather interpreted term, and the map of types occurring in that term*) |
|
817 |
val (t, thy') = |
|
818 |
interpret_formula config lang |
|
819 |
const_map [] type_map tptp_formula thy |
|
820 |
|>> HOLogic.mk_Trueprop |
|
821 |
(*type_maps grow monotonically, so return the newest value (type_mapped); |
|
822 |
there's no need to unify it with the type_map parameter.*) |
|
823 |
in |
|
824 |
((type_map, const_map, |
|
825 |
[(label, role, tptp_formula, |
|
826 |
Syntax.check_term (Proof_Context.init_global thy') t, |
|
827 |
extract_inference_info annotation_opt)]), thy') |
|
828 |
end |
|
829 |
||
830 |
and interpret_problem new_basic_types config path_prefix lines type_map const_map thy = |
|
831 |
let |
|
832 |
val thy_name = Context.theory_name thy |
|
833 |
(*Add interpretation of $o and generate an Isabelle/HOL type to |
|
834 |
interpret $i, unless we've been given a mapping of types.*) |
|
835 |
val (thy', type_map') = |
|
836 |
if not new_basic_types then (thy, type_map) |
|
837 |
else |
|
838 |
let |
|
839 |
val (type_map', thy') = |
|
840 |
declare_type config |
|
841 |
(Defined_type Type_Ind, IND_TYPE) type_map thy |
|
842 |
in |
|
843 |
(thy', type_map') |
|
844 |
end |
|
845 |
in |
|
846 |
fold (fn line => |
|
847 |
fn ((type_map, const_map, lines), thy) => |
|
848 |
let |
|
849 |
(*process the line, ignoring the type-context for variables*) |
|
850 |
val ((type_map', const_map', l'), thy') = |
|
851 |
interpret_line config type_map const_map path_prefix line thy |
|
852 |
in |
|
853 |
((type_map', |
|
854 |
const_map', |
|
855 |
l' @ lines),(*append is sufficient, union would be overkill*) |
|
856 |
thy') |
|
857 |
end) |
|
858 |
lines (*lines of the problem file*) |
|
859 |
((type_map', const_map, []), thy') (*initial values*) |
|
860 |
end |
|
861 |
||
862 |
and interpret_file' new_basic_types config path_prefix file_name = |
|
863 |
let |
|
864 |
val file_name' = Path.expand file_name |
|
865 |
in |
|
866 |
if Path.is_absolute file_name' then |
|
867 |
Path.implode file_name |
|
868 |
|> TPTP_Parser.parse_file |
|
869 |
|> interpret_problem new_basic_types config path_prefix |
|
870 |
else error "Could not determine absolute path to file." |
|
871 |
end |
|
872 |
||
873 |
and interpret_file cautious path_prefix file_name = |
|
874 |
let |
|
875 |
val config = |
|
876 |
{cautious = cautious, |
|
877 |
problem_name = |
|
878 |
SOME (TPTP_Problem_Name.parse_problem_name ((Path.base #> Path.implode) |
|
879 |
file_name)) |
|
880 |
} |
|
881 |
handle _(*FIXME*) => |
|
882 |
{cautious = cautious, |
|
883 |
problem_name = NONE |
|
884 |
} |
|
885 |
in interpret_file' true config path_prefix file_name end |
|
886 |
||
887 |
fun import_file cautious path_prefix file_name type_map const_map thy = |
|
888 |
let |
|
889 |
val prob_name = |
|
890 |
TPTP_Problem_Name.parse_problem_name (Path.implode (Path.base file_name)) |
|
891 |
handle _(*FIXME*) => TPTP_Problem_Name.Nonstandard (Path.implode file_name) |
|
892 |
val (result, thy') = |
|
893 |
interpret_file cautious path_prefix file_name type_map const_map thy |
|
894 |
in TPTP_Data.map (cons ((prob_name, result))) thy' end |
|
895 |
||
896 |
val _ = |
|
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
46951
diff
changeset
|
897 |
Outer_Syntax.improper_command @{command_spec "import_tptp"} "import TPTP problem" |
46951 | 898 |
(Parse.path >> (fn path => |
47366
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
sultana
parents:
47360
diff
changeset
|
899 |
Toplevel.theory (fn thy => |
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
sultana
parents:
47360
diff
changeset
|
900 |
(*NOTE: assumes Axioms directory is on same level as the Problems one |
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
sultana
parents:
47360
diff
changeset
|
901 |
(which is how TPTP is currently organised)*) |
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
sultana
parents:
47360
diff
changeset
|
902 |
import_file true (Path.appends [Path.dir path, Path.parent, Path.parent]) |
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
sultana
parents:
47360
diff
changeset
|
903 |
path [] [] thy))) |
46844 | 904 |
|
905 |
||
906 |
(*Used for debugging. Returns all files contained within a directory or its |
|
907 |
subdirectories. Follows symbolic links, filters away directories.*) |
|
908 |
fun get_file_list path = |
|
909 |
let |
|
910 |
fun check_file_entry f rest = |
|
911 |
let |
|
912 |
(*NOTE needed since don't have File.is_link and File.read_link*) |
|
913 |
val f_str = Path.implode f |
|
914 |
in |
|
915 |
if File.is_dir f then |
|
916 |
rest (*filter out directories*) |
|
917 |
else if OS.FileSys.isLink f_str then |
|
918 |
(*follow links -- NOTE this breaks if links are relative paths*) |
|
919 |
check_file_entry (Path.explode (OS.FileSys.readLink f_str)) rest |
|
920 |
else f :: rest |
|
921 |
end |
|
922 |
in |
|
923 |
File.read_dir path |
|
924 |
|> map |
|
925 |
(Path.explode |
|
926 |
#> Path.append path) |
|
927 |
|> (fn l => fold check_file_entry l []) |
|
928 |
end |
|
929 |
||
930 |
end |