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