| author | paulson <lp15@cam.ac.uk> | 
| Sun, 12 Apr 2015 11:00:56 +0100 | |
| changeset 60026 | 41d81b4a0a21 | 
| parent 48891 | c0eafbd55de3 | 
| child 62390 | 842917225d56 | 
| permissions | -rw-r--r-- | 
| 46844 | 1  | 
(* Title: HOL/TPTP/TPTP_Parser.thy  | 
2  | 
Author: Nik Sultana, Cambridge University Computer Laboratory  | 
|
3  | 
||
| 
47509
 
6f215c2ebd72
split TPTP_Parser thy -- parser can rely on smaller image, whereas TPTP_Interpret requires HOL;
 
sultana 
parents: 
47411 
diff
changeset
 | 
4  | 
Parser for TPTP formulas  | 
| 46844 | 5  | 
*)  | 
6  | 
||
7  | 
theory TPTP_Parser  | 
|
| 
47509
 
6f215c2ebd72
split TPTP_Parser thy -- parser can rely on smaller image, whereas TPTP_Interpret requires HOL;
 
sultana 
parents: 
47411 
diff
changeset
 | 
8  | 
imports Pure  | 
| 48891 | 9  | 
begin  | 
10  | 
||
11  | 
ML_file "TPTP_Parser/ml_yacc_lib.ML" (*generated from ML-Yacc's lib*)  | 
|
| 46844 | 12  | 
|
| 48891 | 13  | 
ML_file "TPTP_Parser/tptp_syntax.ML"  | 
14  | 
ML_file "TPTP_Parser/tptp_lexyacc.ML" (*generated from tptp.lex and tptp.yacc*)  | 
|
15  | 
ML_file "TPTP_Parser/tptp_parser.ML"  | 
|
16  | 
ML_file "TPTP_Parser/tptp_problem_name.ML"  | 
|
17  | 
ML_file "TPTP_Parser/tptp_proof.ML"  | 
|
| 46844 | 18  | 
|
19  | 
text {*The TPTP parser was generated using ML-Yacc, and needs the
 | 
|
20  | 
ML-Yacc library to operate. This library is included with the parser,  | 
|
21  | 
and we include the next section in accordance with ML-Yacc's terms of  | 
|
22  | 
use.*}  | 
|
23  | 
||
24  | 
section "ML-YACC COPYRIGHT NOTICE, LICENSE AND DISCLAIMER."  | 
|
25  | 
text {*
 | 
|
26  | 
Copyright 1989, 1990 by David R. Tarditi Jr. and Andrew W. Appel  | 
|
27  | 
||
28  | 
Permission to use, copy, modify, and distribute this software and its  | 
|
29  | 
documentation for any purpose and without fee is hereby granted,  | 
|
30  | 
provided that the above copyright notice appear in all copies and that  | 
|
31  | 
both the copyright notice and this permission notice and warranty  | 
|
32  | 
disclaimer appear in supporting documentation, and that the names of  | 
|
33  | 
David R. Tarditi Jr. and Andrew W. Appel not be used in advertising or  | 
|
34  | 
publicity pertaining to distribution of the software without specific,  | 
|
35  | 
written prior permission.  | 
|
36  | 
||
37  | 
David R. Tarditi Jr. and Andrew W. Appel disclaim all warranties with  | 
|
38  | 
regard to this software, including all implied warranties of  | 
|
39  | 
merchantability and fitness. In no event shall David R. Tarditi  | 
|
40  | 
Jr. and Andrew W. Appel be liable for any special, indirect or  | 
|
41  | 
consequential damages or any damages whatsoever resulting from loss of  | 
|
42  | 
use, data or profits, whether in an action of contract, negligence or  | 
|
43  | 
other tortious action, arising out of or in connection with the use or  | 
|
44  | 
performance of this software.  | 
|
45  | 
*}  | 
|
46  | 
||
47  | 
end  |