| 
46844
 | 
     1  | 
(*  Title:      HOL/TPTP/TPTP_Parser.thy
  | 
| 
 | 
     2  | 
    Author:     Nik Sultana, Cambridge University Computer Laboratory
  | 
| 
 | 
     3  | 
  | 
| 
 | 
     4  | 
Importing TPTP files into Isabelle/HOL:
  | 
| 
 | 
     5  | 
* parsing TPTP formulas;
  | 
| 
 | 
     6  | 
* interpreting TPTP formulas as HOL terms (i.e. importing types, and
  | 
| 
 | 
     7  | 
    type-checking the terms);
  | 
| 
 | 
     8  | 
*)
  | 
| 
 | 
     9  | 
  | 
| 
 | 
    10  | 
theory TPTP_Parser
  | 
| 
 | 
    11  | 
imports Main
  | 
| 
 | 
    12  | 
uses
  | 
| 
 | 
    13  | 
  "TPTP_Parser/ml_yacc_lib.ML" (*generated from ML-Yacc's lib*)
  | 
| 
 | 
    14  | 
  | 
| 
 | 
    15  | 
  "TPTP_Parser/tptp_syntax.ML"
  | 
| 
 | 
    16  | 
  "TPTP_Parser/tptp_lexyacc.ML" (*generated from tptp.lex and tptp.yacc*)
  | 
| 
 | 
    17  | 
  "TPTP_Parser/tptp_parser.ML"
  | 
| 
 | 
    18  | 
  "TPTP_Parser/tptp_problem_name.ML"
  | 
| 
 | 
    19  | 
  "TPTP_Parser/tptp_interpret.ML"
  | 
| 
 | 
    20  | 
  | 
| 
 | 
    21  | 
begin
  | 
| 
 | 
    22  | 
  | 
| 
 | 
    23  | 
text {*The TPTP parser was generated using ML-Yacc, and needs the
 | 
| 
 | 
    24  | 
ML-Yacc library to operate.  This library is included with the parser,
  | 
| 
 | 
    25  | 
and we include the next section in accordance with ML-Yacc's terms of
  | 
| 
 | 
    26  | 
use.*}
  | 
| 
 | 
    27  | 
  | 
| 
 | 
    28  | 
section "ML-YACC COPYRIGHT NOTICE, LICENSE AND DISCLAIMER."
  | 
| 
 | 
    29  | 
text {*
 | 
| 
 | 
    30  | 
Copyright 1989, 1990 by David R. Tarditi Jr. and Andrew W. Appel
  | 
| 
 | 
    31  | 
  | 
| 
 | 
    32  | 
Permission to use, copy, modify, and distribute this software and its
  | 
| 
 | 
    33  | 
documentation for any purpose and without fee is hereby granted,
  | 
| 
 | 
    34  | 
provided that the above copyright notice appear in all copies and that
  | 
| 
 | 
    35  | 
both the copyright notice and this permission notice and warranty
  | 
| 
 | 
    36  | 
disclaimer appear in supporting documentation, and that the names of
  | 
| 
 | 
    37  | 
David R. Tarditi Jr. and Andrew W. Appel not be used in advertising or
  | 
| 
 | 
    38  | 
publicity pertaining to distribution of the software without specific,
  | 
| 
 | 
    39  | 
written prior permission.
  | 
| 
 | 
    40  | 
  | 
| 
 | 
    41  | 
David R. Tarditi Jr. and Andrew W. Appel disclaim all warranties with
  | 
| 
 | 
    42  | 
regard to this software, including all implied warranties of
  | 
| 
 | 
    43  | 
merchantability and fitness.  In no event shall David R. Tarditi
  | 
| 
 | 
    44  | 
Jr. and Andrew W. Appel be liable for any special, indirect or
  | 
| 
 | 
    45  | 
consequential damages or any damages whatsoever resulting from loss of
  | 
| 
 | 
    46  | 
use, data or profits, whether in an action of contract, negligence or
  | 
| 
 | 
    47  | 
other tortious action, arising out of or in connection with the use or
  | 
| 
 | 
    48  | 
performance of this software.
  | 
| 
 | 
    49  | 
*}
  | 
| 
 | 
    50  | 
  | 
| 
 | 
    51  | 
end  |