| author | sultana | 
| Tue, 17 Apr 2012 16:14:07 +0100 | |
| changeset 47512 | b381d428a725 | 
| parent 47509 | 6f215c2ebd72 | 
| child 47519 | 9c3acd90063a | 
| permissions | -rw-r--r-- | 
(* Title: HOL/TPTP/TPTP_Interpret.thy Author: Nik Sultana, Cambridge University Computer Laboratory Importing TPTP files into Isabelle/HOL: parsing TPTP formulas and interpreting them as HOL terms (i.e. importing types and type-checking the terms) *) theory TPTP_Interpret imports Main TPTP_Parser keywords "import_tptp" :: thy_decl uses "TPTP_Parser/tptp_interpret.ML" begin end