| author | wenzelm | 
| Fri, 27 Nov 2015 18:47:39 +0100 | |
| changeset 61750 | c6c2508f94b8 | 
| parent 61259 | 6dc3d5d50e57 | 
| child 62354 | fdd6989cc8a0 | 
| 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  | 
|
| 47515 | 9  | 
(*Signature extension: typing information for variables and constants*)  | 
| 46844 | 10  | 
type var_types = (string * typ option) list  | 
| 57808 | 11  | 
type const_map = (string * (term * int list)) list  | 
| 46844 | 12  | 
|
| 47515 | 13  | 
(*Mapping from TPTP types to Isabelle/HOL types. This map works over all  | 
14  | 
base types. The map must be total wrt TPTP types.*)  | 
|
| 57808 | 15  | 
type type_map = (string * (string * int)) list  | 
| 46844 | 16  | 
|
| 47515 | 17  | 
(*A parsed annotated formula is represented as a 5-tuple consisting of  | 
18  | 
the formula's label, its role, the TPTP formula, its Isabelle/HOL meaning, and  | 
|
19  | 
inference info*)  | 
|
| 46844 | 20  | 
type tptp_formula_meaning =  | 
| 
53388
 
c878390475f3
using richer annotation from formula annotations in proof;
 
sultana 
parents: 
49323 
diff
changeset
 | 
21  | 
string * TPTP_Syntax.role * term * TPTP_Proof.source_info option  | 
| 46844 | 22  | 
|
23  | 
(*In general, the meaning of a TPTP statement (which, through the Include  | 
|
| 47515 | 24  | 
directive, may include the meaning of an entire TPTP file, is a map from  | 
25  | 
TPTP to Isabelle/HOL types, a map from TPTP constants to their Isabelle/HOL  | 
|
26  | 
counterparts and their types, and a list of interpreted annotated formulas.*)  | 
|
| 46844 | 27  | 
type tptp_general_meaning =  | 
28  | 
(type_map * const_map * tptp_formula_meaning list)  | 
|
29  | 
||
| 47515 | 30  | 
(*cautious = indicates whether additional checks are made to check  | 
31  | 
that all used types have been declared.  | 
|
32  | 
problem_name = if given then it is used to qualify types & constants*)  | 
|
| 46844 | 33  | 
type config =  | 
| 47515 | 34  | 
    {cautious : bool,
 | 
| 47548 | 35  | 
problem_name : TPTP_Problem_Name.problem_name option}  | 
| 46844 | 36  | 
|
| 47515 | 37  | 
(*Generates a fresh Isabelle/HOL type for interpreting a TPTP type in a theory.*)  | 
| 57808 | 38  | 
val declare_type : config -> string * (string * int) -> type_map ->  | 
| 47515 | 39  | 
theory -> type_map * theory  | 
40  | 
||
41  | 
(*Map TPTP types to Isabelle/HOL types*)  | 
|
| 46844 | 42  | 
val interpret_type : config -> theory -> type_map ->  | 
43  | 
TPTP_Syntax.tptp_type -> typ  | 
|
44  | 
||
| 47515 | 45  | 
(*Map TPTP terms to Isabelle/HOL terms*)  | 
| 
47570
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
46  | 
val interpret_term : bool -> config -> TPTP_Syntax.language ->  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
47  | 
const_map -> var_types -> type_map -> TPTP_Syntax.tptp_term -> theory ->  | 
| 47511 | 48  | 
term * theory  | 
49  | 
||
50  | 
val interpret_formula : config -> TPTP_Syntax.language -> const_map ->  | 
|
51  | 
var_types -> type_map -> TPTP_Syntax.tptp_formula -> theory ->  | 
|
52  | 
term * theory  | 
|
53  | 
||
| 47515 | 54  | 
(*Interpret a TPTP line: return a "tptp_general_meaning" for that line,  | 
55  | 
as well as an updated Isabelle theory including any types & constants  | 
|
56  | 
which were specified in that line.  | 
|
| 46844 | 57  | 
Note that type/constant declarations do not result in any formulas being  | 
| 47515 | 58  | 
returned. A typical TPTP line might update the theory, and/or return one or  | 
59  | 
more formulas. For instance, the "include" directive may update the theory  | 
|
60  | 
and also return a list of formulas from the included file.  | 
|
| 46844 | 61  | 
Arguments:  | 
| 47515 | 62  | 
config = (see above)  | 
63  | 
inclusion list = names of annotated formulas to interpret (since "include"  | 
|
64  | 
directive can be selective wrt to such names)  | 
|
| 47360 | 65  | 
type_map = mapping of TPTP-types to Isabelle/HOL types. This argument may be  | 
| 46844 | 66  | 
given to force a specific mapping: this is usually done for using an  | 
| 47360 | 67  | 
imported TPTP problem in Isar.  | 
| 46844 | 68  | 
const_map = as previous, but for constants.  | 
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
69  | 
path_prefixes = paths where TPTP problems etc are located (to help the  | 
| 
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
70  | 
"include" directive find the files.  | 
| 47515 | 71  | 
line = parsed TPTP line  | 
| 46844 | 72  | 
thy = theory where interpreted info will be stored.  | 
73  | 
*)  | 
|
| 47511 | 74  | 
val interpret_line : config -> string list -> type_map -> const_map ->  | 
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
75  | 
Path.T list -> TPTP_Syntax.tptp_line -> theory ->  | 
| 
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
76  | 
tptp_general_meaning * theory  | 
| 46844 | 77  | 
|
78  | 
(*Like "interpret_line" above, but works over a whole parsed problem.  | 
|
79  | 
Arguments:  | 
|
| 47515 | 80  | 
config = as previously  | 
81  | 
inclusion list = as previously  | 
|
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
82  | 
path_prefixes = as previously  | 
| 47515 | 83  | 
lines = parsed TPTP syntax  | 
84  | 
type_map = as previously  | 
|
85  | 
const_map = as previously  | 
|
86  | 
thy = as previously  | 
|
| 46844 | 87  | 
*)  | 
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
88  | 
val interpret_problem : config -> string list -> Path.T list ->  | 
| 46844 | 89  | 
TPTP_Syntax.tptp_line list -> type_map -> const_map -> theory ->  | 
90  | 
tptp_general_meaning * theory  | 
|
91  | 
||
| 47515 | 92  | 
(*Like "interpret_problem" above, but it is given a filename rather than  | 
| 46844 | 93  | 
a parsed problem.*)  | 
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
94  | 
val interpret_file : bool -> Path.T list -> Path.T -> type_map -> const_map ->  | 
| 46844 | 95  | 
theory -> tptp_general_meaning * theory  | 
96  | 
||
| 
47569
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
97  | 
type position = string * int * int  | 
| 
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
98  | 
exception MISINTERPRET of position list * exn  | 
| 46844 | 99  | 
exception MISINTERPRET_FORMULA of string * TPTP_Syntax.tptp_formula  | 
100  | 
exception MISINTERPRET_SYMBOL of string * TPTP_Syntax.symbol  | 
|
101  | 
exception MISINTERPRET_TERM of string * TPTP_Syntax.tptp_term  | 
|
102  | 
exception MISINTERPRET_TYPE of string * TPTP_Syntax.tptp_type  | 
|
103  | 
||
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
104  | 
val import_file : bool -> Path.T list -> Path.T -> type_map -> const_map ->  | 
| 47515 | 105  | 
theory -> theory  | 
106  | 
||
107  | 
(*Imported TPTP files are stored as "manifests" in the theory.*)  | 
|
108  | 
type manifest = TPTP_Problem_Name.problem_name * tptp_general_meaning  | 
|
109  | 
val get_manifests : theory -> manifest list  | 
|
| 46844 | 110  | 
end  | 
111  | 
||
112  | 
structure TPTP_Interpret : TPTP_INTERPRET =  | 
|
113  | 
struct  | 
|
114  | 
||
115  | 
open TPTP_Syntax  | 
|
| 
47569
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
116  | 
type position = string * int * int  | 
| 
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
117  | 
exception MISINTERPRET of position list * exn  | 
| 46844 | 118  | 
exception MISINTERPRET_FORMULA of string * TPTP_Syntax.tptp_formula  | 
119  | 
exception MISINTERPRET_SYMBOL of string * TPTP_Syntax.symbol  | 
|
120  | 
exception MISINTERPRET_TERM of string * TPTP_Syntax.tptp_term  | 
|
121  | 
exception MISINTERPRET_TYPE of string * TPTP_Syntax.tptp_type  | 
|
122  | 
||
| 47515 | 123  | 
|
| 46844 | 124  | 
(* General stuff *)  | 
125  | 
||
126  | 
type config =  | 
|
| 47515 | 127  | 
  {cautious : bool,
 | 
128  | 
problem_name : TPTP_Problem_Name.problem_name option}  | 
|
129  | 
||
| 46844 | 130  | 
|
131  | 
(* Interpretation *)  | 
|
132  | 
||
133  | 
(** Signatures and other type abbrevations **)  | 
|
134  | 
||
| 57808 | 135  | 
type const_map = (string * (term * int list)) list  | 
| 46844 | 136  | 
type var_types = (string * typ option) list  | 
| 57808 | 137  | 
type type_map = (string * (string * int)) list  | 
| 46844 | 138  | 
type tptp_formula_meaning =  | 
| 
53388
 
c878390475f3
using richer annotation from formula annotations in proof;
 
sultana 
parents: 
49323 
diff
changeset
 | 
139  | 
string * TPTP_Syntax.role * term * TPTP_Proof.source_info option  | 
| 46844 | 140  | 
type tptp_general_meaning =  | 
141  | 
(type_map * const_map * tptp_formula_meaning list)  | 
|
142  | 
||
143  | 
type manifest = TPTP_Problem_Name.problem_name * tptp_general_meaning  | 
|
| 47639 | 144  | 
|
| 46844 | 145  | 
structure TPTP_Data = Theory_Data  | 
| 47639 | 146  | 
(  | 
147  | 
type T = manifest list  | 
|
148  | 
val empty = []  | 
|
149  | 
val extend = I  | 
|
150  | 
fun merge data : T = Library.merge (op =) data  | 
|
151  | 
)  | 
|
| 46844 | 152  | 
val get_manifests = TPTP_Data.get  | 
153  | 
||
154  | 
||
155  | 
(** Adding types to a signature **)  | 
|
156  | 
||
| 47514 | 157  | 
(*transform quoted names into acceptable ASCII strings*)  | 
158  | 
fun alphanumerate c =  | 
|
159  | 
let  | 
|
160  | 
val c' = ord c  | 
|
161  | 
val out_of_range =  | 
|
162  | 
((c' > 64) andalso (c' < 91)) orelse ((c' > 96)  | 
|
163  | 
andalso (c' < 123)) orelse ((c' > 47) andalso (c' < 58))  | 
|
164  | 
in  | 
|
165  | 
if (not out_of_range) andalso (not (c = "_")) then  | 
|
166  | 
"_nom_" ^ Int.toString (ord c) ^ "_"  | 
|
167  | 
else c  | 
|
168  | 
end  | 
|
169  | 
fun alphanumerated_name prefix name =  | 
|
170  | 
prefix ^ (raw_explode #> map alphanumerate #> implode) name  | 
|
171  | 
||
| 
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
 | 
172  | 
(*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
 | 
173  | 
fun mk_binding (config : config) ident =  | 
| 46844 | 174  | 
let  | 
| 47514 | 175  | 
val pre_binding = Binding.name (alphanumerated_name "bnd_" ident)  | 
| 46844 | 176  | 
in  | 
177  | 
case #problem_name config of  | 
|
178  | 
NONE => pre_binding  | 
|
179  | 
| SOME prob =>  | 
|
180  | 
Binding.qualify  | 
|
181  | 
false  | 
|
182  | 
(TPTP_Problem_Name.mangle_problem_name prob)  | 
|
183  | 
pre_binding  | 
|
184  | 
end  | 
|
185  | 
||
| 47548 | 186  | 
fun type_exists config thy ty_name =  | 
187  | 
Sign.declared_tyname thy (Sign.full_name thy (mk_binding config ty_name))  | 
|
188  | 
||
| 46844 | 189  | 
(*Returns updated theory and the name of the final type's name -- i.e. if the  | 
190  | 
original name is already taken then the function looks for an available  | 
|
191  | 
alternative. It also returns an updated type_map if one has been passed to it.*)  | 
|
| 57808 | 192  | 
fun declare_type (config : config) (thf_type, (type_name, arity)) ty_map thy =  | 
| 47548 | 193  | 
if type_exists config thy type_name then  | 
| 57808 | 194  | 
    raise MISINTERPRET_TYPE ("Type already exists", Atom_type (type_name, []))
 | 
| 46844 | 195  | 
else  | 
196  | 
let  | 
|
197  | 
val binding = mk_binding config type_name  | 
|
198  | 
val final_fqn = Sign.full_name thy binding  | 
|
| 57808 | 199  | 
val tyargs =  | 
200  | 
        List.tabulate (arity, rpair @{sort type} o prefix "'" o string_of_int)
 | 
|
| 61259 | 201  | 
val (_, thy') =  | 
202  | 
        Typedecl.typedecl_global {final = true} (mk_binding config type_name, tyargs, NoSyn) thy
 | 
|
| 57808 | 203  | 
val typ_map_entry = (thf_type, (final_fqn, arity))  | 
| 46844 | 204  | 
val ty_map' = typ_map_entry :: ty_map  | 
205  | 
in (ty_map', thy') end  | 
|
206  | 
||
207  | 
||
208  | 
(** Adding constants to signature **)  | 
|
209  | 
||
210  | 
fun full_name thy config const_name =  | 
|
211  | 
Sign.full_name thy (mk_binding config const_name)  | 
|
212  | 
||
213  | 
fun const_exists config thy const_name =  | 
|
214  | 
Sign.declared_const thy (full_name thy config const_name)  | 
|
215  | 
||
216  | 
fun declare_constant config const_name ty thy =  | 
|
217  | 
if const_exists config thy const_name then  | 
|
218  | 
raise MISINTERPRET_TERM  | 
|
219  | 
     ("Const already declared", Term_Func (Uninterpreted const_name, []))
 | 
|
220  | 
else  | 
|
| 57808 | 221  | 
Theory.specify_const ((mk_binding config const_name, ty), NoSyn) thy  | 
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
222  | 
|
| 46844 | 223  | 
|
| 47515 | 224  | 
(** Interpretation functions **)  | 
| 46844 | 225  | 
|
| 57808 | 226  | 
(*Types in TFF/THF are encoded as formulas. These functions translate them to type form.*)  | 
227  | 
||
228  | 
fun termtype_to_type (Term_Func (TypeSymbol typ, [])) =  | 
|
229  | 
Defined_type typ  | 
|
230  | 
| termtype_to_type (Term_Func (Uninterpreted str, tms)) =  | 
|
231  | 
Atom_type (str, map termtype_to_type tms)  | 
|
232  | 
| termtype_to_type (Term_Var str) = Var_type str  | 
|
233  | 
||
| 46844 | 234  | 
(*FIXME possibly incomplete hack*)  | 
| 57808 | 235  | 
fun fmlatype_to_type (Atom (THF_Atom_term tm)) = termtype_to_type tm  | 
| 47360 | 236  | 
| fmlatype_to_type (Type_fmla (Fn_type (ty1, ty2))) =  | 
| 46844 | 237  | 
let  | 
238  | 
val ty1' = case ty1 of  | 
|
| 47360 | 239  | 
Fn_type _ => fmlatype_to_type (Type_fmla ty1)  | 
| 46844 | 240  | 
| Fmla_type ty1 => fmlatype_to_type ty1  | 
| 57808 | 241  | 
| _ => ty1  | 
| 46844 | 242  | 
val ty2' = case ty2 of  | 
| 47360 | 243  | 
Fn_type _ => fmlatype_to_type (Type_fmla ty2)  | 
| 46844 | 244  | 
| Fmla_type ty2 => fmlatype_to_type ty2  | 
| 57808 | 245  | 
| _ => ty2  | 
| 46844 | 246  | 
in Fn_type (ty1', ty2') end  | 
| 57808 | 247  | 
| fmlatype_to_type (Type_fmla ty) = ty  | 
248  | 
| fmlatype_to_type (Fmla (Uninterpreted str, fmla)) =  | 
|
249  | 
Atom_type (str, map fmlatype_to_type fmla)  | 
|
250  | 
| fmlatype_to_type (Quant (Dep_Prod, _, fmla)) = fmlatype_to_type fmla  | 
|
251  | 
| fmlatype_to_type (Pred (Interpreted_ExtraLogic Apply, [tm])) =  | 
|
252  | 
termtype_to_type tm  | 
|
253  | 
||
254  | 
fun tfree_name_of_var_type str = "'" ^ Name.desymbolize (SOME false) str  | 
|
255  | 
fun tfree_of_var_type str = TFree (tfree_name_of_var_type str, @{sort type})
 | 
|
| 46844 | 256  | 
|
257  | 
fun interpret_type config thy type_map thf_ty =  | 
|
258  | 
let  | 
|
| 57808 | 259  | 
fun lookup_type ty_map ary str =  | 
260  | 
(case AList.lookup (op =) ty_map str of  | 
|
| 46844 | 261  | 
NONE =>  | 
| 57808 | 262  | 
raise MISINTERPRET_SYMBOL  | 
| 47510 | 263  | 
              ("Could not find the interpretation of this TPTP type in the \
 | 
| 57808 | 264  | 
\mapping to Isabelle/HOL", Uninterpreted str)  | 
265  | 
| SOME (str', ary') =>  | 
|
266  | 
if ary' = ary then  | 
|
267  | 
str'  | 
|
268  | 
else  | 
|
269  | 
              raise MISINTERPRET_SYMBOL ("TPTP type used with wrong arity",
 | 
|
270  | 
Uninterpreted str))  | 
|
| 46844 | 271  | 
in  | 
| 47548 | 272  | 
case thf_ty of  | 
| 46844 | 273  | 
Prod_type (thf_ty1, thf_ty2) =>  | 
| 56256 | 274  | 
         Type (@{type_name prod},
 | 
| 46844 | 275  | 
[interpret_type config thy type_map thf_ty1,  | 
276  | 
interpret_type config thy type_map thf_ty2])  | 
|
277  | 
| Fn_type (thf_ty1, thf_ty2) =>  | 
|
| 56256 | 278  | 
         Type (@{type_name fun},
 | 
| 46844 | 279  | 
[interpret_type config thy type_map thf_ty1,  | 
280  | 
interpret_type config thy type_map thf_ty2])  | 
|
| 57808 | 281  | 
| Atom_type (str, thf_tys) =>  | 
282  | 
Type (lookup_type type_map (length thf_tys) str,  | 
|
283  | 
map (interpret_type config thy type_map) thf_tys)  | 
|
284  | 
| Var_type str => tfree_of_var_type str  | 
|
| 46844 | 285  | 
| Defined_type tptp_base_type =>  | 
286  | 
(case tptp_base_type of  | 
|
| 47519 | 287  | 
            Type_Ind => @{typ ind}
 | 
| 46844 | 288  | 
| Type_Bool => HOLogic.boolT  | 
| 47510 | 289  | 
          | Type_Type => raise MISINTERPRET_TYPE ("No type interpretation", thf_ty)
 | 
| 
60547
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
290  | 
(*FIXME these types are currently unsupported, so they're treated as  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
291  | 
individuals*)  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
292  | 
(*  | 
| 57796 | 293  | 
          | Type_Int => @{typ int}
 | 
294  | 
          | Type_Rat => @{typ rat}
 | 
|
| 57808 | 295  | 
          | Type_Real => @{typ real}
 | 
| 
60547
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
296  | 
*)  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
297  | 
| Type_Int =>  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
298  | 
interpret_type config thy type_map (Defined_type Type_Ind)  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
299  | 
| Type_Rat =>  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
300  | 
interpret_type config thy type_map (Defined_type Type_Ind)  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
301  | 
| Type_Real =>  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
302  | 
interpret_type config thy type_map (Defined_type Type_Ind)  | 
| 57808 | 303  | 
| Type_Dummy => dummyT)  | 
| 46844 | 304  | 
| Sum_type _ =>  | 
305  | 
raise MISINTERPRET_TYPE (*FIXME*)  | 
|
306  | 
          ("No type interpretation (sum type)", thf_ty)
 | 
|
307  | 
| Fmla_type tptp_ty =>  | 
|
308  | 
fmlatype_to_type tptp_ty  | 
|
309  | 
|> interpret_type config thy type_map  | 
|
310  | 
| Subtype _ =>  | 
|
311  | 
raise MISINTERPRET_TYPE (*FIXME*)  | 
|
312  | 
          ("No type interpretation (subtype)", thf_ty)
 | 
|
313  | 
end  | 
|
314  | 
||
| 57808 | 315  | 
fun permute_type_args perm Ts = map (nth Ts) perm  | 
316  | 
||
| 
60547
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
317  | 
fun the_const config thy const_map str tyargs =  | 
| 57808 | 318  | 
(case AList.lookup (op =) const_map str of  | 
319  | 
SOME (Const (s, _), tyarg_perm) =>  | 
|
320  | 
Sign.mk_const thy (s, permute_type_args tyarg_perm tyargs)  | 
|
| 
60547
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
321  | 
| _ =>  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
322  | 
if const_exists config thy str then  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
323  | 
Sign.mk_const thy ((Sign.full_name thy (mk_binding config str)), [])  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
324  | 
else  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
325  | 
raise MISINTERPRET_TERM  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
326  | 
            ("Could not find the interpretation of this constant in the \
 | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
327  | 
\mapping to Isabelle/HOL", Term_Func (Uninterpreted str, [])))  | 
| 46844 | 328  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
329  | 
(*Eta-expands n-ary function.  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
330  | 
"str" is the name of an Isabelle/HOL constant*)  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
331  | 
fun mk_n_fun n str =  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
332  | 
let  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
333  | 
fun body 0 t = t  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
334  | 
| body n t = body (n - 1) (t $ (Bound (n - 1)))  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
335  | 
in  | 
| 49323 | 336  | 
body n (Const (str, dummyT))  | 
337  | 
|> funpow n (Term.absdummy dummyT)  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
338  | 
end  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
339  | 
fun mk_fun_type [] b acc = acc b  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
340  | 
| mk_fun_type (ty :: tys) b acc =  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
341  | 
      mk_fun_type tys b (fn ty_rest => Type ("fun", [ty, acc ty_rest]))
 | 
| 46844 | 342  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
343  | 
(*These arities are not part of the TPTP spec*)  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
344  | 
fun arity_of interpreted_symbol = case interpreted_symbol of  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
345  | 
UMinus => 1  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
346  | 
| Sum => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
347  | 
| Difference => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
348  | 
| Product => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
349  | 
| Quotient => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
350  | 
| Quotient_E => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
351  | 
| Quotient_T => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
352  | 
| Quotient_F => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
353  | 
| Remainder_E => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
354  | 
| Remainder_T => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
355  | 
| Remainder_F => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
356  | 
| Floor => 1  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
357  | 
| Ceiling => 1  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
358  | 
| Truncate => 1  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
359  | 
| Round => 1  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
360  | 
| To_Int => 1  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
361  | 
| To_Rat => 1  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
362  | 
| To_Real => 1  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
363  | 
| Less => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
364  | 
| LessEq => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
365  | 
| Greater => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
366  | 
| GreaterEq => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
367  | 
| EvalEq => 2  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
368  | 
| Is_Int => 1  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
369  | 
| Is_Rat => 1  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
370  | 
| Distinct => 1  | 
| 57808 | 371  | 
| Apply => 2  | 
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
372  | 
|
| 57808 | 373  | 
fun type_arity_of_symbol thy config (Uninterpreted n) =  | 
374  | 
let val s = full_name thy config n in  | 
|
375  | 
length (Sign.const_typargs thy (s, Sign.the_const_type thy s))  | 
|
| 
60547
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
376  | 
handle TYPE _ => 0  | 
| 57808 | 377  | 
end  | 
378  | 
| type_arity_of_symbol _ _ _ = 0  | 
|
379  | 
||
380  | 
fun interpret_symbol config const_map symb tyargs thy =  | 
|
| 46844 | 381  | 
case symb of  | 
382  | 
Uninterpreted n =>  | 
|
383  | 
(*Constant would have been added to the map at the time of  | 
|
384  | 
declaration*)  | 
|
| 
60547
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
385  | 
the_const config thy const_map n tyargs  | 
| 46844 | 386  | 
| Interpreted_ExtraLogic interpreted_symbol =>  | 
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
387  | 
(*FIXME not interpreting*)  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
388  | 
Sign.mk_const thy ((Sign.full_name thy (mk_binding config  | 
| 57808 | 389  | 
(string_of_interpreted_symbol interpreted_symbol))), tyargs)  | 
| 46844 | 390  | 
| Interpreted_Logic logic_symbol =>  | 
391  | 
(case logic_symbol of  | 
|
| 49323 | 392  | 
Equals => HOLogic.eq_const dummyT  | 
| 46844 | 393  | 
| NEquals =>  | 
| 49323 | 394  | 
HOLogic.mk_not (HOLogic.eq_const dummyT $ Bound 1 $ Bound 0)  | 
395  | 
|> (Term.absdummy dummyT #> Term.absdummy dummyT)  | 
|
| 46844 | 396  | 
| Or => HOLogic.disj  | 
397  | 
| And => HOLogic.conj  | 
|
398  | 
| Iff => HOLogic.eq_const HOLogic.boolT  | 
|
399  | 
| If => HOLogic.imp  | 
|
400  | 
        | Fi => @{term "\<lambda> x. \<lambda> y. y \<longrightarrow> x"}
 | 
|
401  | 
| Xor =>  | 
|
402  | 
            @{term
 | 
|
403  | 
"\<lambda> x. \<lambda> y. \<not> (x \<longleftrightarrow> y)"}  | 
|
404  | 
        | Nor => @{term "\<lambda> x. \<lambda> y. \<not> (x \<or> y)"}
 | 
|
405  | 
        | Nand => @{term "\<lambda> x. \<lambda> y. \<not> (x \<and> y)"}
 | 
|
406  | 
| Not => HOLogic.Not  | 
|
| 49323 | 407  | 
| Op_Forall => HOLogic.all_const dummyT  | 
408  | 
| Op_Exists => HOLogic.exists_const dummyT  | 
|
| 46844 | 409  | 
        | True => @{term "True"}
 | 
410  | 
        | False => @{term "False"}
 | 
|
411  | 
)  | 
|
412  | 
| TypeSymbol _ =>  | 
|
413  | 
raise MISINTERPRET_SYMBOL  | 
|
414  | 
        ("Cannot have TypeSymbol in term", symb)
 | 
|
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
415  | 
| System _ =>  | 
| 46844 | 416  | 
raise MISINTERPRET_SYMBOL  | 
| 47510 | 417  | 
        ("Cannot interpret system symbol", symb)
 | 
| 46844 | 418  | 
|
419  | 
(*Apply a function to a list of arguments*)  | 
|
420  | 
val mapply = fold (fn x => fn y => y $ x)  | 
|
| 
47570
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
421  | 
|
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
422  | 
fun mapply' (args, thy) f =  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
423  | 
let  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
424  | 
val (f', thy') = f thy  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
425  | 
in (mapply args f', thy') end  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
426  | 
|
| 46844 | 427  | 
(*As above, but for products*)  | 
428  | 
fun mtimes thy =  | 
|
429  | 
fold (fn x => fn y =>  | 
|
| 56256 | 430  | 
    Sign.mk_const thy (@{const_name Pair}, [dummyT, dummyT]) $ y $ x)
 | 
| 46844 | 431  | 
|
| 
47570
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
432  | 
fun mtimes' (args, thy) f =  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
433  | 
let  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
434  | 
val (f', thy') = f thy  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
435  | 
in (mtimes thy' args f', thy') end  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
436  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
437  | 
datatype tptp_atom_value =  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
438  | 
Atom_App of tptp_term  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
439  | 
| Atom_Arity of string * int (*FIXME instead of int could use type list*)  | 
| 
47570
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
440  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
441  | 
(*Adds constants to signature when explicit type declaration isn't  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
442  | 
expected, e.g. FOL. "formula_level" means that the constants are to be interpreted  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
443  | 
as predicates, otherwise as functions over individuals.*)  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
444  | 
fun type_atoms_to_thy config formula_level type_map tptp_atom_value thy =  | 
| 46844 | 445  | 
let  | 
446  | 
val ind_type = interpret_type config thy type_map (Defined_type Type_Ind)  | 
|
447  | 
val value_type =  | 
|
448  | 
if formula_level then  | 
|
449  | 
interpret_type config thy type_map (Defined_type Type_Bool)  | 
|
450  | 
else ind_type  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
451  | 
in case tptp_atom_value of  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
452  | 
Atom_App (Term_Func (symb, tptp_ts)) =>  | 
| 46844 | 453  | 
let  | 
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
454  | 
val thy' = fold (fn t =>  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
455  | 
type_atoms_to_thy config false type_map (Atom_App t)) tptp_ts thy  | 
| 46844 | 456  | 
in  | 
457  | 
case symb of  | 
|
458  | 
Uninterpreted const_name =>  | 
|
| 57808 | 459  | 
perhaps (try (snd o declare_constant config const_name  | 
460  | 
(mk_fun_type (replicate (length tptp_ts) ind_type) value_type I))) thy'  | 
|
| 46844 | 461  | 
| _ => thy'  | 
462  | 
end  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
463  | 
| Atom_App _ => thy  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
464  | 
| Atom_Arity (const_name, n_args) =>  | 
| 57808 | 465  | 
perhaps (try (snd o declare_constant config const_name  | 
466  | 
(mk_fun_type (replicate n_args ind_type) value_type I))) thy  | 
|
| 46844 | 467  | 
end  | 
468  | 
||
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
469  | 
(*FIXME only used until interpretation is implemented*)  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
470  | 
fun add_interp_symbols_to_thy config type_map thy =  | 
| 46844 | 471  | 
let  | 
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
472  | 
val ind_symbols = [UMinus, Sum, Difference, Product, Quotient, Quotient_E,  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
473  | 
Quotient_T, Quotient_F, Remainder_E, Remainder_T, Remainder_F, Floor,  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
474  | 
Ceiling, Truncate, Round, To_Int, To_Rat, To_Real, Distinct, Apply]  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
475  | 
val bool_symbols = [Less, LessEq, Greater, GreaterEq, EvalEq, Is_Int, Is_Rat]  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
476  | 
fun add_interp_symbol_to_thy formula_level symbol =  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
477  | 
type_atoms_to_thy config formula_level type_map  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
478  | 
(Atom_Arity (string_of_interpreted_symbol symbol, arity_of symbol))  | 
| 46844 | 479  | 
in  | 
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
480  | 
fold (add_interp_symbol_to_thy false) ind_symbols thy  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
481  | 
|> fold (add_interp_symbol_to_thy true) bool_symbols  | 
| 46844 | 482  | 
end  | 
483  | 
||
484  | 
(*Next batch of functions give info about Isabelle/HOL types*)  | 
|
| 56256 | 485  | 
fun is_fun (Type (@{type_name fun}, _)) = true
 | 
| 46844 | 486  | 
| is_fun _ = false  | 
| 56256 | 487  | 
fun is_prod (Type (@{type_name prod}, _)) = true
 | 
| 46844 | 488  | 
| is_prod _ = false  | 
| 56256 | 489  | 
fun dom_type (Type (@{type_name fun}, [ty1, _])) = ty1
 | 
| 46844 | 490  | 
fun is_prod_typed thy config symb =  | 
491  | 
let  | 
|
492  | 
fun symb_type const_name =  | 
|
493  | 
Sign.the_const_type thy (full_name thy config const_name)  | 
|
494  | 
in  | 
|
495  | 
case symb of  | 
|
496  | 
Uninterpreted const_name =>  | 
|
497  | 
if is_fun (symb_type const_name) then  | 
|
498  | 
symb_type const_name |> dom_type |> is_prod  | 
|
499  | 
else false  | 
|
500  | 
| _ => false  | 
|
501  | 
end  | 
|
502  | 
||
503  | 
(*  | 
|
504  | 
Information needed to be carried around:  | 
|
505  | 
- constant mapping: maps constant names to Isabelle terms with fully-qualified  | 
|
506  | 
names. This is fixed, and it is determined by declaration lines earlier  | 
|
507  | 
in the script (i.e. constants must be declared before appearing in terms.  | 
|
508  | 
- type context: maps bound variables to their Isabelle type. This is discarded  | 
|
| 47515 | 509  | 
after each call of interpret_term since variables' cannot be bound across  | 
| 46844 | 510  | 
terms.  | 
511  | 
*)  | 
|
| 
47570
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
512  | 
fun interpret_term formula_level config language const_map var_types type_map  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
513  | 
tptp_t thy =  | 
| 46844 | 514  | 
case tptp_t of  | 
515  | 
Term_Func (symb, tptp_ts) =>  | 
|
516  | 
let  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
517  | 
val thy' =  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
518  | 
type_atoms_to_thy config formula_level type_map (Atom_App tptp_t) thy  | 
| 46844 | 519  | 
in  | 
520  | 
case symb of  | 
|
521  | 
Interpreted_ExtraLogic Apply =>  | 
|
522  | 
(*apply the head of the argument list to the tail*)  | 
|
| 
47570
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
523  | 
(mapply'  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
524  | 
(fold_map (interpret_term false config language const_map  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
525  | 
var_types type_map) (tl tptp_ts) thy')  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
526  | 
(interpret_term formula_level config language const_map  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
527  | 
var_types type_map (hd tptp_ts)))  | 
| 46844 | 528  | 
| _ =>  | 
| 57808 | 529  | 
let  | 
| 
60547
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
530  | 
val typ_arity = type_arity_of_symbol thy' config symb  | 
| 57808 | 531  | 
val (tptp_type_args, tptp_term_args) = chop typ_arity tptp_ts  | 
532  | 
val tyargs =  | 
|
| 
60547
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
533  | 
map (interpret_type config thy' type_map o termtype_to_type)  | 
| 57808 | 534  | 
tptp_type_args  | 
535  | 
in  | 
|
536  | 
(*apply symb to tptp_ts*)  | 
|
537  | 
if is_prod_typed thy' config symb then  | 
|
538  | 
let  | 
|
539  | 
val (t, thy'') =  | 
|
540  | 
mtimes'  | 
|
541  | 
(fold_map (interpret_term false config language const_map  | 
|
542  | 
var_types type_map) (tl tptp_term_args) thy')  | 
|
543  | 
(interpret_term false config language const_map  | 
|
544  | 
var_types type_map (hd tptp_term_args))  | 
|
545  | 
in (interpret_symbol config const_map symb tyargs thy' $ t, thy'')  | 
|
546  | 
end  | 
|
547  | 
else  | 
|
548  | 
(  | 
|
549  | 
mapply'  | 
|
550  | 
(fold_map (interpret_term false config language const_map  | 
|
551  | 
var_types type_map) tptp_term_args thy')  | 
|
552  | 
(`(interpret_symbol config const_map symb tyargs))  | 
|
553  | 
)  | 
|
554  | 
end  | 
|
| 46844 | 555  | 
end  | 
556  | 
| Term_Var n =>  | 
|
557  | 
(if language = THF orelse language = TFF then  | 
|
558  | 
(case AList.lookup (op =) var_types n of  | 
|
559  | 
SOME ty =>  | 
|
560  | 
(case ty of  | 
|
561  | 
SOME ty => Free (n, ty)  | 
|
| 49323 | 562  | 
| NONE => Free (n, dummyT) (*to infer the variable's type*)  | 
| 46844 | 563  | 
)  | 
564  | 
| NONE =>  | 
|
| 55593 | 565  | 
Free (n, dummyT)  | 
566  | 
             (*raise MISINTERPRET_TERM ("Could not type variable", tptp_t)*))
 | 
|
| 46844 | 567  | 
else (*variables range over individuals*)  | 
568  | 
Free (n, interpret_type config thy type_map (Defined_type Type_Ind)),  | 
|
569  | 
thy)  | 
|
570  | 
| Term_Conditional (tptp_formula, tptp_t1, tptp_t2) =>  | 
|
| 47691 | 571  | 
let  | 
572  | 
val (t_fmla, thy') =  | 
|
573  | 
interpret_formula config language const_map var_types type_map tptp_formula thy  | 
|
574  | 
val ([t1, t2], thy'') =  | 
|
575  | 
fold_map (interpret_term formula_level config language const_map var_types type_map)  | 
|
576  | 
[tptp_t1, tptp_t2] thy'  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
577  | 
      in (mk_n_fun 3 @{const_name If} $ t_fmla $ t1 $ t2, thy'') end
 | 
| 46844 | 578  | 
| Term_Num (number_kind, num) =>  | 
579  | 
let  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
580  | 
(*FIXME hack*)  | 
| 
60547
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
581  | 
(*  | 
| 57808 | 582  | 
val tptp_type = case number_kind of  | 
583  | 
Int_num => Type_Int  | 
|
584  | 
| Real_num => Type_Real  | 
|
585  | 
| Rat_num => Type_Rat  | 
|
586  | 
val T = interpret_type config thy type_map (Defined_type tptp_type)  | 
|
587  | 
in (HOLogic.mk_number T (the (Int.fromString num)), thy) end  | 
|
| 
60547
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
588  | 
*)  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
589  | 
|
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
590  | 
val ind_type = interpret_type config thy type_map (Defined_type Type_Ind)  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
591  | 
val prefix = case number_kind of  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
592  | 
Int_num => "intn_"  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
593  | 
| Real_num => "realn_"  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
594  | 
| Rat_num => "ratn_"  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
595  | 
val const_name = prefix ^ num  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
596  | 
in  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
597  | 
if const_exists config thy const_name then  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
598  | 
(Sign.mk_const thy ((Sign.full_name thy (mk_binding config const_name)), []), thy)  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
599  | 
else  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
600  | 
declare_constant config const_name ind_type thy  | 
| 
 
a62655f925c8
reverted some too aggressive TPTP interpreter changes
 
blanchet 
parents: 
59936 
diff
changeset
 | 
601  | 
end  | 
| 46844 | 602  | 
| Term_Distinct_Object str =>  | 
| 
47570
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
603  | 
      declare_constant config ("do_" ^ str)
 | 
| 47514 | 604  | 
(interpret_type config thy type_map (Defined_type Type_Ind)) thy  | 
| 46844 | 605  | 
|
| 47691 | 606  | 
and interpret_formula config language const_map var_types type_map tptp_fmla thy =  | 
| 46844 | 607  | 
case tptp_fmla of  | 
608  | 
Pred app =>  | 
|
| 
47570
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
609  | 
interpret_term true config language const_map  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
610  | 
var_types type_map (Term_Func app) thy  | 
| 46844 | 611  | 
| Fmla (symbol, tptp_formulas) =>  | 
612  | 
(case symbol of  | 
|
613  | 
Interpreted_ExtraLogic Apply =>  | 
|
| 
47570
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
614  | 
mapply'  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
615  | 
(fold_map (interpret_formula config language const_map  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
616  | 
var_types type_map) (tl tptp_formulas) thy)  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
617  | 
(interpret_formula config language const_map  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
618  | 
var_types type_map (hd tptp_formulas))  | 
| 46844 | 619  | 
| Uninterpreted const_name =>  | 
620  | 
let  | 
|
621  | 
val (args, thy') = (fold_map (interpret_formula config language  | 
|
622  | 
const_map var_types type_map) tptp_formulas thy)  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
623  | 
val thy' =  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
624  | 
type_atoms_to_thy config true type_map  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
625  | 
(Atom_Arity (const_name, length tptp_formulas)) thy'  | 
| 46844 | 626  | 
in  | 
627  | 
(if is_prod_typed thy' config symbol then  | 
|
| 57808 | 628  | 
mtimes thy' args (interpret_symbol config const_map symbol [] thy')  | 
| 46844 | 629  | 
else  | 
| 57808 | 630  | 
mapply args (interpret_symbol config const_map symbol [] thy'),  | 
| 46844 | 631  | 
thy')  | 
632  | 
end  | 
|
633  | 
| _ =>  | 
|
634  | 
let  | 
|
635  | 
val (args, thy') =  | 
|
636  | 
fold_map  | 
|
637  | 
(interpret_formula config language  | 
|
638  | 
const_map var_types type_map)  | 
|
639  | 
tptp_formulas thy  | 
|
640  | 
in  | 
|
641  | 
(if is_prod_typed thy' config symbol then  | 
|
| 57808 | 642  | 
mtimes thy' args (interpret_symbol config const_map symbol [] thy')  | 
| 46844 | 643  | 
else  | 
| 57808 | 644  | 
mapply args (interpret_symbol config const_map symbol [] thy'),  | 
| 46844 | 645  | 
thy')  | 
646  | 
end)  | 
|
647  | 
| Sequent _ =>  | 
|
| 47510 | 648  | 
(*FIXME unsupported*)  | 
649  | 
        raise MISINTERPRET_FORMULA ("'Sequent' unsupported", tptp_fmla)
 | 
|
| 46844 | 650  | 
| Quant (quantifier, bindings, tptp_formula) =>  | 
651  | 
let  | 
|
652  | 
val var_types' =  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
653  | 
ListPair.unzip bindings  | 
| 47729 | 654  | 
|> (apsnd ((map o Option.map) (interpret_type config thy type_map)))  | 
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
655  | 
|> ListPair.zip  | 
| 
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
656  | 
|> List.rev  | 
| 46844 | 657  | 
fun fold_bind f =  | 
658  | 
fold  | 
|
659  | 
(fn (n, ty) => fn t =>  | 
|
660  | 
case ty of  | 
|
661  | 
NONE =>  | 
|
662  | 
f (n,  | 
|
| 49323 | 663  | 
if language = THF then dummyT  | 
| 46844 | 664  | 
else  | 
665  | 
interpret_type config thy type_map  | 
|
666  | 
(Defined_type Type_Ind),  | 
|
667  | 
t)  | 
|
668  | 
| SOME ty => f (n, ty, t)  | 
|
669  | 
)  | 
|
670  | 
var_types'  | 
|
671  | 
in case quantifier of  | 
|
672  | 
Forall =>  | 
|
673  | 
interpret_formula config language const_map (var_types' @ var_types)  | 
|
674  | 
type_map tptp_formula thy  | 
|
675  | 
|>> fold_bind HOLogic.mk_all  | 
|
676  | 
| Exists =>  | 
|
677  | 
interpret_formula config language const_map (var_types' @ var_types)  | 
|
678  | 
type_map tptp_formula thy  | 
|
679  | 
|>> fold_bind HOLogic.mk_exists  | 
|
680  | 
| Lambda =>  | 
|
681  | 
interpret_formula config language const_map (var_types' @ var_types)  | 
|
682  | 
type_map tptp_formula thy  | 
|
683  | 
|>> fold_bind (fn (n, ty, rest) => lambda (Free (n, ty)) rest)  | 
|
684  | 
| Epsilon =>  | 
|
685  | 
let val (t, thy') =  | 
|
686  | 
interpret_formula config language const_map var_types type_map  | 
|
687  | 
(Quant (Lambda, bindings, tptp_formula)) thy  | 
|
| 49323 | 688  | 
in ((HOLogic.choice_const dummyT) $ t, thy') end  | 
| 46844 | 689  | 
| Iota =>  | 
690  | 
let val (t, thy') =  | 
|
691  | 
interpret_formula config language const_map var_types type_map  | 
|
692  | 
(Quant (Lambda, bindings, tptp_formula)) thy  | 
|
| 49323 | 693  | 
            in (Const (@{const_name The}, dummyT) $ t, thy') end
 | 
| 46844 | 694  | 
| Dep_Prod =>  | 
695  | 
            raise MISINTERPRET_FORMULA ("Unsupported", tptp_fmla)
 | 
|
696  | 
| Dep_Sum =>  | 
|
697  | 
            raise MISINTERPRET_FORMULA ("Unsupported", tptp_fmla)
 | 
|
698  | 
end  | 
|
| 47359 | 699  | 
| Conditional (fmla1, fmla2, fmla3) =>  | 
700  | 
let  | 
|
701  | 
val interp = interpret_formula config language const_map var_types type_map  | 
|
702  | 
val (((fmla1', fmla2'), fmla3'), thy') =  | 
|
703  | 
interp fmla1 thy  | 
|
704  | 
||>> interp fmla2  | 
|
705  | 
||>> interp fmla3  | 
|
706  | 
in (HOLogic.mk_conj (HOLogic.mk_imp (fmla1', fmla2'),  | 
|
707  | 
HOLogic.mk_imp (HOLogic.mk_not fmla1', fmla3')),  | 
|
708  | 
thy')  | 
|
709  | 
end  | 
|
710  | 
| Let (tptp_let_list, tptp_formula) => (*FIXME not yet implemented*)  | 
|
711  | 
        raise MISINTERPRET_FORMULA ("'Let' not yet implemented", tptp_fmla)
 | 
|
| 46844 | 712  | 
| Atom tptp_atom =>  | 
713  | 
(case tptp_atom of  | 
|
714  | 
TFF_Typed_Atom (symbol, tptp_type_opt) =>  | 
|
715  | 
(*FIXME ignoring type info*)  | 
|
| 57808 | 716  | 
(interpret_symbol config const_map symbol [] thy, thy)  | 
| 46844 | 717  | 
| THF_Atom_term tptp_term =>  | 
| 
47570
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
718  | 
interpret_term true config language const_map var_types  | 
| 
 
df3c9aa6c9dc
improved threading of thy-values through interpret functions;
 
sultana 
parents: 
47569 
diff
changeset
 | 
719  | 
type_map tptp_term thy  | 
| 46844 | 720  | 
| THF_Atom_conn_term symbol =>  | 
| 57808 | 721  | 
(interpret_symbol config const_map symbol [] thy, thy))  | 
| 47510 | 722  | 
| Type_fmla _ =>  | 
| 46844 | 723  | 
raise MISINTERPRET_FORMULA  | 
724  | 
          ("Cannot interpret types as formulas", tptp_fmla)
 | 
|
725  | 
| THF_typing (tptp_formula, _) => (*FIXME ignoring type info*)  | 
|
726  | 
interpret_formula config language  | 
|
727  | 
const_map var_types type_map tptp_formula thy  | 
|
728  | 
||
729  | 
(*Extract the type from a typing*)  | 
|
730  | 
local  | 
|
| 57808 | 731  | 
fun type_vars_of_fmlatype (Quant (Dep_Prod, varlist, fmla)) =  | 
732  | 
map fst varlist @ type_vars_of_fmlatype fmla  | 
|
733  | 
| type_vars_of_fmlatype _ = []  | 
|
734  | 
||
| 46844 | 735  | 
fun extract_type tptp_type =  | 
736  | 
case tptp_type of  | 
|
| 57808 | 737  | 
Fmla_type fmla => (type_vars_of_fmlatype fmla, fmlatype_to_type fmla)  | 
738  | 
| _ => ([], tptp_type)  | 
|
| 46844 | 739  | 
in  | 
740  | 
fun typeof_typing (THF_typing (_, tptp_type)) = extract_type tptp_type  | 
|
741  | 
fun typeof_tff_typing (Atom (TFF_Typed_Atom (_, SOME tptp_type))) =  | 
|
742  | 
extract_type tptp_type  | 
|
743  | 
end  | 
|
| 57808 | 744  | 
|
| 46844 | 745  | 
fun nameof_typing  | 
746  | 
(THF_typing  | 
|
747  | 
((Atom (THF_Atom_term (Term_Func (Uninterpreted str, [])))), _)) = str  | 
|
748  | 
fun nameof_tff_typing (Atom (TFF_Typed_Atom (Uninterpreted str, _))) = str  | 
|
749  | 
||
| 57808 | 750  | 
fun strip_prod_type (Prod_type (ty1, ty2)) = ty1 :: strip_prod_type ty2  | 
751  | 
| strip_prod_type ty = [ty]  | 
|
752  | 
||
753  | 
fun dest_fn_type (Fn_type (ty1, ty2)) = (strip_prod_type ty1, ty2)  | 
|
754  | 
| dest_fn_type ty = ([], ty)  | 
|
755  | 
||
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
756  | 
fun resolve_include_path path_prefixes path_suffix =  | 
| 
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
757  | 
case find_first (fn prefix => File.exists (Path.append prefix path_suffix))  | 
| 
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
758  | 
path_prefixes of  | 
| 
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
759  | 
SOME prefix => Path.append prefix path_suffix  | 
| 
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
760  | 
| NONE =>  | 
| 
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
761  | 
    error ("Cannot find include file " ^ quote (Path.implode path_suffix))
 | 
| 
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
762  | 
|
| 57808 | 763  | 
(* Ideally, to be in sync with TFF1, we should perform full type skolemization here.  | 
764  | 
But the problems originating from HOL systems are restricted to top-level  | 
|
765  | 
universal quantification for types. *)  | 
|
766  | 
fun remove_leading_type_quantifiers (Quant (Forall, varlist, tptp_fmla)) =  | 
|
767  | 
(case filter_out (curry (op =) (SOME (Defined_type Type_Type)) o snd) varlist of  | 
|
768  | 
[] => remove_leading_type_quantifiers tptp_fmla  | 
|
769  | 
| varlist' => Quant (Forall, varlist', tptp_fmla))  | 
|
770  | 
| remove_leading_type_quantifiers tptp_fmla = tptp_fmla  | 
|
771  | 
||
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
772  | 
fun interpret_line config inc_list type_map const_map path_prefixes line thy =  | 
| 46844 | 773  | 
case line of  | 
| 
47569
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
774  | 
Include (_, quoted_path, inc_list) =>  | 
| 46844 | 775  | 
interpret_file'  | 
776  | 
config  | 
|
| 47511 | 777  | 
inc_list  | 
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
778  | 
path_prefixes  | 
| 
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
779  | 
(resolve_include_path path_prefixes  | 
| 
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
780  | 
(quoted_path |> unenclose |> Path.explode))  | 
| 46844 | 781  | 
type_map  | 
782  | 
const_map  | 
|
783  | 
thy  | 
|
| 
47569
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
784  | 
| Annotated_Formula (_, lang, label, role, tptp_formula, annotation_opt) =>  | 
| 47511 | 785  | 
(*interpret a line only if "label" is in "inc_list", or if "inc_list" is  | 
786  | 
empty (in which case interpret all lines)*)  | 
|
787  | 
(*FIXME could work better if inc_list were sorted*)  | 
|
788  | 
if null inc_list orelse is_some (List.find (fn x => x = label) inc_list) then  | 
|
789  | 
case role of  | 
|
790  | 
Role_Type =>  | 
|
791  | 
let  | 
|
| 57808 | 792  | 
val ((tptp_type_vars, tptp_ty), name) =  | 
| 47511 | 793  | 
if lang = THF then  | 
794  | 
(typeof_typing tptp_formula (*assuming tptp_formula is a typing*),  | 
|
795  | 
nameof_typing tptp_formula (*and that the LHS is a (atom) name*))  | 
|
796  | 
else  | 
|
797  | 
(typeof_tff_typing tptp_formula,  | 
|
798  | 
nameof_tff_typing tptp_formula)  | 
|
799  | 
in  | 
|
| 57808 | 800  | 
case dest_fn_type tptp_ty of  | 
801  | 
(tptp_binders, Defined_type Type_Type) => (*add new type*)  | 
|
| 47511 | 802  | 
(*generate an Isabelle/HOL type to interpret this TPTP type,  | 
803  | 
and add it to both the Isabelle/HOL theory and to the type_map*)  | 
|
| 46844 | 804  | 
let  | 
| 47511 | 805  | 
val (type_map', thy') =  | 
| 57808 | 806  | 
declare_type config  | 
807  | 
(name, (name, length tptp_binders)) type_map thy  | 
|
808  | 
in ((type_map', const_map, []), thy') end  | 
|
| 
47513
 
50a424b89952
improved naming of 'distinct objects' in tptp import
 
sultana 
parents: 
47511 
diff
changeset
 | 
809  | 
|
| 47511 | 810  | 
| _ => (*declaration of constant*)  | 
811  | 
(*Here we populate the map from constants to the Isabelle/HOL  | 
|
812  | 
terms they map to (which in turn contain their types).*)  | 
|
813  | 
let  | 
|
814  | 
val ty = interpret_type config thy type_map tptp_ty  | 
|
815  | 
(*make sure that the theory contains all the types appearing  | 
|
816  | 
in this constant's signature. Exception is thrown if encounter  | 
|
817  | 
an undeclared types.*)  | 
|
| 57808 | 818  | 
val (t as Const (name', _), thy') =  | 
| 47511 | 819  | 
let  | 
820  | 
fun analyse_type thy thf_ty =  | 
|
821  | 
if #cautious config then  | 
|
822  | 
case thf_ty of  | 
|
823  | 
Fn_type (thf_ty1, thf_ty2) =>  | 
|
824  | 
(analyse_type thy thf_ty1;  | 
|
825  | 
analyse_type thy thf_ty2)  | 
|
| 57808 | 826  | 
| Atom_type (ty_name, _) =>  | 
| 47548 | 827  | 
if type_exists config thy ty_name then ()  | 
| 47511 | 828  | 
else  | 
829  | 
raise MISINTERPRET_TYPE  | 
|
830  | 
                                   ("Type (in signature of " ^
 | 
|
831  | 
name ^ ") has not been declared",  | 
|
| 57808 | 832  | 
Atom_type (ty_name, []))  | 
| 47511 | 833  | 
| _ => ()  | 
834  | 
else () (*skip test if we're not being cautious.*)  | 
|
835  | 
in  | 
|
836  | 
analyse_type thy tptp_ty;  | 
|
837  | 
declare_constant config name ty thy  | 
|
838  | 
end  | 
|
839  | 
(*register a mapping from name to t. Constants' type  | 
|
840  | 
declarations should occur at most once, so it's justified to  | 
|
841  | 
use "::". Note that here we use a constant's name rather  | 
|
842  | 
than the possibly- new one, since all references in the  | 
|
843  | 
theory will be to this name.*)  | 
|
| 57808 | 844  | 
|
845  | 
val tf_tyargs = map tfree_of_var_type tptp_type_vars  | 
|
846  | 
val isa_tyargs = Sign.const_typargs thy' (name', ty)  | 
|
847  | 
val _ =  | 
|
848  | 
if length isa_tyargs < length tf_tyargs then  | 
|
849  | 
raise MISINTERPRET_SYMBOL  | 
|
850  | 
                          ("Cannot handle phantom types for constant",
 | 
|
851  | 
Uninterpreted name)  | 
|
852  | 
else  | 
|
853  | 
()  | 
|
854  | 
val tyarg_perm =  | 
|
855  | 
map (fn T => find_index (curry (op =) T) tf_tyargs) isa_tyargs  | 
|
856  | 
val const_map' = ((name, (t, tyarg_perm)) :: const_map)  | 
|
| 47511 | 857  | 
in ((type_map,(*type_map unchanged, since a constant's  | 
858  | 
declaration shouldn't include any new types.*)  | 
|
859  | 
const_map',(*typing table of constant was extended*)  | 
|
860  | 
[]),(*no formulas were interpreted*)  | 
|
861  | 
thy')(*theory was extended with new constant*)  | 
|
862  | 
end  | 
|
863  | 
end  | 
|
| 
47513
 
50a424b89952
improved naming of 'distinct objects' in tptp import
 
sultana 
parents: 
47511 
diff
changeset
 | 
864  | 
|
| 47511 | 865  | 
| _ => (*i.e. the AF is not a type declaration*)  | 
866  | 
let  | 
|
867  | 
(*gather interpreted term, and the map of types occurring in that term*)  | 
|
868  | 
val (t, thy') =  | 
|
869  | 
interpret_formula config lang  | 
|
| 57808 | 870  | 
const_map [] type_map (remove_leading_type_quantifiers tptp_formula) thy  | 
| 47511 | 871  | 
|>> HOLogic.mk_Trueprop  | 
872  | 
(*type_maps grow monotonically, so return the newest value (type_mapped);  | 
|
873  | 
there's no need to unify it with the type_map parameter.*)  | 
|
874  | 
in  | 
|
875  | 
((type_map, const_map,  | 
|
| 47548 | 876  | 
[(label, role,  | 
| 47511 | 877  | 
Syntax.check_term (Proof_Context.init_global thy') t,  | 
| 
53388
 
c878390475f3
using richer annotation from formula annotations in proof;
 
sultana 
parents: 
49323 
diff
changeset
 | 
878  | 
TPTP_Proof.extract_source_info annotation_opt)]), thy')  | 
| 47511 | 879  | 
end  | 
880  | 
else (*do nothing if we're not to includ this AF*)  | 
|
881  | 
((type_map, const_map, []), thy)  | 
|
| 
47569
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
882  | 
|
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
883  | 
and interpret_problem config inc_list path_prefixes lines type_map const_map thy =  | 
| 
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
884  | 
let val thy_with_symbols = add_interp_symbols_to_thy config type_map thy in  | 
| 47511 | 885  | 
fold (fn line =>  | 
886  | 
fn ((type_map, const_map, lines), thy) =>  | 
|
887  | 
let  | 
|
888  | 
(*process the line, ignoring the type-context for variables*)  | 
|
889  | 
val ((type_map', const_map', l'), thy') =  | 
|
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
890  | 
interpret_line config inc_list type_map const_map path_prefixes  | 
| 
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
891  | 
line thy  | 
| 47688 | 892  | 
(*FIXME  | 
| 
47569
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
893  | 
handle  | 
| 
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
894  | 
(*package all exceptions to include position information,  | 
| 
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
895  | 
even relating to multiple levels of "include"ed files*)  | 
| 
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
896  | 
(*FIXME "exn" contents may appear garbled due to markup  | 
| 
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
897  | 
FIXME this appears as ML source position *)  | 
| 
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
898  | 
MISINTERPRET (pos_list, exn) =>  | 
| 
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
899  | 
raise MISINTERPRET  | 
| 
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
900  | 
(TPTP_Syntax.pos_of_line line :: pos_list, exn)  | 
| 
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
901  | 
| exn => raise MISINTERPRET  | 
| 
 
fce9d97a3258
exceptions related to interpreting tptp problems now mention the relevant position in the tptp file;
 
sultana 
parents: 
47558 
diff
changeset
 | 
902  | 
([TPTP_Syntax.pos_of_line line], exn)  | 
| 47688 | 903  | 
*)  | 
| 47511 | 904  | 
in  | 
905  | 
((type_map',  | 
|
906  | 
const_map',  | 
|
907  | 
l' @ lines),(*append is sufficient, union would be overkill*)  | 
|
908  | 
thy')  | 
|
909  | 
end)  | 
|
910  | 
lines (*lines of the problem file*)  | 
|
| 
47692
 
3d76c81b5ed2
improved non-interpretation of constants and numbers;
 
sultana 
parents: 
47691 
diff
changeset
 | 
911  | 
((type_map, const_map, []), thy_with_symbols) (*initial values*)  | 
| 47511 | 912  | 
end  | 
| 46844 | 913  | 
|
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
914  | 
and interpret_file' config inc_list path_prefixes file_name =  | 
| 46844 | 915  | 
let  | 
916  | 
val file_name' = Path.expand file_name  | 
|
917  | 
in  | 
|
918  | 
if Path.is_absolute file_name' then  | 
|
919  | 
Path.implode file_name  | 
|
920  | 
|> TPTP_Parser.parse_file  | 
|
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
921  | 
|> interpret_problem config inc_list path_prefixes  | 
| 47510 | 922  | 
else error "Could not determine absolute path to file"  | 
| 46844 | 923  | 
end  | 
924  | 
||
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
925  | 
and interpret_file cautious path_prefixes file_name =  | 
| 46844 | 926  | 
let  | 
927  | 
val config =  | 
|
928  | 
      {cautious = cautious,
 | 
|
929  | 
problem_name =  | 
|
930  | 
SOME (TPTP_Problem_Name.parse_problem_name ((Path.base #> Path.implode)  | 
|
| 47520 | 931  | 
file_name))}  | 
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
932  | 
in interpret_file' config [] path_prefixes file_name end  | 
| 46844 | 933  | 
|
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
934  | 
fun import_file cautious path_prefixes file_name type_map const_map thy =  | 
| 46844 | 935  | 
let  | 
936  | 
val prob_name =  | 
|
937  | 
TPTP_Problem_Name.parse_problem_name (Path.implode (Path.base file_name))  | 
|
938  | 
val (result, thy') =  | 
|
| 
48829
 
6ed588c4f963
look in current directory first before looking up includes in the TPTP directory, as required by Geoff
 
blanchet 
parents: 
47729 
diff
changeset
 | 
939  | 
interpret_file cautious path_prefixes file_name type_map const_map thy  | 
| 47520 | 940  | 
(*FIXME doesn't check if problem has already been interpreted*)  | 
| 46844 | 941  | 
in TPTP_Data.map (cons ((prob_name, result))) thy' end  | 
942  | 
||
943  | 
val _ =  | 
|
| 
59936
 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
944  | 
  Outer_Syntax.command @{command_keyword import_tptp} "import TPTP problem"
 | 
| 
48881
 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 
wenzelm 
parents: 
48829 
diff
changeset
 | 
945  | 
(Parse.path >> (fn name =>  | 
| 
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
 | 
946  | 
Toplevel.theory (fn thy =>  | 
| 
48881
 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 
wenzelm 
parents: 
48829 
diff
changeset
 | 
947  | 
let val path = Path.explode name  | 
| 
 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 
wenzelm 
parents: 
48829 
diff
changeset
 | 
948  | 
(*NOTE: assumes include files are located at $TPTP/Axioms  | 
| 
 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 
wenzelm 
parents: 
48829 
diff
changeset
 | 
949  | 
(which is how TPTP is organised)*)  | 
| 
 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 
wenzelm 
parents: 
48829 
diff
changeset
 | 
950  | 
in import_file true [Path.dir path, Path.explode "$TPTP"] path [] [] thy end)))  | 
| 46844 | 951  | 
|
952  | 
end  |