merges
authorpaulson
Fri, 09 Mar 2012 17:24:42 +0000
changeset 46848 13eeb06847cb
parent 46847 8740cea39a4a (current diff)
parent 46846 9e99afaade17 (diff)
child 46849 36f392239b6a
child 46850 7b04cfc24eb6
merges
--- a/src/HOL/TPTP/ROOT.ML	Fri Mar 09 17:24:00 2012 +0000
+++ b/src/HOL/TPTP/ROOT.ML	Fri Mar 09 17:24:42 2012 +0000
@@ -7,7 +7,8 @@
 *)
 
 use_thys [
-  "ATP_Theory_Export"
+  "ATP_Theory_Export",
+  "TPTP_Parser"
 ];
 
 Unsynchronized.setmp Proofterm.proofs (!Proofterm.proofs)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser.thy	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,51 @@
+(*  Title:      HOL/TPTP/TPTP_Parser.thy
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+Importing TPTP files into Isabelle/HOL:
+* parsing TPTP formulas;
+* interpreting TPTP formulas as HOL terms (i.e. importing types, and
+    type-checking the terms);
+*)
+
+theory TPTP_Parser
+imports Main
+uses
+  "TPTP_Parser/ml_yacc_lib.ML" (*generated from ML-Yacc's lib*)
+
+  "TPTP_Parser/tptp_syntax.ML"
+  "TPTP_Parser/tptp_lexyacc.ML" (*generated from tptp.lex and tptp.yacc*)
+  "TPTP_Parser/tptp_parser.ML"
+  "TPTP_Parser/tptp_problem_name.ML"
+  "TPTP_Parser/tptp_interpret.ML"
+
+begin
+
+text {*The TPTP parser was generated using ML-Yacc, and needs the
+ML-Yacc library to operate.  This library is included with the parser,
+and we include the next section in accordance with ML-Yacc's terms of
+use.*}
+
+section "ML-YACC COPYRIGHT NOTICE, LICENSE AND DISCLAIMER."
+text {*
+Copyright 1989, 1990 by David R. Tarditi Jr. and Andrew W. Appel
+
+Permission to use, copy, modify, and distribute this software and its
+documentation for any purpose and without fee is hereby granted,
+provided that the above copyright notice appear in all copies and that
+both the copyright notice and this permission notice and warranty
+disclaimer appear in supporting documentation, and that the names of
+David R. Tarditi Jr. and Andrew W. Appel not be used in advertising or
+publicity pertaining to distribution of the software without specific,
+written prior permission.
+
+David R. Tarditi Jr. and Andrew W. Appel disclaim all warranties with
+regard to this software, including all implied warranties of
+merchantability and fitness.  In no event shall David R. Tarditi
+Jr. and Andrew W. Appel be liable for any special, indirect or
+consequential damages or any damages whatsoever resulting from loss of
+use, data or profits, whether in an action of contract, negligence or
+other tortious action, arising out of or in connection with the use or
+performance of this software.
+*}
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/README	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,32 @@
+The TPTP parser is generated using ML-Yacc and relies on the ML-Yacc
+library to function. The ML-Yacc library is an external piece of
+software that needs a small modification for use in Isabelle. The
+relationship between Isabelle and ML-Yacc is similar to that with
+Metis (see src/Tools/Metis).
+
+The file "make_tptp_parser" generates the TPTP parser and patches it
+to conform to Isabelle's naming conventions.
+
+In order to generate the parser from its lex/yacc definition you need
+to have the ML-Yacc binaries. The sources can be downloaded via SVN as
+follows:
+
+ svn co --username anonsvn \
+ https://smlnj-gforge.cs.uchicago.edu/svn/smlnj/ml-yacc/trunk ml-yacc
+
+ML-Yacc is usually distributed with Standard ML of New Jersey, and its
+binaries can also be obtained as packages for some distributions of
+Linux. The script "make_tptp_parser" will produce a file called
+tptp_lexyacc.ML -- this is a compilation of the SML files (generated
+by ML-Yacc) making up the TPTP parser.
+
+The generated parser needs ML-Yacc's library. This is distributed with
+ML-Yacc's source code, in the lib/ directory. The ML-Yacc library
+cannot be used directly and must be patched. The script
+"make_mlyacclib" takes the ML-Yacc library (for instance, as
+downloaded from the ML-Yacc repo) and produces the file ml_yacc_lib.ML
+-- this is a compilation of slightly modified files making up
+ML-Yacc's library.
+
+Nik Sultana
+9th March 2012
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/make_mlyacclib	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,54 @@
+#!/bin/bash
+#
+# make_mlyacclib - Generates Isabelle-friendly version of ML-Yacc's library.
+#
+# This code is based on that used in src/Tools/Metis to adapt Metis for
+# use in Isabelle.
+
+MLYACCDIR=./ml-yacc
+MLYACCLIB_FILES="base.sig join.sml lrtable.sml stream.sml parser2.sml"
+
+echo "Cleaning"
+rm -f ml_yacc_lib.ML
+echo "Generating ml_yacc_lib.ML"
+(
+  cat <<EOF
+
+(******************************************************************)
+(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
+(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
+(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
+(******************************************************************)
+
+print_depth 0;
+
+(*
+  This file is generated from the contents of ML-Yacc's lib directory.
+  ML-Yacc's COPYRIGHT-file contents follow:
+
+EOF
+  perl -pe 'print "  ";' ml-yacc/COPYRIGHT
+  echo "*)"
+
+for FILE in $MLYACCLIB_FILES
+do
+  echo
+  echo "(**** Original file: $FILE ****)"
+  echo
+  echo -e "  $FILE" >&2
+  perl -p -e 's/\bref\b/Unsynchronized.ref/g;' \
+          -e 's/Unsafe\.(.*)/\1/g;' \
+          -e 's/\bconcat\b/String.concat/g;' \
+          -e 's/(?<!List\.)foldr\b/List.foldr/g;' \
+          -e 's/\bfoldl\b/List.foldl/g;' \
+          -e 's/val print = fn s => TextIO\.output\(TextIO\.stdOut,s\)$//g;' \
+          -e 's/\bprint\b/TextIO.print/g;' \
+          $MLYACCDIR/lib/$FILE
+  done
+
+  cat <<EOF
+;
+print_depth 10;
+EOF
+
+) > ml_yacc_lib.ML
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/make_tptp_parser	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,37 @@
+#!/bin/bash
+#
+# make_tptp_parser - Runs ML-Yacc to generate TPTP parser and makes it
+#                    Isabelle-friendly.
+#
+# This code is based on that used in src/Tools/Metis to adapt Metis for
+# use in Isabelle.
+
+INTERMEDIATE_FILES="tptp.yacc.sig tptp.lex.sml tptp.yacc.sml"
+
+echo "Cleaning"
+rm -f tptp_lexyacc.ML
+echo "Generating lexer and parser"
+ml-lex tptp.lex && ml-yacc tptp.yacc
+echo "Generating tptp_lexyacc.ML"
+(
+  cat <<EOF
+
+(******************************************************************)
+(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
+(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
+(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
+(******************************************************************)
+
+(*
+  This file is produced from the parser generated by ML-Yacc from the
+  source files tptp.lex and tptp.yacc.
+*)
+EOF
+
+for FILE in $INTERMEDIATE_FILES
+do
+  perl -p -e 's/\bref\b/Unsynchronized.ref/g;' -e 's/Unsafe\.(.*)/\1/g;' $FILE
+done
+) > tptp_lexyacc.ML
+
+rm -f $INTERMEDIATE_FILES tptp.yacc.desc
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/ml-yacc/COPYRIGHT	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,20 @@
+ML-YACC COPYRIGHT NOTICE, LICENSE AND DISCLAIMER.
+
+Copyright 1989, 1990 by David R. Tarditi Jr. and Andrew W. Appel
+
+Permission to use, copy, modify, and distribute this software and its
+documentation for any purpose and without fee is hereby granted,
+provided that the above copyright notice appear in all copies and that
+both the copyright notice and this permission notice and warranty
+disclaimer appear in supporting documentation, and that the names of
+David R. Tarditi Jr. and Andrew W. Appel not be used in advertising
+or publicity pertaining to distribution of the software without
+specific, written prior permission.
+
+David R. Tarditi Jr. and Andrew W. Appel disclaim all warranties with regard to
+this software, including all implied warranties of merchantability and fitness.
+In no event shall David R. Tarditi Jr. and Andrew W. Appel be liable for any
+special, indirect or consequential damages or any damages whatsoever resulting
+from loss of use, data or profits, whether in an action of contract, negligence
+or other tortious action, arising out of or in connection with the use or
+performance of this software.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/ml-yacc/lib/base.sig	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,299 @@
+(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
+
+(* base.sig: Base signature file for SML-Yacc.  This file contains signatures
+   that must be loaded before any of the files produced by ML-Yacc are loaded
+*)
+
+(* STREAM: signature for a lazy stream.*)
+
+signature STREAM =
+ sig type 'xa stream
+     val streamify : (unit -> '_a) -> '_a stream
+     val cons : '_a * '_a stream -> '_a stream
+     val get : '_a stream -> '_a * '_a stream
+ end
+
+(* LR_TABLE: signature for an LR Table.
+
+   The list of actions and gotos passed to mkLrTable must be ordered by state
+   number. The values for state 0 are the first in the list, the values for
+    state 1 are next, etc.
+*)
+
+signature LR_TABLE =
+    sig
+        datatype ('a,'b) pairlist = EMPTY | PAIR of 'a * 'b * ('a,'b) pairlist
+	datatype state = STATE of int
+	datatype term = T of int
+	datatype nonterm = NT of int
+	datatype action = SHIFT of state
+			| REDUCE of int
+			| ACCEPT
+			| ERROR
+	type table
+	
+	val numStates : table -> int
+	val numRules : table -> int
+	val describeActions : table -> state ->
+				(term,action) pairlist * action
+	val describeGoto : table -> state -> (nonterm,state) pairlist
+	val action : table -> state * term -> action
+	val goto : table -> state * nonterm -> state
+	val initialState : table -> state
+	exception Goto of state * nonterm
+
+	val mkLrTable : {actions : ((term,action) pairlist * action) array,
+			 gotos : (nonterm,state) pairlist array,
+			 numStates : int, numRules : int,
+			 initialState : state} -> table
+    end
+
+(* TOKEN: signature revealing the internal structure of a token. This signature
+   TOKEN distinct from the signature {parser name}_TOKENS produced by ML-Yacc.
+   The {parser name}_TOKENS structures contain some types and functions to
+    construct tokens from values and positions.
+
+   The representation of token was very carefully chosen here to allow the
+   polymorphic parser to work without knowing the types of semantic values
+   or line numbers.
+
+   This has had an impact on the TOKENS structure produced by SML-Yacc, which
+   is a structure parameter to lexer functors.  We would like to have some
+   type 'a token which functions to construct tokens would create.  A
+   constructor function for a integer token might be
+
+	  INT: int * 'a * 'a -> 'a token.
+ 
+   This is not possible because we need to have tokens with the representation
+   given below for the polymorphic parser.
+
+   Thus our constructur functions for tokens have the form:
+
+	  INT: int * 'a * 'a -> (svalue,'a) token
+
+   This in turn has had an impact on the signature that lexers for SML-Yacc
+   must match and the types that a user must declare in the user declarations
+   section of lexers.
+*)
+
+signature TOKEN =
+    sig
+	structure LrTable : LR_TABLE
+        datatype ('a,'b) token = TOKEN of LrTable.term * ('a * 'b * 'b)
+	val sameToken : ('a,'b) token * ('a,'b) token -> bool
+    end
+
+(* LR_PARSER: signature for a polymorphic LR parser *)
+
+signature LR_PARSER =
+    sig
+	structure Stream: STREAM
+	structure LrTable : LR_TABLE
+	structure Token : TOKEN
+
+	sharing LrTable = Token.LrTable
+
+	exception ParseError
+
+	val parse : {table : LrTable.table,
+		     lexer : ('_b,'_c) Token.token Stream.stream,
+		     arg: 'arg,
+		     saction : int *
+			       '_c *
+				(LrTable.state * ('_b * '_c * '_c)) list * 
+				'arg ->
+				     LrTable.nonterm *
+				     ('_b * '_c * '_c) *
+				     ((LrTable.state *('_b * '_c * '_c)) list),
+		     void : '_b,
+		     ec : { is_keyword : LrTable.term -> bool,
+			    noShift : LrTable.term -> bool,
+			    preferred_change : (LrTable.term list * LrTable.term list) list,
+			    errtermvalue : LrTable.term -> '_b,
+			    showTerminal : LrTable.term -> string,
+			    terms: LrTable.term list,
+			    error : string * '_c * '_c -> unit
+			   },
+		     lookahead : int  (* max amount of lookahead used in *)
+				      (* error correction *)
+			} -> '_b *
+			     (('_b,'_c) Token.token Stream.stream)
+    end
+
+(* LEXER: a signature that most lexers produced for use with SML-Yacc's
+   output will match.  The user is responsible for declaring type token,
+   type pos, and type svalue in the UserDeclarations section of a lexer.
+
+   Note that type token is abstract in the lexer.  This allows SML-Yacc to
+   create a TOKENS signature for use with lexers produced by ML-Lex that
+   treats the type token abstractly.  Lexers that are functors parametrized by
+   a Tokens structure matching a TOKENS signature cannot examine the structure
+   of tokens.
+*)
+
+signature LEXER =
+   sig
+       structure UserDeclarations :
+	   sig
+	        type ('a,'b) token
+		type pos
+		type svalue
+	   end
+	val makeLexer : (int -> string) -> unit -> 
+         (UserDeclarations.svalue,UserDeclarations.pos) UserDeclarations.token
+   end
+
+(* ARG_LEXER: the %arg option of ML-Lex allows users to produce lexers which
+   also take an argument before yielding a function from unit to a token
+*)
+
+signature ARG_LEXER =
+   sig
+       structure UserDeclarations :
+	   sig
+	        type ('a,'b) token
+		type pos
+		type svalue
+		type arg
+	   end
+	val makeLexer : (int -> string) -> UserDeclarations.arg -> unit -> 
+         (UserDeclarations.svalue,UserDeclarations.pos) UserDeclarations.token
+   end
+
+(* PARSER_DATA: the signature of ParserData structures in {parser name}LrValsFun
+   produced by  SML-Yacc.  All such structures match this signature.  
+
+   The {parser name}LrValsFun produces a structure which contains all the values
+   except for the lexer needed to call the polymorphic parser mentioned
+   before.
+
+*)
+
+signature PARSER_DATA =
+   sig
+        (* the type of line numbers *)
+
+	type pos
+
+	(* the type of semantic values *)
+
+	type svalue
+
+         (* the type of the user-supplied argument to the parser *)
+ 	type arg
+ 
+	(* the intended type of the result of the parser.  This value is
+	   produced by applying extract from the structure Actions to the
+	   final semantic value resultiing from a parse.
+	 *)
+
+	type result
+
+	structure LrTable : LR_TABLE
+	structure Token : TOKEN
+	sharing Token.LrTable = LrTable
+
+	(* structure Actions contains the functions which mantain the
+	   semantic values stack in the parser.  Void is used to provide
+	   a default value for the semantic stack.
+	 *)
+
+	structure Actions : 
+	  sig
+	      val actions : int * pos *
+		   (LrTable.state * (svalue * pos * pos)) list * arg->
+		         LrTable.nonterm * (svalue * pos * pos) *
+			 ((LrTable.state *(svalue * pos * pos)) list)
+	      val void : svalue
+	      val extract : svalue -> result
+	  end
+
+	(* structure EC contains information used to improve error
+	   recovery in an error-correcting parser *)
+
+	structure EC :
+	   sig
+	     val is_keyword : LrTable.term -> bool
+	     val noShift : LrTable.term -> bool
+ 	     val preferred_change : (LrTable.term list * LrTable.term list) list
+	     val errtermvalue : LrTable.term -> svalue
+	     val showTerminal : LrTable.term -> string
+	     val terms: LrTable.term list
+	   end
+
+	(* table is the LR table for the parser *)
+
+	val table : LrTable.table
+    end
+
+(* signature PARSER is the signature that most user parsers created by 
+   SML-Yacc will match.
+*)
+
+signature PARSER =
+    sig
+        structure Token : TOKEN
+	structure Stream : STREAM
+	exception ParseError
+
+	(* type pos is the type of line numbers *)
+
+	type pos
+
+	(* type result is the type of the result from the parser *)
+
+	type result
+
+         (* the type of the user-supplied argument to the parser *)
+ 	type arg
+	
+	(* type svalue is the type of semantic values for the semantic value
+	   stack
+	 *)
+
+	type svalue
+
+	(* val makeLexer is used to create a stream of tokens for the parser *)
+
+	val makeLexer : (int -> string) ->
+			 (svalue,pos) Token.token Stream.stream
+
+	(* val parse takes a stream of tokens and a function to print
+	   errors and returns a value of type result and a stream containing
+	   the unused tokens
+	 *)
+
+	val parse : int * ((svalue,pos) Token.token Stream.stream) *
+		    (string * pos * pos -> unit) * arg ->
+				result * (svalue,pos) Token.token Stream.stream
+
+	val sameToken : (svalue,pos) Token.token * (svalue,pos) Token.token ->
+				bool
+     end
+
+(* signature ARG_PARSER is the signature that will be matched by parsers whose
+    lexer takes an additional argument.
+*)
+
+signature ARG_PARSER = 
+    sig
+        structure Token : TOKEN
+	structure Stream : STREAM
+	exception ParseError
+
+	type arg
+	type lexarg
+	type pos
+	type result
+	type svalue
+
+	val makeLexer : (int -> string) -> lexarg ->
+			 (svalue,pos) Token.token Stream.stream
+	val parse : int * ((svalue,pos) Token.token Stream.stream) *
+		    (string * pos * pos -> unit) * arg ->
+				result * (svalue,pos) Token.token Stream.stream
+
+	val sameToken : (svalue,pos) Token.token * (svalue,pos) Token.token ->
+				bool
+     end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/ml-yacc/lib/join.sml	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,94 @@
+(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
+
+(* functor Join creates a user parser by putting together a Lexer structure,
+   an LrValues structure, and a polymorphic parser structure.  Note that
+   the Lexer and LrValues structure must share the type pos (i.e. the type
+   of line numbers), the type svalues for semantic values, and the type
+   of tokens.
+*)
+
+functor Join(structure Lex : LEXER
+	     structure ParserData: PARSER_DATA
+	     structure LrParser : LR_PARSER
+	     sharing ParserData.LrTable = LrParser.LrTable
+	     sharing ParserData.Token = LrParser.Token
+	     sharing type Lex.UserDeclarations.svalue = ParserData.svalue
+	     sharing type Lex.UserDeclarations.pos = ParserData.pos
+	     sharing type Lex.UserDeclarations.token = ParserData.Token.token)
+		 : PARSER =
+struct
+    structure Token = ParserData.Token
+    structure Stream = LrParser.Stream
+ 
+    exception ParseError = LrParser.ParseError
+
+    type arg = ParserData.arg
+    type pos = ParserData.pos
+    type result = ParserData.result
+    type svalue = ParserData.svalue
+    val makeLexer = LrParser.Stream.streamify o Lex.makeLexer
+    val parse = fn (lookahead,lexer,error,arg) =>
+	(fn (a,b) => (ParserData.Actions.extract a,b))
+     (LrParser.parse {table = ParserData.table,
+	        lexer=lexer,
+		lookahead=lookahead,
+		saction = ParserData.Actions.actions,
+		arg=arg,
+		void= ParserData.Actions.void,
+	        ec = {is_keyword = ParserData.EC.is_keyword,
+		      noShift = ParserData.EC.noShift,
+		      preferred_change = ParserData.EC.preferred_change,
+		      errtermvalue = ParserData.EC.errtermvalue,
+		      error=error,
+		      showTerminal = ParserData.EC.showTerminal,
+		      terms = ParserData.EC.terms}}
+      )
+     val sameToken = Token.sameToken
+end
+
+(* functor JoinWithArg creates a variant of the parser structure produced 
+   above.  In this case, the makeLexer take an additional argument before
+   yielding a value of type unit -> (svalue,pos) token
+ *)
+
+functor JoinWithArg(structure Lex : ARG_LEXER
+	     structure ParserData: PARSER_DATA
+	     structure LrParser : LR_PARSER
+	     sharing ParserData.LrTable = LrParser.LrTable
+	     sharing ParserData.Token = LrParser.Token
+	     sharing type Lex.UserDeclarations.svalue = ParserData.svalue
+	     sharing type Lex.UserDeclarations.pos = ParserData.pos
+	     sharing type Lex.UserDeclarations.token = ParserData.Token.token)
+		 : ARG_PARSER  =
+struct
+    structure Token = ParserData.Token
+    structure Stream = LrParser.Stream
+
+    exception ParseError = LrParser.ParseError
+
+    type arg = ParserData.arg
+    type lexarg = Lex.UserDeclarations.arg
+    type pos = ParserData.pos
+    type result = ParserData.result
+    type svalue = ParserData.svalue
+
+    val makeLexer = fn s => fn arg =>
+		 LrParser.Stream.streamify (Lex.makeLexer s arg)
+    val parse = fn (lookahead,lexer,error,arg) =>
+	(fn (a,b) => (ParserData.Actions.extract a,b))
+     (LrParser.parse {table = ParserData.table,
+	        lexer=lexer,
+		lookahead=lookahead,
+		saction = ParserData.Actions.actions,
+		arg=arg,
+		void= ParserData.Actions.void,
+	        ec = {is_keyword = ParserData.EC.is_keyword,
+		      noShift = ParserData.EC.noShift,
+		      preferred_change = ParserData.EC.preferred_change,
+		      errtermvalue = ParserData.EC.errtermvalue,
+		      error=error,
+		      showTerminal = ParserData.EC.showTerminal,
+		      terms = ParserData.EC.terms}}
+      )
+    val sameToken = Token.sameToken
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/ml-yacc/lib/lrtable.sml	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,59 @@
+(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
+structure LrTable : LR_TABLE = 
+    struct
+	open Array List
+	infix 9 sub
+	datatype ('a,'b) pairlist = EMPTY
+				  | PAIR of 'a * 'b * ('a,'b) pairlist
+	datatype term = T of int
+	datatype nonterm = NT of int
+	datatype state = STATE of int
+	datatype action = SHIFT of state
+			| REDUCE of int (* rulenum from grammar *)
+			| ACCEPT
+			| ERROR
+	exception Goto of state * nonterm
+	type table = {states: int, rules : int,initialState: state,
+		      action: ((term,action) pairlist * action) array,
+		      goto :  (nonterm,state) pairlist array}
+	val numStates = fn ({states,...} : table) => states
+	val numRules = fn ({rules,...} : table) => rules
+	val describeActions =
+	   fn ({action,...} : table) => 
+	           fn (STATE s) => action sub s
+	val describeGoto =
+	   fn ({goto,...} : table) =>
+	           fn (STATE s) => goto sub s
+	fun findTerm (T term,row,default) =
+	    let fun find (PAIR (T key,data,r)) =
+		       if key < term then find r
+		       else if key=term then data
+		       else default
+		   | find EMPTY = default
+	    in find row
+	    end
+	fun findNonterm (NT nt,row) =
+	    let fun find (PAIR (NT key,data,r)) =
+		       if key < nt then find r
+		       else if key=nt then SOME data
+		       else NONE
+		   | find EMPTY = NONE
+	    in find row
+	    end
+	val action = fn ({action,...} : table) =>
+		fn (STATE state,term) =>
+		  let val (row,default) = action sub state
+		  in findTerm(term,row,default)
+		  end
+	val goto = fn ({goto,...} : table) =>
+			fn (a as (STATE state,nonterm)) =>
+			  case findNonterm(nonterm,goto sub state)
+			  of SOME state => state
+			   | NONE => raise (Goto a)
+	val initialState = fn ({initialState,...} : table) => initialState
+	val mkLrTable = fn {actions,gotos,initialState,numStates,numRules} =>
+	     ({action=actions,goto=gotos,
+	       states=numStates,
+	       rules=numRules,
+               initialState=initialState} : table)
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/ml-yacc/lib/ml-yacc-lib.cm	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,30 @@
+(* sources file for ML-Yacc library *)
+
+Library
+
+signature STREAM
+signature LR_TABLE
+signature TOKEN
+signature LR_PARSER
+signature LEXER
+signature ARG_LEXER
+signature PARSER_DATA
+signature PARSER
+signature ARG_PARSER
+functor Join
+functor JoinWithArg
+structure LrTable
+structure Stream
+structure LrParser
+
+is
+
+#if defined(NEW_CM)
+  $/basis.cm
+#endif
+
+base.sig
+join.sml
+lrtable.sml
+stream.sml
+parser2.sml	(* error correcting version *)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/ml-yacc/lib/parser1.sml	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,98 @@
+(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
+
+(* drt (12/15/89) -- the functor should be used during development work,
+   but it is wastes space in the release version.
+   
+functor ParserGen(structure LrTable : LR_TABLE
+		  structure Stream : STREAM) : LR_PARSER =
+*)
+
+structure LrParser :> LR_PARSER =
+ struct
+     val print = fn s => output(std_out,s)
+     val println = fn s => (print s; print "\n")
+     structure LrTable = LrTable
+     structure Stream = Stream
+     structure Token : TOKEN =
+	struct
+	    structure LrTable = LrTable
+	    datatype ('a,'b) token = TOKEN of LrTable.term * ('a * 'b * 'b)
+	    val sameToken = fn (TOKEN (t,_),TOKEN(t',_)) => t=t'
+	end
+     
+
+     open LrTable 
+     open Token
+
+     val DEBUG = false
+     exception ParseError
+
+      type ('a,'b) elem = (state * ('a * 'b * 'b))
+      type ('a,'b) stack = ('a,'b) elem list
+
+      val showState = fn (STATE s) => ("STATE " ^ (makestring s))
+
+      fun printStack(stack: ('a,'b) elem list, n: int) =
+         case stack
+           of (state, _) :: rest =>
+                 (print("          " ^ makestring n ^ ": ");
+                  println(showState state);
+                  printStack(rest, n+1)
+                 )
+            | nil => ()
+
+      val parse = fn {arg : 'a,
+		      table : LrTable.table,
+		      lexer : ('_b,'_c) token Stream.stream,
+		      saction : int * '_c * ('_b,'_c) stack * 'a ->
+				nonterm * ('_b * '_c * '_c) * ('_b,'_c) stack,
+		      void : '_b,
+		      ec = {is_keyword,preferred_change,
+			    errtermvalue,showTerminal,
+			    error,terms,noShift},
+		      lookahead} =>
+ let fun prAction(stack as (state, _) :: _, 
+		  next as (TOKEN (term,_),_), action) =
+             (println "Parse: state stack:";
+              printStack(stack, 0);
+              print("       state="
+                         ^ showState state	
+                         ^ " next="
+                         ^ showTerminal term
+                         ^ " action="
+                        );
+              case action
+                of SHIFT s => println ("SHIFT " ^ showState s)
+                 | REDUCE i => println ("REDUCE " ^ (makestring i))
+                 | ERROR => println "ERROR"
+		 | ACCEPT => println "ACCEPT";
+              action)
+        | prAction (_,_,action) = action
+
+      val action = LrTable.action table
+      val goto = LrTable.goto table
+
+      fun parseStep(next as (TOKEN (terminal, value as (_,leftPos,_)),lexer) :
+			('_b,'_c) token * ('_b,'_c) token Stream.stream,
+		    stack as (state,_) :: _ : ('_b ,'_c) stack) =
+         case (if DEBUG then prAction(stack, next,action(state, terminal))
+               else action(state, terminal))
+              of SHIFT s => parseStep(Stream.get lexer, (s,value) :: stack)
+               | REDUCE i =>
+		    let val (nonterm,value,stack as (state,_) :: _ ) =
+					 saction(i,leftPos,stack,arg)
+		    in parseStep(next,(goto(state,nonterm),value)::stack)
+		    end
+               | ERROR => let val (_,leftPos,rightPos) = value
+		          in error("syntax error\n",leftPos,rightPos);
+			     raise ParseError
+			  end
+  	       | ACCEPT => let val (_,(topvalue,_,_)) :: _ = stack
+			       val (token,restLexer) = next
+			   in (topvalue,Stream.cons(token,lexer))
+			   end
+      val next as (TOKEN (terminal,(_,leftPos,_)),_) = Stream.get lexer
+   in parseStep(next,[(initialState table,(void,leftPos,leftPos))])
+   end
+end;
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/ml-yacc/lib/parser2.sml	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,542 @@
+(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
+
+(* parser.sml:  This is a parser driver for LR tables with an error-recovery
+   routine added to it.  The routine used is described in detail in this
+   article:
+
+	'A Practical Method for LR and LL Syntactic Error Diagnosis and
+	 Recovery', by M. Burke and G. Fisher, ACM Transactions on
+	 Programming Langauges and Systems, Vol. 9, No. 2, April 1987,
+	 pp. 164-197.
+
+    This program is an implementation is the partial, deferred method discussed
+    in the article.  The algorithm and data structures used in the program
+    are described below.  
+
+    This program assumes that all semantic actions are delayed.  A semantic
+    action should produce a function from unit -> value instead of producing the
+    normal value.  The parser returns the semantic value on the top of the
+    stack when accept is encountered.  The user can deconstruct this value
+    and apply the unit -> value function in it to get the answer.
+
+    It also assumes that the lexer is a lazy stream.
+
+    Data Structures:
+    ----------------
+	
+	* The parser:
+
+	   The state stack has the type
+
+		 (state * (semantic value * line # * line #)) list
+
+	   The parser keeps a queue of (state stack * lexer pair).  A lexer pair
+	 consists of a terminal * value pair and a lexer.  This allows the 
+	 parser to reconstruct the states for terminals to the left of a
+	 syntax error, and attempt to make error corrections there.
+
+	   The queue consists of a pair of lists (x,y).  New additions to
+	 the queue are cons'ed onto y.  The first element of x is the top
+	 of the queue.  If x is nil, then y is reversed and used
+	 in place of x.
+
+    Algorithm:
+    ----------
+
+	* The steady-state parser:  
+
+	    This parser keeps the length of the queue of state stacks at
+	a steady state by always removing an element from the front when
+	another element is placed on the end.
+
+	    It has these arguments:
+
+	   stack: current stack
+	   queue: value of the queue
+	   lexPair ((terminal,value),lex stream)
+
+	When SHIFT is encountered, the state to shift to and the value are
+	are pushed onto the state stack.  The state stack and lexPair are
+	placed on the queue.  The front element of the queue is removed.
+
+	When REDUCTION is encountered, the rule is applied to the current
+	stack to yield a triple (nonterm,value,new stack).  A new
+	stack is formed by adding (goto(top state of stack,nonterm),value)
+	to the stack.
+
+	When ACCEPT is encountered, the top value from the stack and the
+	lexer are returned.
+
+	When an ERROR is encountered, fixError is called.  FixError
+	takes the arguments to the parser, fixes the error if possible and
+        returns a new set of arguments.
+
+	* The distance-parser:
+
+	This parser includes an additional argument distance.  It pushes
+	elements on the queue until it has parsed distance tokens, or an
+	ACCEPT or ERROR occurs.  It returns a stack, lexer, the number of
+	tokens left unparsed, a queue, and an action option.
+*)
+
+signature FIFO = 
+  sig type 'a queue
+      val empty : 'a queue
+      exception Empty
+      val get : 'a queue -> 'a * 'a queue
+      val put : 'a * 'a queue -> 'a queue
+  end
+
+(* drt (12/15/89) -- the functor should be used in development work, but
+   it wastes space in the release version.
+
+functor ParserGen(structure LrTable : LR_TABLE
+		  structure Stream : STREAM) : LR_PARSER =
+*)
+
+structure LrParser :> LR_PARSER =
+   struct
+      structure LrTable = LrTable
+      structure Stream = Stream
+
+      fun eqT (LrTable.T i, LrTable.T i') = i = i'
+
+      structure Token : TOKEN =
+	struct
+	    structure LrTable = LrTable
+	    datatype ('a,'b) token = TOKEN of LrTable.term * ('a * 'b * 'b)
+	    val sameToken = fn (TOKEN(t,_),TOKEN(t',_)) => eqT (t,t')
+        end
+
+      open LrTable
+      open Token
+
+      val DEBUG1 = false
+      val DEBUG2 = false
+      exception ParseError
+      exception ParseImpossible of int
+
+      structure Fifo :> FIFO =
+        struct
+	  type 'a queue = ('a list * 'a list)
+	  val empty = (nil,nil)
+	  exception Empty
+	  fun get(a::x, y) = (a, (x,y))
+	    | get(nil, nil) = raise Empty
+	    | get(nil, y) = get(rev y, nil)
+ 	  fun put(a,(x,y)) = (x,a::y)
+        end
+
+      type ('a,'b) elem = (state * ('a * 'b * 'b))
+      type ('a,'b) stack = ('a,'b) elem list
+      type ('a,'b) lexv = ('a,'b) token
+      type ('a,'b) lexpair = ('a,'b) lexv * (('a,'b) lexv Stream.stream)
+      type ('a,'b) distanceParse =
+		 ('a,'b) lexpair *
+		 ('a,'b) stack * 
+		 (('a,'b) stack * ('a,'b) lexpair) Fifo.queue *
+		 int ->
+		   ('a,'b) lexpair *
+		   ('a,'b) stack * 
+		   (('a,'b) stack * ('a,'b) lexpair) Fifo.queue *
+		   int *
+		   action option
+
+      type ('a,'b) ecRecord =
+	 {is_keyword : term -> bool,
+          preferred_change : (term list * term list) list,
+	  error : string * 'b * 'b -> unit,
+	  errtermvalue : term -> 'a,
+	  terms : term list,
+	  showTerminal : term -> string,
+	  noShift : term -> bool}
+
+      local 
+	 val print = fn s => TextIO.output(TextIO.stdOut,s)
+	 val println = fn s => (print s; print "\n")
+	 val showState = fn (STATE s) => "STATE " ^ (Int.toString s)
+      in
+        fun printStack(stack: ('a,'b) stack, n: int) =
+         case stack
+           of (state,_) :: rest =>
+                 (print("\t" ^ Int.toString n ^ ": ");
+                  println(showState state);
+                  printStack(rest, n+1))
+            | nil => ()
+                
+        fun prAction showTerminal
+		 (stack as (state,_) :: _, next as (TOKEN (term,_),_), action) =
+             (println "Parse: state stack:";
+              printStack(stack, 0);
+              print("       state="
+                         ^ showState state	
+                         ^ " next="
+                         ^ showTerminal term
+                         ^ " action="
+                        );
+              case action
+                of SHIFT state => println ("SHIFT " ^ (showState state))
+                 | REDUCE i => println ("REDUCE " ^ (Int.toString i))
+                 | ERROR => println "ERROR"
+		 | ACCEPT => println "ACCEPT")
+        | prAction _ (_,_,action) = ()
+     end
+
+    (* ssParse: parser which maintains the queue of (state * lexvalues) in a
+	steady-state.  It takes a table, showTerminal function, saction
+	function, and fixError function.  It parses until an ACCEPT is
+	encountered, or an exception is raised.  When an error is encountered,
+	fixError is called with the arguments of parseStep (lexv,stack,and
+	queue).  It returns the lexv, and a new stack and queue adjusted so
+	that the lexv can be parsed *)
+	
+    val ssParse =
+      fn (table,showTerminal,saction,fixError,arg) =>
+	let val prAction = prAction showTerminal
+	    val action = LrTable.action table
+	    val goto = LrTable.goto table
+	    fun parseStep(args as
+			 (lexPair as (TOKEN (terminal, value as (_,leftPos,_)),
+				      lexer
+				      ),
+			  stack as (state,_) :: _,
+			  queue)) =
+	      let val nextAction = action (state,terminal)
+	          val _ = if DEBUG1 then prAction(stack,lexPair,nextAction)
+			  else ()
+	      in case nextAction
+		 of SHIFT s =>
+		  let val newStack = (s,value) :: stack
+		      val newLexPair = Stream.get lexer
+		      val (_,newQueue) =Fifo.get(Fifo.put((newStack,newLexPair),
+							    queue))
+		  in parseStep(newLexPair,(s,value)::stack,newQueue)
+		  end
+		 | REDUCE i =>
+		     (case saction(i,leftPos,stack,arg)
+		      of (nonterm,value,stack as (state,_) :: _) =>
+		          parseStep(lexPair,(goto(state,nonterm),value)::stack,
+				    queue)
+		       | _ => raise (ParseImpossible 197))
+		 | ERROR => parseStep(fixError args)
+		 | ACCEPT => 
+			(case stack
+			 of (_,(topvalue,_,_)) :: _ =>
+				let val (token,restLexer) = lexPair
+				in (topvalue,Stream.cons(token,restLexer))
+				end
+			  | _ => raise (ParseImpossible 202))
+	      end
+	    | parseStep _ = raise (ParseImpossible 204)
+	in parseStep
+	end
+
+    (*  distanceParse: parse until n tokens are shifted, or accept or
+	error are encountered.  Takes a table, showTerminal function, and
+	semantic action function.  Returns a parser which takes a lexPair
+	(lex result * lexer), a state stack, a queue, and a distance
+	(must be > 0) to parse.  The parser returns a new lex-value, a stack
+	with the nth token shifted on top, a queue, a distance, and action
+	option. *)
+
+    val distanceParse =
+      fn (table,showTerminal,saction,arg) =>
+	let val prAction = prAction showTerminal
+	    val action = LrTable.action table
+	    val goto = LrTable.goto table
+	    fun parseStep(lexPair,stack,queue,0) = (lexPair,stack,queue,0,NONE)
+	      | parseStep(lexPair as (TOKEN (terminal, value as (_,leftPos,_)),
+				      lexer
+				     ),
+			  stack as (state,_) :: _,
+			  queue,distance) =
+	      let val nextAction = action(state,terminal)
+	          val _ = if DEBUG1 then prAction(stack,lexPair,nextAction)
+			  else ()
+	      in case nextAction
+		 of SHIFT s =>
+		  let val newStack = (s,value) :: stack
+		      val newLexPair = Stream.get lexer
+		  in parseStep(newLexPair,(s,value)::stack,
+			       Fifo.put((newStack,newLexPair),queue),distance-1)
+		  end
+		 | REDUCE i =>
+		    (case saction(i,leftPos,stack,arg)
+		      of (nonterm,value,stack as (state,_) :: _) =>
+		         parseStep(lexPair,(goto(state,nonterm),value)::stack,
+				 queue,distance)
+		      | _ => raise (ParseImpossible 240))
+		 | ERROR => (lexPair,stack,queue,distance,SOME nextAction)
+		 | ACCEPT => (lexPair,stack,queue,distance,SOME nextAction)
+	      end
+	   | parseStep _ = raise (ParseImpossible 242)
+	in parseStep : ('_a,'_b) distanceParse 
+	end
+
+(* mkFixError: function to create fixError function which adjusts parser state
+   so that parse may continue in the presence of an error *)
+
+fun mkFixError({is_keyword,terms,errtermvalue,
+	      preferred_change,noShift,
+	      showTerminal,error,...} : ('_a,'_b) ecRecord,
+	     distanceParse : ('_a,'_b) distanceParse,
+	     minAdvance,maxAdvance) 
+
+            (lexv as (TOKEN (term,value as (_,leftPos,_)),_),stack,queue) =
+    let val _ = if DEBUG2 then
+			error("syntax error found at " ^ (showTerminal term),
+			      leftPos,leftPos)
+		else ()
+
+        fun tokAt(t,p) = TOKEN(t,(errtermvalue t,p,p))
+
+	val minDelta = 3
+
+	(* pull all the state * lexv elements from the queue *)
+
+	val stateList = 
+	   let fun f q = let val (elem,newQueue) = Fifo.get q
+			 in elem :: (f newQueue)
+			 end handle Fifo.Empty => nil
+	   in f queue
+	   end
+
+	(* now number elements of stateList, giving distance from
+	   error token *)
+
+	val (_, numStateList) =
+	      List.foldr (fn (a,(num,r)) => (num+1,(a,num)::r)) (0, []) stateList
+
+	(* Represent the set of potential changes as a linked list.
+
+	   Values of datatype Change hold information about a potential change.
+
+	   oper = oper to be applied
+	   pos = the # of the element in stateList that would be altered.
+	   distance = the number of tokens beyond the error token which the
+	     change allows us to parse.
+	   new = new terminal * value pair at that point
+	   orig = original terminal * value pair at the point being changed.
+	 *)
+
+	datatype ('a,'b) change = CHANGE of
+	   {pos : int, distance : int, leftPos: 'b, rightPos: 'b,
+	    new : ('a,'b) lexv list, orig : ('a,'b) lexv list}
+
+
+         val showTerms = concat o map (fn TOKEN(t,_) => " " ^ showTerminal t)
+
+	 val printChange = fn c =>
+	  let val CHANGE {distance,new,orig,pos,...} = c
+	  in (print ("{distance= " ^ (Int.toString distance));
+	      print (",orig ="); print(showTerms orig);
+	      print (",new ="); print(showTerms new);
+	      print (",pos= " ^ (Int.toString pos));
+	      print "}\n")
+	  end
+
+	val printChangeList = app printChange
+
+(* parse: given a lexPair, a stack, and the distance from the error
+   token, return the distance past the error token that we are able to parse.*)
+
+	fun parse (lexPair,stack,queuePos : int) =
+	    case distanceParse(lexPair,stack,Fifo.empty,queuePos+maxAdvance+1)
+             of (_,_,_,distance,SOME ACCEPT) => 
+		        if maxAdvance-distance-1 >= 0 
+			    then maxAdvance 
+			    else maxAdvance-distance-1
+	      | (_,_,_,distance,_) => maxAdvance - distance - 1
+
+(* catList: concatenate results of scanning list *)
+
+	fun catList l f = List.foldr (fn(a,r)=> f a @ r) [] l
+
+        fun keywordsDelta new = if List.exists (fn(TOKEN(t,_))=>is_keyword t) new
+	               then minDelta else 0
+
+        fun tryChange{lex,stack,pos,leftPos,rightPos,orig,new} =
+	     let val lex' = List.foldr (fn (t',p)=>(t',Stream.cons p)) lex new
+		 val distance = parse(lex',stack,pos+length new-length orig)
+	      in if distance >= minAdvance + keywordsDelta new 
+		   then [CHANGE{pos=pos,leftPos=leftPos,rightPos=rightPos,
+				distance=distance,orig=orig,new=new}] 
+		   else []
+	     end
+
+
+(* tryDelete: Try to delete n terminals.
+              Return single-element [success] or nil.
+	      Do not delete unshiftable terminals. *)
+
+
+    fun tryDelete n ((stack,lexPair as (TOKEN(term,(_,l,r)),_)),qPos) =
+	let fun del(0,accum,left,right,lexPair) =
+	          tryChange{lex=lexPair,stack=stack,
+			    pos=qPos,leftPos=left,rightPos=right,
+			    orig=rev accum, new=[]}
+	      | del(n,accum,left,right,(tok as TOKEN(term,(_,_,r)),lexer)) =
+		   if noShift term then []
+		   else del(n-1,tok::accum,left,r,Stream.get lexer)
+         in del(n,[],l,r,lexPair)
+        end
+
+(* tryInsert: try to insert tokens before the current terminal;
+       return a list of the successes  *)
+
+        fun tryInsert((stack,lexPair as (TOKEN(_,(_,l,_)),_)),queuePos) =
+	       catList terms (fn t =>
+		 tryChange{lex=lexPair,stack=stack,
+			   pos=queuePos,orig=[],new=[tokAt(t,l)],
+			   leftPos=l,rightPos=l})
+			      
+(* trySubst: try to substitute tokens for the current terminal;
+       return a list of the successes  *)
+
+        fun trySubst ((stack,lexPair as (orig as TOKEN (term,(_,l,r)),lexer)),
+		      queuePos) =
+	      if noShift term then []
+	      else
+		  catList terms (fn t =>
+		      tryChange{lex=Stream.get lexer,stack=stack,
+				pos=queuePos,
+				leftPos=l,rightPos=r,orig=[orig],
+				new=[tokAt(t,r)]})
+
+     (* do_delete(toks,lexPair) tries to delete tokens "toks" from "lexPair".
+         If it succeeds, returns SOME(toks',l,r,lp), where
+	     toks' is the actual tokens (with positions and values) deleted,
+	     (l,r) are the (leftmost,rightmost) position of toks', 
+	     lp is what remains of the stream after deletion 
+     *)
+        fun do_delete(nil,lp as (TOKEN(_,(_,l,_)),_)) = SOME(nil,l,l,lp)
+          | do_delete([t],(tok as TOKEN(t',(_,l,r)),lp')) =
+	       if eqT (t, t')
+		   then SOME([tok],l,r,Stream.get lp')
+                   else NONE
+          | do_delete(t::rest,(tok as TOKEN(t',(_,l,r)),lp')) =
+	       if eqT (t,t')
+		   then case do_delete(rest,Stream.get lp')
+                         of SOME(deleted,l',r',lp'') =>
+			       SOME(tok::deleted,l,r',lp'')
+			  | NONE => NONE
+		   else NONE
+			     
+        fun tryPreferred((stack,lexPair),queuePos) =
+	    catList preferred_change (fn (delete,insert) =>
+	       if List.exists noShift delete then [] (* should give warning at
+						 parser-generation time *)
+               else case do_delete(delete,lexPair)
+                     of SOME(deleted,l,r,lp) => 
+			    tryChange{lex=lp,stack=stack,pos=queuePos,
+				      leftPos=l,rightPos=r,orig=deleted,
+				      new=map (fn t=>(tokAt(t,r))) insert}
+		      | NONE => [])
+
+	val changes = catList numStateList tryPreferred @
+	                catList numStateList tryInsert @
+			  catList numStateList trySubst @
+			    catList numStateList (tryDelete 1) @
+			      catList numStateList (tryDelete 2) @
+			        catList numStateList (tryDelete 3)
+
+	val findMaxDist = fn l => 
+	  foldr (fn (CHANGE {distance,...},high) => Int.max(distance,high)) 0 l
+
+(* maxDist: max distance past error taken that we could parse *)
+
+	val maxDist = findMaxDist changes
+
+(* remove changes which did not parse maxDist tokens past the error token *)
+
+        val changes = catList changes 
+	      (fn(c as CHANGE{distance,...}) => 
+		  if distance=maxDist then [c] else [])
+
+      in case changes 
+	  of (l as change :: _) =>
+	      let fun print_msg (CHANGE {new,orig,leftPos,rightPos,...}) =
+		  let val s = 
+		      case (orig,new)
+			  of (_::_,[]) => "deleting " ^ (showTerms orig)
+	                   | ([],_::_) => "inserting " ^ (showTerms new)
+			   | _ => "replacing " ^ (showTerms orig) ^
+				 " with " ^ (showTerms new)
+		  in error ("syntax error: " ^ s,leftPos,rightPos)
+		  end
+		   
+		  val _ = 
+		      (if length l > 1 andalso DEBUG2 then
+			   (print "multiple fixes possible; could fix it by:\n";
+			    app print_msg l;
+			    print "chosen correction:\n")
+		       else ();
+		       print_msg change)
+
+		  (* findNth: find nth queue entry from the error
+		   entry.  Returns the Nth queue entry and the  portion of
+		   the queue from the beginning to the nth-1 entry.  The
+		   error entry is at the end of the queue.
+
+		   Examples:
+
+		   queue = a b c d e
+		   findNth 0 = (e,a b c d)
+		   findNth 1 =  (d,a b c)
+		   *)
+
+		  val findNth = fn n =>
+		      let fun f (h::t,0) = (h,rev t)
+			    | f (h::t,n) = f(t,n-1)
+			    | f (nil,_) = let exception FindNth
+					  in raise FindNth
+					  end
+		      in f (rev stateList,n)
+		      end
+		
+		  val CHANGE {pos,orig,new,...} = change
+		  val (last,queueFront) = findNth pos
+		  val (stack,lexPair) = last
+
+		  val lp1 = foldl(fn (_,(_,r)) => Stream.get r) lexPair orig
+		  val lp2 = foldr(fn(t,r)=>(t,Stream.cons r)) lp1 new
+
+		  val restQueue = 
+		      Fifo.put((stack,lp2),
+			       foldl Fifo.put Fifo.empty queueFront)
+
+		  val (lexPair,stack,queue,_,_) =
+		      distanceParse(lp2,stack,restQueue,pos)
+
+	      in (lexPair,stack,queue)
+	      end
+	| nil => (error("syntax error found at " ^ (showTerminal term),
+			leftPos,leftPos); raise ParseError)
+    end
+
+   val parse = fn {arg,table,lexer,saction,void,lookahead,
+		   ec=ec as {showTerminal,...} : ('_a,'_b) ecRecord} =>
+	let val distance = 15   (* defer distance tokens *)
+	    val minAdvance = 1  (* must parse at least 1 token past error *)
+	    val maxAdvance = Int.max(lookahead,0)(* max distance for parse check *)
+	    val lexPair = Stream.get lexer
+	    val (TOKEN (_,(_,leftPos,_)),_) = lexPair
+	    val startStack = [(initialState table,(void,leftPos,leftPos))]
+	    val startQueue = Fifo.put((startStack,lexPair),Fifo.empty)
+	    val distanceParse = distanceParse(table,showTerminal,saction,arg)
+	    val fixError = mkFixError(ec,distanceParse,minAdvance,maxAdvance)
+	    val ssParse = ssParse(table,showTerminal,saction,fixError,arg)
+	    fun loop (lexPair,stack,queue,_,SOME ACCEPT) =
+		   ssParse(lexPair,stack,queue)
+	      | loop (lexPair,stack,queue,0,_) = ssParse(lexPair,stack,queue)
+	      | loop (lexPair,stack,queue,distance,SOME ERROR) =
+		 let val (lexPair,stack,queue) = fixError(lexPair,stack,queue)
+		 in loop (distanceParse(lexPair,stack,queue,distance))
+		 end
+	      | loop _ = let exception ParseInternal
+			 in raise ParseInternal
+			 end
+	in loop (distanceParse(lexPair,startStack,startQueue,distance))
+	end
+ end;
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/ml-yacc/lib/sources.cm	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,30 @@
+(* sources file for ML-Yacc library *)
+
+Library
+
+signature STREAM
+signature LR_TABLE
+signature TOKEN
+signature LR_PARSER
+signature LEXER
+signature ARG_LEXER
+signature PARSER_DATA
+signature PARSER
+signature ARG_PARSER
+functor Join
+functor JoinWithArg
+structure LrTable
+structure Stream
+structure LrParser
+
+is
+
+#if defined(NEW_CM)
+  $basis.cm
+#endif
+
+base.sig
+join.sml
+lrtable.sml
+stream.sml
+parser2.sml	(* error correcting version *)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/ml-yacc/lib/stream.sml	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,19 @@
+(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
+
+(* Stream: a structure implementing a lazy stream.  The signature STREAM
+   is found in base.sig *)
+
+structure Stream :> STREAM =
+struct
+   datatype 'a str = EVAL of 'a * 'a str ref | UNEVAL of (unit->'a)
+
+   type 'a stream = 'a str ref
+
+   fun get(ref(EVAL t)) = t
+     | get(s as ref(UNEVAL f)) = 
+	    let val t = (f(), ref(UNEVAL f)) in s := EVAL t; t end
+
+   fun streamify f = ref(UNEVAL f)
+   fun cons(a,s) = ref(EVAL(a,s))
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/ml_yacc_lib.ML	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,1064 @@
+
+(******************************************************************)
+(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
+(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
+(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
+(******************************************************************)
+
+print_depth 0;
+
+(*
+  This file is generated from the contents of ML-Yacc's lib directory.
+  ML-Yacc's COPYRIGHT-file contents follow:
+
+  ML-YACC COPYRIGHT NOTICE, LICENSE AND DISCLAIMER.
+  
+  Copyright 1989, 1990 by David R. Tarditi Jr. and Andrew W. Appel
+  
+  Permission to use, copy, modify, and distribute this software and its
+  documentation for any purpose and without fee is hereby granted,
+  provided that the above copyright notice appear in all copies and that
+  both the copyright notice and this permission notice and warranty
+  disclaimer appear in supporting documentation, and that the names of
+  David R. Tarditi Jr. and Andrew W. Appel not be used in advertising
+  or publicity pertaining to distribution of the software without
+  specific, written prior permission.
+  
+  David R. Tarditi Jr. and Andrew W. Appel disclaim all warranties with regard to
+  this software, including all implied warranties of merchantability and fitness.
+  In no event shall David R. Tarditi Jr. and Andrew W. Appel be liable for any
+  special, indirect or consequential damages or any damages whatsoever resulting
+  from loss of use, data or profits, whether in an action of contract, negligence
+  or other tortious action, arising out of or in connection with the use or
+  performance of this software.
+*)
+
+(**** Original file: base.sig ****)
+
+(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
+
+(* base.sig: Base signature file for SML-Yacc.  This file contains signatures
+   that must be loaded before any of the files produced by ML-Yacc are loaded
+*)
+
+(* STREAM: signature for a lazy stream.*)
+
+signature STREAM =
+ sig type 'xa stream
+     val streamify : (unit -> '_a) -> '_a stream
+     val cons : '_a * '_a stream -> '_a stream
+     val get : '_a stream -> '_a * '_a stream
+ end
+
+(* LR_TABLE: signature for an LR Table.
+
+   The list of actions and gotos passed to mkLrTable must be ordered by state
+   number. The values for state 0 are the first in the list, the values for
+    state 1 are next, etc.
+*)
+
+signature LR_TABLE =
+    sig
+        datatype ('a,'b) pairlist = EMPTY | PAIR of 'a * 'b * ('a,'b) pairlist
+	datatype state = STATE of int
+	datatype term = T of int
+	datatype nonterm = NT of int
+	datatype action = SHIFT of state
+			| REDUCE of int
+			| ACCEPT
+			| ERROR
+	type table
+	
+	val numStates : table -> int
+	val numRules : table -> int
+	val describeActions : table -> state ->
+				(term,action) pairlist * action
+	val describeGoto : table -> state -> (nonterm,state) pairlist
+	val action : table -> state * term -> action
+	val goto : table -> state * nonterm -> state
+	val initialState : table -> state
+	exception Goto of state * nonterm
+
+	val mkLrTable : {actions : ((term,action) pairlist * action) array,
+			 gotos : (nonterm,state) pairlist array,
+			 numStates : int, numRules : int,
+			 initialState : state} -> table
+    end
+
+(* TOKEN: signature revealing the internal structure of a token. This signature
+   TOKEN distinct from the signature {parser name}_TOKENS produced by ML-Yacc.
+   The {parser name}_TOKENS structures contain some types and functions to
+    construct tokens from values and positions.
+
+   The representation of token was very carefully chosen here to allow the
+   polymorphic parser to work without knowing the types of semantic values
+   or line numbers.
+
+   This has had an impact on the TOKENS structure produced by SML-Yacc, which
+   is a structure parameter to lexer functors.  We would like to have some
+   type 'a token which functions to construct tokens would create.  A
+   constructor function for a integer token might be
+
+	  INT: int * 'a * 'a -> 'a token.
+ 
+   This is not possible because we need to have tokens with the representation
+   given below for the polymorphic parser.
+
+   Thus our constructur functions for tokens have the form:
+
+	  INT: int * 'a * 'a -> (svalue,'a) token
+
+   This in turn has had an impact on the signature that lexers for SML-Yacc
+   must match and the types that a user must declare in the user declarations
+   section of lexers.
+*)
+
+signature TOKEN =
+    sig
+	structure LrTable : LR_TABLE
+        datatype ('a,'b) token = TOKEN of LrTable.term * ('a * 'b * 'b)
+	val sameToken : ('a,'b) token * ('a,'b) token -> bool
+    end
+
+(* LR_PARSER: signature for a polymorphic LR parser *)
+
+signature LR_PARSER =
+    sig
+	structure Stream: STREAM
+	structure LrTable : LR_TABLE
+	structure Token : TOKEN
+
+	sharing LrTable = Token.LrTable
+
+	exception ParseError
+
+	val parse : {table : LrTable.table,
+		     lexer : ('_b,'_c) Token.token Stream.stream,
+		     arg: 'arg,
+		     saction : int *
+			       '_c *
+				(LrTable.state * ('_b * '_c * '_c)) list * 
+				'arg ->
+				     LrTable.nonterm *
+				     ('_b * '_c * '_c) *
+				     ((LrTable.state *('_b * '_c * '_c)) list),
+		     void : '_b,
+		     ec : { is_keyword : LrTable.term -> bool,
+			    noShift : LrTable.term -> bool,
+			    preferred_change : (LrTable.term list * LrTable.term list) list,
+			    errtermvalue : LrTable.term -> '_b,
+			    showTerminal : LrTable.term -> string,
+			    terms: LrTable.term list,
+			    error : string * '_c * '_c -> unit
+			   },
+		     lookahead : int  (* max amount of lookahead used in *)
+				      (* error correction *)
+			} -> '_b *
+			     (('_b,'_c) Token.token Stream.stream)
+    end
+
+(* LEXER: a signature that most lexers produced for use with SML-Yacc's
+   output will match.  The user is responsible for declaring type token,
+   type pos, and type svalue in the UserDeclarations section of a lexer.
+
+   Note that type token is abstract in the lexer.  This allows SML-Yacc to
+   create a TOKENS signature for use with lexers produced by ML-Lex that
+   treats the type token abstractly.  Lexers that are functors parametrized by
+   a Tokens structure matching a TOKENS signature cannot examine the structure
+   of tokens.
+*)
+
+signature LEXER =
+   sig
+       structure UserDeclarations :
+	   sig
+	        type ('a,'b) token
+		type pos
+		type svalue
+	   end
+	val makeLexer : (int -> string) -> unit -> 
+         (UserDeclarations.svalue,UserDeclarations.pos) UserDeclarations.token
+   end
+
+(* ARG_LEXER: the %arg option of ML-Lex allows users to produce lexers which
+   also take an argument before yielding a function from unit to a token
+*)
+
+signature ARG_LEXER =
+   sig
+       structure UserDeclarations :
+	   sig
+	        type ('a,'b) token
+		type pos
+		type svalue
+		type arg
+	   end
+	val makeLexer : (int -> string) -> UserDeclarations.arg -> unit -> 
+         (UserDeclarations.svalue,UserDeclarations.pos) UserDeclarations.token
+   end
+
+(* PARSER_DATA: the signature of ParserData structures in {parser name}LrValsFun
+   produced by  SML-Yacc.  All such structures match this signature.  
+
+   The {parser name}LrValsFun produces a structure which contains all the values
+   except for the lexer needed to call the polymorphic parser mentioned
+   before.
+
+*)
+
+signature PARSER_DATA =
+   sig
+        (* the type of line numbers *)
+
+	type pos
+
+	(* the type of semantic values *)
+
+	type svalue
+
+         (* the type of the user-supplied argument to the parser *)
+ 	type arg
+ 
+	(* the intended type of the result of the parser.  This value is
+	   produced by applying extract from the structure Actions to the
+	   final semantic value resultiing from a parse.
+	 *)
+
+	type result
+
+	structure LrTable : LR_TABLE
+	structure Token : TOKEN
+	sharing Token.LrTable = LrTable
+
+	(* structure Actions contains the functions which mantain the
+	   semantic values stack in the parser.  Void is used to provide
+	   a default value for the semantic stack.
+	 *)
+
+	structure Actions : 
+	  sig
+	      val actions : int * pos *
+		   (LrTable.state * (svalue * pos * pos)) list * arg->
+		         LrTable.nonterm * (svalue * pos * pos) *
+			 ((LrTable.state *(svalue * pos * pos)) list)
+	      val void : svalue
+	      val extract : svalue -> result
+	  end
+
+	(* structure EC contains information used to improve error
+	   recovery in an error-correcting parser *)
+
+	structure EC :
+	   sig
+	     val is_keyword : LrTable.term -> bool
+	     val noShift : LrTable.term -> bool
+ 	     val preferred_change : (LrTable.term list * LrTable.term list) list
+	     val errtermvalue : LrTable.term -> svalue
+	     val showTerminal : LrTable.term -> string
+	     val terms: LrTable.term list
+	   end
+
+	(* table is the LR table for the parser *)
+
+	val table : LrTable.table
+    end
+
+(* signature PARSER is the signature that most user parsers created by 
+   SML-Yacc will match.
+*)
+
+signature PARSER =
+    sig
+        structure Token : TOKEN
+	structure Stream : STREAM
+	exception ParseError
+
+	(* type pos is the type of line numbers *)
+
+	type pos
+
+	(* type result is the type of the result from the parser *)
+
+	type result
+
+         (* the type of the user-supplied argument to the parser *)
+ 	type arg
+	
+	(* type svalue is the type of semantic values for the semantic value
+	   stack
+	 *)
+
+	type svalue
+
+	(* val makeLexer is used to create a stream of tokens for the parser *)
+
+	val makeLexer : (int -> string) ->
+			 (svalue,pos) Token.token Stream.stream
+
+	(* val parse takes a stream of tokens and a function to TextIO.print
+	   errors and returns a value of type result and a stream containing
+	   the unused tokens
+	 *)
+
+	val parse : int * ((svalue,pos) Token.token Stream.stream) *
+		    (string * pos * pos -> unit) * arg ->
+				result * (svalue,pos) Token.token Stream.stream
+
+	val sameToken : (svalue,pos) Token.token * (svalue,pos) Token.token ->
+				bool
+     end
+
+(* signature ARG_PARSER is the signature that will be matched by parsers whose
+    lexer takes an additional argument.
+*)
+
+signature ARG_PARSER = 
+    sig
+        structure Token : TOKEN
+	structure Stream : STREAM
+	exception ParseError
+
+	type arg
+	type lexarg
+	type pos
+	type result
+	type svalue
+
+	val makeLexer : (int -> string) -> lexarg ->
+			 (svalue,pos) Token.token Stream.stream
+	val parse : int * ((svalue,pos) Token.token Stream.stream) *
+		    (string * pos * pos -> unit) * arg ->
+				result * (svalue,pos) Token.token Stream.stream
+
+	val sameToken : (svalue,pos) Token.token * (svalue,pos) Token.token ->
+				bool
+     end
+
+
+(**** Original file: join.sml ****)
+
+(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
+
+(* functor Join creates a user parser by putting together a Lexer structure,
+   an LrValues structure, and a polymorphic parser structure.  Note that
+   the Lexer and LrValues structure must share the type pos (i.e. the type
+   of line numbers), the type svalues for semantic values, and the type
+   of tokens.
+*)
+
+functor Join(structure Lex : LEXER
+	     structure ParserData: PARSER_DATA
+	     structure LrParser : LR_PARSER
+	     sharing ParserData.LrTable = LrParser.LrTable
+	     sharing ParserData.Token = LrParser.Token
+	     sharing type Lex.UserDeclarations.svalue = ParserData.svalue
+	     sharing type Lex.UserDeclarations.pos = ParserData.pos
+	     sharing type Lex.UserDeclarations.token = ParserData.Token.token)
+		 : PARSER =
+struct
+    structure Token = ParserData.Token
+    structure Stream = LrParser.Stream
+ 
+    exception ParseError = LrParser.ParseError
+
+    type arg = ParserData.arg
+    type pos = ParserData.pos
+    type result = ParserData.result
+    type svalue = ParserData.svalue
+    val makeLexer = LrParser.Stream.streamify o Lex.makeLexer
+    val parse = fn (lookahead,lexer,error,arg) =>
+	(fn (a,b) => (ParserData.Actions.extract a,b))
+     (LrParser.parse {table = ParserData.table,
+	        lexer=lexer,
+		lookahead=lookahead,
+		saction = ParserData.Actions.actions,
+		arg=arg,
+		void= ParserData.Actions.void,
+	        ec = {is_keyword = ParserData.EC.is_keyword,
+		      noShift = ParserData.EC.noShift,
+		      preferred_change = ParserData.EC.preferred_change,
+		      errtermvalue = ParserData.EC.errtermvalue,
+		      error=error,
+		      showTerminal = ParserData.EC.showTerminal,
+		      terms = ParserData.EC.terms}}
+      )
+     val sameToken = Token.sameToken
+end
+
+(* functor JoinWithArg creates a variant of the parser structure produced 
+   above.  In this case, the makeLexer take an additional argument before
+   yielding a value of type unit -> (svalue,pos) token
+ *)
+
+functor JoinWithArg(structure Lex : ARG_LEXER
+	     structure ParserData: PARSER_DATA
+	     structure LrParser : LR_PARSER
+	     sharing ParserData.LrTable = LrParser.LrTable
+	     sharing ParserData.Token = LrParser.Token
+	     sharing type Lex.UserDeclarations.svalue = ParserData.svalue
+	     sharing type Lex.UserDeclarations.pos = ParserData.pos
+	     sharing type Lex.UserDeclarations.token = ParserData.Token.token)
+		 : ARG_PARSER  =
+struct
+    structure Token = ParserData.Token
+    structure Stream = LrParser.Stream
+
+    exception ParseError = LrParser.ParseError
+
+    type arg = ParserData.arg
+    type lexarg = Lex.UserDeclarations.arg
+    type pos = ParserData.pos
+    type result = ParserData.result
+    type svalue = ParserData.svalue
+
+    val makeLexer = fn s => fn arg =>
+		 LrParser.Stream.streamify (Lex.makeLexer s arg)
+    val parse = fn (lookahead,lexer,error,arg) =>
+	(fn (a,b) => (ParserData.Actions.extract a,b))
+     (LrParser.parse {table = ParserData.table,
+	        lexer=lexer,
+		lookahead=lookahead,
+		saction = ParserData.Actions.actions,
+		arg=arg,
+		void= ParserData.Actions.void,
+	        ec = {is_keyword = ParserData.EC.is_keyword,
+		      noShift = ParserData.EC.noShift,
+		      preferred_change = ParserData.EC.preferred_change,
+		      errtermvalue = ParserData.EC.errtermvalue,
+		      error=error,
+		      showTerminal = ParserData.EC.showTerminal,
+		      terms = ParserData.EC.terms}}
+      )
+    val sameToken = Token.sameToken
+end;
+
+(**** Original file: lrtable.sml ****)
+
+(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
+structure LrTable : LR_TABLE = 
+    struct
+	open Array List
+	infix 9 sub
+	datatype ('a,'b) pairlist = EMPTY
+				  | PAIR of 'a * 'b * ('a,'b) pairlist
+	datatype term = T of int
+	datatype nonterm = NT of int
+	datatype state = STATE of int
+	datatype action = SHIFT of state
+			| REDUCE of int (* rulenum from grammar *)
+			| ACCEPT
+			| ERROR
+	exception Goto of state * nonterm
+	type table = {states: int, rules : int,initialState: state,
+		      action: ((term,action) pairlist * action) array,
+		      goto :  (nonterm,state) pairlist array}
+	val numStates = fn ({states,...} : table) => states
+	val numRules = fn ({rules,...} : table) => rules
+	val describeActions =
+	   fn ({action,...} : table) => 
+	           fn (STATE s) => action sub s
+	val describeGoto =
+	   fn ({goto,...} : table) =>
+	           fn (STATE s) => goto sub s
+	fun findTerm (T term,row,default) =
+	    let fun find (PAIR (T key,data,r)) =
+		       if key < term then find r
+		       else if key=term then data
+		       else default
+		   | find EMPTY = default
+	    in find row
+	    end
+	fun findNonterm (NT nt,row) =
+	    let fun find (PAIR (NT key,data,r)) =
+		       if key < nt then find r
+		       else if key=nt then SOME data
+		       else NONE
+		   | find EMPTY = NONE
+	    in find row
+	    end
+	val action = fn ({action,...} : table) =>
+		fn (STATE state,term) =>
+		  let val (row,default) = action sub state
+		  in findTerm(term,row,default)
+		  end
+	val goto = fn ({goto,...} : table) =>
+			fn (a as (STATE state,nonterm)) =>
+			  case findNonterm(nonterm,goto sub state)
+			  of SOME state => state
+			   | NONE => raise (Goto a)
+	val initialState = fn ({initialState,...} : table) => initialState
+	val mkLrTable = fn {actions,gotos,initialState,numStates,numRules} =>
+	     ({action=actions,goto=gotos,
+	       states=numStates,
+	       rules=numRules,
+               initialState=initialState} : table)
+end;
+
+(**** Original file: stream.sml ****)
+
+(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
+
+(* Stream: a structure implementing a lazy stream.  The signature STREAM
+   is found in base.sig *)
+
+structure Stream :> STREAM =
+struct
+   datatype 'a str = EVAL of 'a * 'a str Unsynchronized.ref | UNEVAL of (unit->'a)
+
+   type 'a stream = 'a str Unsynchronized.ref
+
+   fun get(Unsynchronized.ref(EVAL t)) = t
+     | get(s as Unsynchronized.ref(UNEVAL f)) = 
+	    let val t = (f(), Unsynchronized.ref(UNEVAL f)) in s := EVAL t; t end
+
+   fun streamify f = Unsynchronized.ref(UNEVAL f)
+   fun cons(a,s) = Unsynchronized.ref(EVAL(a,s))
+
+end;
+
+(**** Original file: parser2.sml ****)
+
+(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
+
+(* parser.sml:  This is a parser driver for LR tables with an error-recovery
+   routine added to it.  The routine used is described in detail in this
+   article:
+
+	'A Practical Method for LR and LL Syntactic Error Diagnosis and
+	 Recovery', by M. Burke and G. Fisher, ACM Transactions on
+	 Programming Langauges and Systems, Vol. 9, No. 2, April 1987,
+	 pp. 164-197.
+
+    This program is an implementation is the partial, deferred method discussed
+    in the article.  The algorithm and data structures used in the program
+    are described below.  
+
+    This program assumes that all semantic actions are delayed.  A semantic
+    action should produce a function from unit -> value instead of producing the
+    normal value.  The parser returns the semantic value on the top of the
+    stack when accept is encountered.  The user can deconstruct this value
+    and apply the unit -> value function in it to get the answer.
+
+    It also assumes that the lexer is a lazy stream.
+
+    Data Structures:
+    ----------------
+	
+	* The parser:
+
+	   The state stack has the type
+
+		 (state * (semantic value * line # * line #)) list
+
+	   The parser keeps a queue of (state stack * lexer pair).  A lexer pair
+	 consists of a terminal * value pair and a lexer.  This allows the 
+	 parser to reconstruct the states for terminals to the left of a
+	 syntax error, and attempt to make error corrections there.
+
+	   The queue consists of a pair of lists (x,y).  New additions to
+	 the queue are cons'ed onto y.  The first element of x is the top
+	 of the queue.  If x is nil, then y is reversed and used
+	 in place of x.
+
+    Algorithm:
+    ----------
+
+	* The steady-state parser:  
+
+	    This parser keeps the length of the queue of state stacks at
+	a steady state by always removing an element from the front when
+	another element is placed on the end.
+
+	    It has these arguments:
+
+	   stack: current stack
+	   queue: value of the queue
+	   lexPair ((terminal,value),lex stream)
+
+	When SHIFT is encountered, the state to shift to and the value are
+	are pushed onto the state stack.  The state stack and lexPair are
+	placed on the queue.  The front element of the queue is removed.
+
+	When REDUCTION is encountered, the rule is applied to the current
+	stack to yield a triple (nonterm,value,new stack).  A new
+	stack is formed by adding (goto(top state of stack,nonterm),value)
+	to the stack.
+
+	When ACCEPT is encountered, the top value from the stack and the
+	lexer are returned.
+
+	When an ERROR is encountered, fixError is called.  FixError
+	takes the arguments to the parser, fixes the error if possible and
+        returns a new set of arguments.
+
+	* The distance-parser:
+
+	This parser includes an additional argument distance.  It pushes
+	elements on the queue until it has parsed distance tokens, or an
+	ACCEPT or ERROR occurs.  It returns a stack, lexer, the number of
+	tokens left unparsed, a queue, and an action option.
+*)
+
+signature FIFO = 
+  sig type 'a queue
+      val empty : 'a queue
+      exception Empty
+      val get : 'a queue -> 'a * 'a queue
+      val put : 'a * 'a queue -> 'a queue
+  end
+
+(* drt (12/15/89) -- the functor should be used in development work, but
+   it wastes space in the release version.
+
+functor ParserGen(structure LrTable : LR_TABLE
+		  structure Stream : STREAM) : LR_PARSER =
+*)
+
+structure LrParser :> LR_PARSER =
+   struct
+      structure LrTable = LrTable
+      structure Stream = Stream
+
+      fun eqT (LrTable.T i, LrTable.T i') = i = i'
+
+      structure Token : TOKEN =
+	struct
+	    structure LrTable = LrTable
+	    datatype ('a,'b) token = TOKEN of LrTable.term * ('a * 'b * 'b)
+	    val sameToken = fn (TOKEN(t,_),TOKEN(t',_)) => eqT (t,t')
+        end
+
+      open LrTable
+      open Token
+
+      val DEBUG1 = false
+      val DEBUG2 = false
+      exception ParseError
+      exception ParseImpossible of int
+
+      structure Fifo :> FIFO =
+        struct
+	  type 'a queue = ('a list * 'a list)
+	  val empty = (nil,nil)
+	  exception Empty
+	  fun get(a::x, y) = (a, (x,y))
+	    | get(nil, nil) = raise Empty
+	    | get(nil, y) = get(rev y, nil)
+ 	  fun put(a,(x,y)) = (x,a::y)
+        end
+
+      type ('a,'b) elem = (state * ('a * 'b * 'b))
+      type ('a,'b) stack = ('a,'b) elem list
+      type ('a,'b) lexv = ('a,'b) token
+      type ('a,'b) lexpair = ('a,'b) lexv * (('a,'b) lexv Stream.stream)
+      type ('a,'b) distanceParse =
+		 ('a,'b) lexpair *
+		 ('a,'b) stack * 
+		 (('a,'b) stack * ('a,'b) lexpair) Fifo.queue *
+		 int ->
+		   ('a,'b) lexpair *
+		   ('a,'b) stack * 
+		   (('a,'b) stack * ('a,'b) lexpair) Fifo.queue *
+		   int *
+		   action option
+
+      type ('a,'b) ecRecord =
+	 {is_keyword : term -> bool,
+          preferred_change : (term list * term list) list,
+	  error : string * 'b * 'b -> unit,
+	  errtermvalue : term -> 'a,
+	  terms : term list,
+	  showTerminal : term -> string,
+	  noShift : term -> bool}
+
+      local 
+	 
+	 val println = fn s => (TextIO.print s; TextIO.print "\n")
+	 val showState = fn (STATE s) => "STATE " ^ (Int.toString s)
+      in
+        fun printStack(stack: ('a,'b) stack, n: int) =
+         case stack
+           of (state,_) :: rest =>
+                 (TextIO.print("\t" ^ Int.toString n ^ ": ");
+                  println(showState state);
+                  printStack(rest, n+1))
+            | nil => ()
+                
+        fun prAction showTerminal
+		 (stack as (state,_) :: _, next as (TOKEN (term,_),_), action) =
+             (println "Parse: state stack:";
+              printStack(stack, 0);
+              TextIO.print("       state="
+                         ^ showState state	
+                         ^ " next="
+                         ^ showTerminal term
+                         ^ " action="
+                        );
+              case action
+                of SHIFT state => println ("SHIFT " ^ (showState state))
+                 | REDUCE i => println ("REDUCE " ^ (Int.toString i))
+                 | ERROR => println "ERROR"
+		 | ACCEPT => println "ACCEPT")
+        | prAction _ (_,_,action) = ()
+     end
+
+    (* ssParse: parser which maintains the queue of (state * lexvalues) in a
+	steady-state.  It takes a table, showTerminal function, saction
+	function, and fixError function.  It parses until an ACCEPT is
+	encountered, or an exception is raised.  When an error is encountered,
+	fixError is called with the arguments of parseStep (lexv,stack,and
+	queue).  It returns the lexv, and a new stack and queue adjusted so
+	that the lexv can be parsed *)
+	
+    val ssParse =
+      fn (table,showTerminal,saction,fixError,arg) =>
+	let val prAction = prAction showTerminal
+	    val action = LrTable.action table
+	    val goto = LrTable.goto table
+	    fun parseStep(args as
+			 (lexPair as (TOKEN (terminal, value as (_,leftPos,_)),
+				      lexer
+				      ),
+			  stack as (state,_) :: _,
+			  queue)) =
+	      let val nextAction = action (state,terminal)
+	          val _ = if DEBUG1 then prAction(stack,lexPair,nextAction)
+			  else ()
+	      in case nextAction
+		 of SHIFT s =>
+		  let val newStack = (s,value) :: stack
+		      val newLexPair = Stream.get lexer
+		      val (_,newQueue) =Fifo.get(Fifo.put((newStack,newLexPair),
+							    queue))
+		  in parseStep(newLexPair,(s,value)::stack,newQueue)
+		  end
+		 | REDUCE i =>
+		     (case saction(i,leftPos,stack,arg)
+		      of (nonterm,value,stack as (state,_) :: _) =>
+		          parseStep(lexPair,(goto(state,nonterm),value)::stack,
+				    queue)
+		       | _ => raise (ParseImpossible 197))
+		 | ERROR => parseStep(fixError args)
+		 | ACCEPT => 
+			(case stack
+			 of (_,(topvalue,_,_)) :: _ =>
+				let val (token,restLexer) = lexPair
+				in (topvalue,Stream.cons(token,restLexer))
+				end
+			  | _ => raise (ParseImpossible 202))
+	      end
+	    | parseStep _ = raise (ParseImpossible 204)
+	in parseStep
+	end
+
+    (*  distanceParse: parse until n tokens are shifted, or accept or
+	error are encountered.  Takes a table, showTerminal function, and
+	semantic action function.  Returns a parser which takes a lexPair
+	(lex result * lexer), a state stack, a queue, and a distance
+	(must be > 0) to parse.  The parser returns a new lex-value, a stack
+	with the nth token shifted on top, a queue, a distance, and action
+	option. *)
+
+    val distanceParse =
+      fn (table,showTerminal,saction,arg) =>
+	let val prAction = prAction showTerminal
+	    val action = LrTable.action table
+	    val goto = LrTable.goto table
+	    fun parseStep(lexPair,stack,queue,0) = (lexPair,stack,queue,0,NONE)
+	      | parseStep(lexPair as (TOKEN (terminal, value as (_,leftPos,_)),
+				      lexer
+				     ),
+			  stack as (state,_) :: _,
+			  queue,distance) =
+	      let val nextAction = action(state,terminal)
+	          val _ = if DEBUG1 then prAction(stack,lexPair,nextAction)
+			  else ()
+	      in case nextAction
+		 of SHIFT s =>
+		  let val newStack = (s,value) :: stack
+		      val newLexPair = Stream.get lexer
+		  in parseStep(newLexPair,(s,value)::stack,
+			       Fifo.put((newStack,newLexPair),queue),distance-1)
+		  end
+		 | REDUCE i =>
+		    (case saction(i,leftPos,stack,arg)
+		      of (nonterm,value,stack as (state,_) :: _) =>
+		         parseStep(lexPair,(goto(state,nonterm),value)::stack,
+				 queue,distance)
+		      | _ => raise (ParseImpossible 240))
+		 | ERROR => (lexPair,stack,queue,distance,SOME nextAction)
+		 | ACCEPT => (lexPair,stack,queue,distance,SOME nextAction)
+	      end
+	   | parseStep _ = raise (ParseImpossible 242)
+	in parseStep : ('_a,'_b) distanceParse 
+	end
+
+(* mkFixError: function to create fixError function which adjusts parser state
+   so that parse may continue in the presence of an error *)
+
+fun mkFixError({is_keyword,terms,errtermvalue,
+	      preferred_change,noShift,
+	      showTerminal,error,...} : ('_a,'_b) ecRecord,
+	     distanceParse : ('_a,'_b) distanceParse,
+	     minAdvance,maxAdvance) 
+
+            (lexv as (TOKEN (term,value as (_,leftPos,_)),_),stack,queue) =
+    let val _ = if DEBUG2 then
+			error("syntax error found at " ^ (showTerminal term),
+			      leftPos,leftPos)
+		else ()
+
+        fun tokAt(t,p) = TOKEN(t,(errtermvalue t,p,p))
+
+	val minDelta = 3
+
+	(* pull all the state * lexv elements from the queue *)
+
+	val stateList = 
+	   let fun f q = let val (elem,newQueue) = Fifo.get q
+			 in elem :: (f newQueue)
+			 end handle Fifo.Empty => nil
+	   in f queue
+	   end
+
+	(* now number elements of stateList, giving distance from
+	   error token *)
+
+	val (_, numStateList) =
+	      List.foldr (fn (a,(num,r)) => (num+1,(a,num)::r)) (0, []) stateList
+
+	(* Represent the set of potential changes as a linked list.
+
+	   Values of datatype Change hold information about a potential change.
+
+	   oper = oper to be applied
+	   pos = the # of the element in stateList that would be altered.
+	   distance = the number of tokens beyond the error token which the
+	     change allows us to parse.
+	   new = new terminal * value pair at that point
+	   orig = original terminal * value pair at the point being changed.
+	 *)
+
+	datatype ('a,'b) change = CHANGE of
+	   {pos : int, distance : int, leftPos: 'b, rightPos: 'b,
+	    new : ('a,'b) lexv list, orig : ('a,'b) lexv list}
+
+
+         val showTerms = String.concat o map (fn TOKEN(t,_) => " " ^ showTerminal t)
+
+	 val printChange = fn c =>
+	  let val CHANGE {distance,new,orig,pos,...} = c
+	  in (TextIO.print ("{distance= " ^ (Int.toString distance));
+	      TextIO.print (",orig ="); TextIO.print(showTerms orig);
+	      TextIO.print (",new ="); TextIO.print(showTerms new);
+	      TextIO.print (",pos= " ^ (Int.toString pos));
+	      TextIO.print "}\n")
+	  end
+
+	val printChangeList = app printChange
+
+(* parse: given a lexPair, a stack, and the distance from the error
+   token, return the distance past the error token that we are able to parse.*)
+
+	fun parse (lexPair,stack,queuePos : int) =
+	    case distanceParse(lexPair,stack,Fifo.empty,queuePos+maxAdvance+1)
+             of (_,_,_,distance,SOME ACCEPT) => 
+		        if maxAdvance-distance-1 >= 0 
+			    then maxAdvance 
+			    else maxAdvance-distance-1
+	      | (_,_,_,distance,_) => maxAdvance - distance - 1
+
+(* catList: concatenate results of scanning list *)
+
+	fun catList l f = List.foldr (fn(a,r)=> f a @ r) [] l
+
+        fun keywordsDelta new = if List.exists (fn(TOKEN(t,_))=>is_keyword t) new
+	               then minDelta else 0
+
+        fun tryChange{lex,stack,pos,leftPos,rightPos,orig,new} =
+	     let val lex' = List.foldr (fn (t',p)=>(t',Stream.cons p)) lex new
+		 val distance = parse(lex',stack,pos+length new-length orig)
+	      in if distance >= minAdvance + keywordsDelta new 
+		   then [CHANGE{pos=pos,leftPos=leftPos,rightPos=rightPos,
+				distance=distance,orig=orig,new=new}] 
+		   else []
+	     end
+
+
+(* tryDelete: Try to delete n terminals.
+              Return single-element [success] or nil.
+	      Do not delete unshiftable terminals. *)
+
+
+    fun tryDelete n ((stack,lexPair as (TOKEN(term,(_,l,r)),_)),qPos) =
+	let fun del(0,accum,left,right,lexPair) =
+	          tryChange{lex=lexPair,stack=stack,
+			    pos=qPos,leftPos=left,rightPos=right,
+			    orig=rev accum, new=[]}
+	      | del(n,accum,left,right,(tok as TOKEN(term,(_,_,r)),lexer)) =
+		   if noShift term then []
+		   else del(n-1,tok::accum,left,r,Stream.get lexer)
+         in del(n,[],l,r,lexPair)
+        end
+
+(* tryInsert: try to insert tokens before the current terminal;
+       return a list of the successes  *)
+
+        fun tryInsert((stack,lexPair as (TOKEN(_,(_,l,_)),_)),queuePos) =
+	       catList terms (fn t =>
+		 tryChange{lex=lexPair,stack=stack,
+			   pos=queuePos,orig=[],new=[tokAt(t,l)],
+			   leftPos=l,rightPos=l})
+			      
+(* trySubst: try to substitute tokens for the current terminal;
+       return a list of the successes  *)
+
+        fun trySubst ((stack,lexPair as (orig as TOKEN (term,(_,l,r)),lexer)),
+		      queuePos) =
+	      if noShift term then []
+	      else
+		  catList terms (fn t =>
+		      tryChange{lex=Stream.get lexer,stack=stack,
+				pos=queuePos,
+				leftPos=l,rightPos=r,orig=[orig],
+				new=[tokAt(t,r)]})
+
+     (* do_delete(toks,lexPair) tries to delete tokens "toks" from "lexPair".
+         If it succeeds, returns SOME(toks',l,r,lp), where
+	     toks' is the actual tokens (with positions and values) deleted,
+	     (l,r) are the (leftmost,rightmost) position of toks', 
+	     lp is what remains of the stream after deletion 
+     *)
+        fun do_delete(nil,lp as (TOKEN(_,(_,l,_)),_)) = SOME(nil,l,l,lp)
+          | do_delete([t],(tok as TOKEN(t',(_,l,r)),lp')) =
+	       if eqT (t, t')
+		   then SOME([tok],l,r,Stream.get lp')
+                   else NONE
+          | do_delete(t::rest,(tok as TOKEN(t',(_,l,r)),lp')) =
+	       if eqT (t,t')
+		   then case do_delete(rest,Stream.get lp')
+                         of SOME(deleted,l',r',lp'') =>
+			       SOME(tok::deleted,l,r',lp'')
+			  | NONE => NONE
+		   else NONE
+			     
+        fun tryPreferred((stack,lexPair),queuePos) =
+	    catList preferred_change (fn (delete,insert) =>
+	       if List.exists noShift delete then [] (* should give warning at
+						 parser-generation time *)
+               else case do_delete(delete,lexPair)
+                     of SOME(deleted,l,r,lp) => 
+			    tryChange{lex=lp,stack=stack,pos=queuePos,
+				      leftPos=l,rightPos=r,orig=deleted,
+				      new=map (fn t=>(tokAt(t,r))) insert}
+		      | NONE => [])
+
+	val changes = catList numStateList tryPreferred @
+	                catList numStateList tryInsert @
+			  catList numStateList trySubst @
+			    catList numStateList (tryDelete 1) @
+			      catList numStateList (tryDelete 2) @
+			        catList numStateList (tryDelete 3)
+
+	val findMaxDist = fn l => 
+	  List.foldr (fn (CHANGE {distance,...},high) => Int.max(distance,high)) 0 l
+
+(* maxDist: max distance past error taken that we could parse *)
+
+	val maxDist = findMaxDist changes
+
+(* remove changes which did not parse maxDist tokens past the error token *)
+
+        val changes = catList changes 
+	      (fn(c as CHANGE{distance,...}) => 
+		  if distance=maxDist then [c] else [])
+
+      in case changes 
+	  of (l as change :: _) =>
+	      let fun print_msg (CHANGE {new,orig,leftPos,rightPos,...}) =
+		  let val s = 
+		      case (orig,new)
+			  of (_::_,[]) => "deleting " ^ (showTerms orig)
+	                   | ([],_::_) => "inserting " ^ (showTerms new)
+			   | _ => "replacing " ^ (showTerms orig) ^
+				 " with " ^ (showTerms new)
+		  in error ("syntax error: " ^ s,leftPos,rightPos)
+		  end
+		   
+		  val _ = 
+		      (if length l > 1 andalso DEBUG2 then
+			   (TextIO.print "multiple fixes possible; could fix it by:\n";
+			    app print_msg l;
+			    TextIO.print "chosen correction:\n")
+		       else ();
+		       print_msg change)
+
+		  (* findNth: find nth queue entry from the error
+		   entry.  Returns the Nth queue entry and the  portion of
+		   the queue from the beginning to the nth-1 entry.  The
+		   error entry is at the end of the queue.
+
+		   Examples:
+
+		   queue = a b c d e
+		   findNth 0 = (e,a b c d)
+		   findNth 1 =  (d,a b c)
+		   *)
+
+		  val findNth = fn n =>
+		      let fun f (h::t,0) = (h,rev t)
+			    | f (h::t,n) = f(t,n-1)
+			    | f (nil,_) = let exception FindNth
+					  in raise FindNth
+					  end
+		      in f (rev stateList,n)
+		      end
+		
+		  val CHANGE {pos,orig,new,...} = change
+		  val (last,queueFront) = findNth pos
+		  val (stack,lexPair) = last
+
+		  val lp1 = List.foldl(fn (_,(_,r)) => Stream.get r) lexPair orig
+		  val lp2 = List.foldr(fn(t,r)=>(t,Stream.cons r)) lp1 new
+
+		  val restQueue = 
+		      Fifo.put((stack,lp2),
+			       List.foldl Fifo.put Fifo.empty queueFront)
+
+		  val (lexPair,stack,queue,_,_) =
+		      distanceParse(lp2,stack,restQueue,pos)
+
+	      in (lexPair,stack,queue)
+	      end
+	| nil => (error("syntax error found at " ^ (showTerminal term),
+			leftPos,leftPos); raise ParseError)
+    end
+
+   val parse = fn {arg,table,lexer,saction,void,lookahead,
+		   ec=ec as {showTerminal,...} : ('_a,'_b) ecRecord} =>
+	let val distance = 15   (* defer distance tokens *)
+	    val minAdvance = 1  (* must parse at least 1 token past error *)
+	    val maxAdvance = Int.max(lookahead,0)(* max distance for parse check *)
+	    val lexPair = Stream.get lexer
+	    val (TOKEN (_,(_,leftPos,_)),_) = lexPair
+	    val startStack = [(initialState table,(void,leftPos,leftPos))]
+	    val startQueue = Fifo.put((startStack,lexPair),Fifo.empty)
+	    val distanceParse = distanceParse(table,showTerminal,saction,arg)
+	    val fixError = mkFixError(ec,distanceParse,minAdvance,maxAdvance)
+	    val ssParse = ssParse(table,showTerminal,saction,fixError,arg)
+	    fun loop (lexPair,stack,queue,_,SOME ACCEPT) =
+		   ssParse(lexPair,stack,queue)
+	      | loop (lexPair,stack,queue,0,_) = ssParse(lexPair,stack,queue)
+	      | loop (lexPair,stack,queue,distance,SOME ERROR) =
+		 let val (lexPair,stack,queue) = fixError(lexPair,stack,queue)
+		 in loop (distanceParse(lexPair,stack,queue,distance))
+		 end
+	      | loop _ = let exception ParseInternal
+			 in raise ParseInternal
+			 end
+	in loop (distanceParse(lexPair,startStack,startQueue,distance))
+	end
+ end;
+
+;
+print_depth 10;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/tptp.lex	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,185 @@
+(*  Title:      HOL/TPTP/TPTP_Parser/tptp.lex
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+ Notes:
+ * Omit %full in definitions to restrict alphabet to ascii.
+ * Could include %posarg to ensure that start counting character positions from
+   0, but it would punish performance.
+ * %s AF F COMMENT; -- could improve by making stateful.
+
+ Acknowledgements:
+ * Geoff Sutcliffe for help with TPTP.
+ * Timothy Bourke for his tips on getting ML-Yacc working with Poly/ML.
+ * An early version of this was ported from the specification shipped with
+   Leo-II, written by Frank Theiss.
+ * Some boilerplate bits were taken from the ml-yacc/ml-lex manual by Roger Price.
+ * Jasmin Blanchette and Makarius Wenzel for help with Isabelle integration.
+*)
+
+structure T = Tokens
+type pos = int             (* Position in file *)
+type lineNo = int
+type svalue = T.svalue
+type ('a,'b) token = ('a,'b) T.token
+type lexresult = (svalue,pos) token
+type lexarg = string
+type arg = lexarg
+val col = ref 0;
+val linep = ref 1;         (* Line pointer *)
+val eolpos = ref 0;
+
+val badCh : string * string * int * int -> unit = fn
+    (file_name, bad, line, col) =>
+    TextIO.output(TextIO.stdOut, file_name ^ "["
+          ^ Int.toString line ^ "." ^ Int.toString col
+          ^ "] Invalid character \"" ^ bad ^ "\"\n");
+
+val eof = fn file_name =>
+  let
+    val result = T.EOF (!linep,!col);
+    val _ = linep := 0;
+  in result end
+(*here could check whether file ended prematurely:
+ see if have open brackets, or if we're in some state other than INITIAL*)
+
+val count_commentlines : string -> unit = fn str =>
+  let
+    val str' = String.explode str
+    val newlines = List.filter (fn x => x = #"\n") str'
+  in linep := (!linep) + (List.length newlines) end
+
+%%
+%header (functor TPTPLexFun(structure Tokens: TPTP_TOKENS));
+%arg (file_name:string);
+
+printable_char            = .;
+viewable_char             = [.\n];
+numeric                   = [0-9];
+lower_alpha               = [a-z];
+upper_alpha               = [A-Z];
+alpha_numeric             = ({lower_alpha}|{upper_alpha}|{numeric}|_);
+zero_numeric              = [0];
+non_zero_numeric          = [1-9];
+slash                     = [/];
+exponent                  = [Ee];
+dot                       = [.];
+any_char                  = [^\n];
+dollar                    = \$;
+ddollar                   = \$\$;
+unsigned_integer          = {numeric}+;
+sign                      = [+-];
+divide                    = [/];
+
+signed_integer            = {sign}{unsigned_integer};
+dot_decimal               = {dot}{numeric}+;
+exp_suffix                = {exponent}({signed_integer}|{unsigned_integer});
+real                      = ({signed_integer}|{unsigned_integer}){dot_decimal}{exp_suffix}?;
+upper_word                = {upper_alpha}{alpha_numeric}*;
+rational                  = ({signed_integer}|{unsigned_integer}){divide}{unsigned_integer};
+
+percentage_sign           = "%";
+
+sq_char                   = ([\040-\041\043-\126]|[\\]['\\]);
+
+ws                        = ([\ ]|[\t]);
+eol                       = ("\013\010"|"\010"|"\013");
+
+single_quote              = ['];
+single_quoted             = {single_quote}({alpha_numeric}|{sq_char}|{ws})+{single_quote};
+
+lower_word                = {lower_alpha}{alpha_numeric}*;
+atomic_system_word        = {ddollar}{lower_word};
+atomic_defined_word       = {dollar}{lower_word};
+
+system_comment_one        = [%][\ ]*{ddollar}[_]*;
+system_comment_multi      = [/][\*][\ ]*(ddollar)([^\*]*[\*][\*]*[^/\*])*[^\*]*[\*][\*]*[/];
+system_comment            = (system_comment_one)|(system_comment_multi);
+comment_one               = {percentage_sign}[^\n]*;
+comment_multi             = [/][\*]([^\*]*[\*]+[^/\*])*[^\*]*[\*]+[/];
+comment                   = ({comment_one}|{comment_multi})+;
+
+do_char                   = ([^"]|[\\]["\\]);
+double_quote              = ["];
+distinct_object           = {double_quote}{do_char}+{double_quote};
+
+%%
+
+{ws}*           => (col:=(!col)+size yytext; continue () );
+
+{eol}           => (linep:=(!linep)+1;
+                   eolpos:=yypos+size yytext; continue ());
+
+"&"             => (col:=yypos-(!eolpos); T.AMPERSAND(!linep,!col));
+
+"@+"            => (col:=yypos-(!eolpos); T.INDEF_CHOICE(!linep,!col));
+"@-"            => (col:=yypos-(!eolpos); T.DEFIN_CHOICE(!linep,!col));
+
+"!!"            => (col:=yypos-(!eolpos); T.OPERATOR_FORALL(!linep,!col));
+"??"            => (col:=yypos-(!eolpos); T.OPERATOR_EXISTS(!linep,!col));
+
+"@"             => (col:=yypos-(!eolpos); T.AT_SIGN(!linep,!col));
+"^"             => (col:=yypos-(!eolpos); T.CARET(!linep,!col));
+
+":"             => (col:=yypos-(!eolpos); T.COLON(!linep,!col));
+","             => (col:=yypos-(!eolpos); T.COMMA(!linep,!col));
+"="             => (col:=yypos-(!eolpos); T.EQUALS(!linep,!col));
+"!"             => (col:=yypos-(!eolpos); T.EXCLAMATION(!linep,!col));
+":="            => (col:=yypos-(!eolpos); T.LET(!linep,!col));
+">"             => (col:=yypos-(!eolpos); T.ARROW(!linep,!col));
+
+"<="            => (col:=yypos-(!eolpos); T.IF(!linep,!col));
+"<=>"           => (col:=yypos-(!eolpos); T.IFF(!linep,!col));
+"=>"            => (col:=yypos-(!eolpos); T.IMPLIES(!linep,!col));
+"["             => (col:=yypos-(!eolpos); T.LBRKT(!linep,!col));
+"("             => (col:=yypos-(!eolpos); T.LPAREN(!linep,!col));
+"->"            => (col:=yypos-(!eolpos); T.MAP_TO(!linep,!col));
+"--"            => (col:=yypos-(!eolpos); T.MMINUS(!linep,!col));
+"~&"            => (col:=yypos-(!eolpos); T.NAND(!linep,!col));
+"!="            => (col:=yypos-(!eolpos); T.NEQUALS(!linep,!col));
+"<~>"           => (col:=yypos-(!eolpos); T.XOR(!linep,!col));
+"~|"            => (col:=yypos-(!eolpos); T.NOR(!linep,!col));
+"."             => (col:=yypos-(!eolpos); T.PERIOD(!linep,!col));
+"++"            => (col:=yypos-(!eolpos); T.PPLUS(!linep,!col));
+"?"             => (col:=yypos-(!eolpos); T.QUESTION(!linep,!col));
+"]"             => (col:=yypos-(!eolpos); T.RBRKT(!linep,!col));
+")"             => (col:=yypos-(!eolpos); T.RPAREN(!linep,!col));
+"~"             => (col:=yypos-(!eolpos); T.TILDE(!linep,!col));
+"|"             => (col:=yypos-(!eolpos); T.VLINE(!linep,!col));
+
+{distinct_object}    => (col:=yypos-(!eolpos); T.DISTINCT_OBJECT(yytext,!linep,!col));
+{rational}           => (col:=yypos-(!eolpos); T.RATIONAL(yytext,!linep,!col));
+{real}               => (col:=yypos-(!eolpos); T.REAL(yytext,!linep,!col));
+{signed_integer}     => (col:=yypos-(!eolpos); T.SIGNED_INTEGER(yytext,!linep,!col));
+{unsigned_integer}   => (col:=yypos-(!eolpos); T.UNSIGNED_INTEGER(yytext,!linep,!col));
+{dot_decimal}        => (col:=yypos-(!eolpos); T.DOT_DECIMAL(yytext,!linep,!col));
+{single_quoted}      => (col:=yypos-(!eolpos); T.SINGLE_QUOTED(yytext,!linep,!col));
+{upper_word}         => (col:=yypos-(!eolpos); T.UPPER_WORD(yytext,!linep,!col));
+{comment}            => (col:=yypos-(!eolpos); count_commentlines yytext;T.COMMENT(yytext,!linep,!col));
+
+"thf"          => (col:=yypos-(!eolpos); T.THF(!linep,!col));
+"fof"          => (col:=yypos-(!eolpos); T.FOF(!linep,!col));
+"cnf"          => (col:=yypos-(!eolpos); T.CNF(!linep,!col));
+"tff"          => (col:=yypos-(!eolpos); T.TFF(!linep,!col));
+"include"      => (col:=yypos-(!eolpos); T.INCLUDE(!linep,!col));
+
+"$thf"          => (col:=yypos-(!eolpos); T.DTHF(!linep,!col));
+"$fof"          => (col:=yypos-(!eolpos); T.DFOF(!linep,!col));
+"$cnf"          => (col:=yypos-(!eolpos); T.DCNF(!linep,!col));
+"$fot"          => (col:=yypos-(!eolpos); T.DFOT(!linep,!col));
+"$tff"          => (col:=yypos-(!eolpos); T.DTFF(!linep,!col));
+
+"$ite_f"        => (col:=yypos-(!eolpos); T.ITE_F(!linep,!col));
+"$ite_t"        => (col:=yypos-(!eolpos); T.ITE_T(!linep,!col));
+
+{lower_word}          => (col:=yypos-(!eolpos); T.LOWER_WORD(yytext,!linep,!col));
+{atomic_system_word}  => (col:=yypos-(!eolpos); T.ATOMIC_SYSTEM_WORD(yytext,!linep,!col));
+{atomic_defined_word} => (col:=yypos-(!eolpos); T.ATOMIC_DEFINED_WORD(yytext,!linep,!col));
+
+"+"           => (col:=yypos-(!eolpos); T.PLUS(!linep,!col));
+"*"           => (col:=yypos-(!eolpos); T.TIMES(!linep,!col));
+"-->"         => (col:=yypos-(!eolpos); T.GENTZEN_ARROW(!linep,!col));
+"<<"          => (col:=yypos-(!eolpos); T.SUBTYPE(!linep,!col));
+"!>"          => (col:=yypos-(!eolpos); T.DEP_PROD(!linep,!col));
+"?*"          => (col:=yypos-(!eolpos); T.DEP_SUM(!linep,!col));
+
+":-"          => (col:=yypos-(!eolpos); T.LET_TERM(!linep,!col));
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/tptp.yacc	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,653 @@
+open TPTP_Syntax
+
+exception UNRECOGNISED_SYMBOL of string * string
+
+exception UNRECOGNISED_ROLE of string
+fun classify_role role =
+  case role of
+    "axiom" => Role_Axiom
+  | "hypothesis" => Role_Hypothesis
+  | "definition" => Role_Definition
+  | "assumption" => Role_Assumption
+  | "lemma" => Role_Lemma
+  | "theorem" => Role_Theorem
+  | "conjecture" => Role_Conjecture
+  | "negated_conjecture" => Role_Negated_Conjecture
+  | "plain" => Role_Plain
+  | "fi_domain" => Role_Fi_Domain
+  | "fi_functors" => Role_Fi_Functors
+  | "fi_predicates" => Role_Fi_Predicates
+  | "type" => Role_Type
+  | "unknown" => Role_Unknown
+  | thing => raise (UNRECOGNISED_ROLE thing)
+
+%%
+%name TPTP
+%term AMPERSAND | AT_SIGN | CARET | COLON | COMMA | EQUALS | EXCLAMATION
+  | LET | ARROW | IF | IFF | IMPLIES | INCLUDE
+  | LAMBDA | LBRKT | LPAREN | MAP_TO | MMINUS | NAND
+  | NEQUALS | XOR | NOR | PERIOD | PPLUS | QUESTION | RBRKT | RPAREN
+  | TILDE | TOK_FALSE | TOK_I | TOK_O | TOK_INT | TOK_REAL | TOK_RAT | TOK_TRUE
+  | TOK_TYPE | VLINE | EOF | DTHF | DFOF | DCNF | DFOT | DTFF | REAL of string
+  | RATIONAL of string | SIGNED_INTEGER of string | UNSIGNED_INTEGER of string
+  | DOT_DECIMAL of string | SINGLE_QUOTED of string | UPPER_WORD of string
+  | LOWER_WORD of string | COMMENT of string
+  | DISTINCT_OBJECT of string
+  | DUD | INDEF_CHOICE | DEFIN_CHOICE
+  | OPERATOR_FORALL | OPERATOR_EXISTS
+  | PLUS | TIMES | GENTZEN_ARROW | DEP_SUM | DEP_PROD
+  | ATOMIC_DEFINED_WORD of string | ATOMIC_SYSTEM_WORD of string
+  | SUBTYPE | LET_TERM
+  | THF | TFF | FOF | CNF
+  | ITE_F | ITE_T
+%nonterm
+    annotations of annotation option
+  | name of string
+  | name_list of string list
+  | formula_selection of string list
+  | optional_info of general_term list
+  | general_list of general_list | general_terms of general_term list
+  | general_term of general_term
+
+  | atomic_word of string
+  | general_data of general_data | variable_ of string
+  | number of number_kind * string | formula_data of general_data
+  | integer of string
+  | identifier of string
+  | general_function of general_data | useful_info of general_list
+  | file_name of string
+  | functor_ of symbol
+  | term of tptp_term
+  | arguments of tptp_term list
+  | defined_functor of symbol
+  | system_functor of symbol
+  | system_constant of symbol
+  | system_term of symbol * tptp_term list
+  | defined_constant of symbol
+  | defined_plain_term of symbol * tptp_term list
+  | defined_atomic_term of tptp_term
+  | defined_atom of tptp_term
+  | defined_term of tptp_term
+  | constant of symbol
+  | plain_term of symbol * tptp_term list
+  | function_term of tptp_term
+  | conditional_term of tptp_term
+  | system_atomic_formula of tptp_formula
+  | infix_equality of symbol
+  | infix_inequality of symbol
+  | defined_infix_pred of symbol
+  | defined_infix_formula of tptp_formula
+  | defined_prop of string
+  | defined_pred of string
+  | defined_plain_formula of tptp_formula
+  | defined_atomic_formula of tptp_formula
+  | plain_atomic_formula of tptp_formula
+  | atomic_formula of tptp_formula
+  | unary_connective of symbol
+
+  | defined_type of tptp_base_type
+  | system_type of string
+  | assoc_connective of symbol
+  | binary_connective of symbol
+  | fol_quantifier of quantifier
+  | thf_unary_connective of symbol
+  | thf_pair_connective of symbol
+  | thf_quantifier of quantifier
+  | fol_infix_unary of tptp_formula
+  | thf_conn_term of symbol
+  | literal of tptp_formula
+  | disjunction of tptp_formula
+  | cnf_formula of tptp_formula
+  | fof_tuple_list of tptp_formula list
+  | fof_tuple of tptp_formula list
+  | fof_sequent of tptp_formula
+  | fof_unary_formula of tptp_formula
+  | fof_variable_list of string list
+  | fof_quantified_formula of tptp_formula
+  | fof_unitary_formula of tptp_formula
+  | fof_and_formula of tptp_formula
+  | fof_or_formula of tptp_formula
+  | fof_binary_assoc of tptp_formula
+  | fof_binary_nonassoc of tptp_formula
+  | fof_binary_formula of tptp_formula
+  | fof_logic_formula of tptp_formula
+  | fof_formula of tptp_formula
+  | tff_tuple of tptp_formula list
+  | tff_tuple_list of tptp_formula list
+  | tff_sequent of tptp_formula
+  | tff_conditional of tptp_formula
+  | tff_defined_var of tptp_let
+  | tff_let_list of tptp_let list
+  | tptp_let of tptp_formula
+  | tff_xprod_type of tptp_type
+  | tff_mapping_type of tptp_type
+  | tff_atomic_type of tptp_type
+  | tff_unitary_type of tptp_type
+  | tff_top_level_type of tptp_type
+  | tff_untyped_atom of symbol * tptp_type option
+  | tff_typed_atom of symbol * tptp_type option
+  | tff_unary_formula of tptp_formula
+  | tff_typed_variable of string * tptp_type option
+  | tff_variable of string * tptp_type option
+  | tff_variable_list of (string * tptp_type option) list
+  | tff_quantified_formula of tptp_formula
+  | tff_unitary_formula of tptp_formula
+  | tff_and_formula of tptp_formula
+  | tff_or_formula of tptp_formula
+  | tff_binary_assoc of tptp_formula
+  | tff_binary_nonassoc of tptp_formula
+  | tff_binary_formula of tptp_formula
+  | tff_logic_formula of tptp_formula
+  | tff_formula of tptp_formula
+
+  | thf_tuple of tptp_formula list
+  | thf_tuple_list of tptp_formula list
+  | thf_sequent of tptp_formula
+  | thf_conditional of tptp_formula
+  | thf_defined_var of tptp_let
+  | thf_let_list of tptp_let list
+  | thf_let of tptp_formula
+  | thf_atom of tptp_formula
+  | thf_union_type of tptp_type
+  | thf_xprod_type of tptp_type
+  | thf_mapping_type of tptp_type
+  | thf_binary_type of tptp_type
+  | thf_unitary_type of tptp_type
+  | thf_top_level_type of tptp_type
+  | thf_subtype of tptp_type
+  | thf_typeable_formula of tptp_formula
+  | thf_type_formula of tptp_formula * tptp_type
+  | thf_unary_formula of tptp_formula
+  | thf_typed_variable of string * tptp_type option
+  | thf_variable of string * tptp_type option
+  | thf_variable_list of (string * tptp_type option) list
+  | thf_quantified_formula of tptp_formula
+  | thf_unitary_formula of tptp_formula
+  | thf_apply_formula of tptp_formula
+  | thf_and_formula of tptp_formula
+  | thf_or_formula of tptp_formula
+  | thf_binary_tuple of tptp_formula
+  | thf_binary_pair of tptp_formula
+  | thf_binary_formula of tptp_formula
+  | thf_logic_formula of tptp_formula
+  | thf_formula of tptp_formula
+  | formula_role of role
+
+  | cnf_annotated of tptp_line
+  | fof_annotated of tptp_line
+  | tff_annotated of tptp_line
+  | thf_annotated of tptp_line
+  | annotated_formula of tptp_line
+  | include_ of tptp_line
+  | tptp_input of tptp_line
+  | tptp_file of tptp_problem
+  | tptp of tptp_problem
+
+%pos int
+%eop EOF
+%noshift EOF
+%arg (file_name) : string
+
+%nonassoc LET
+%nonassoc RPAREN
+%nonassoc DUD
+%right COMMA
+%left COLON
+
+%left AT_SIGN
+%nonassoc IFF XOR
+%right IMPLIES IF
+%nonassoc EQUALS NEQUALS
+%right VLINE NOR
+%left AMPERSAND NAND
+%right ARROW
+%left PLUS
+%left TIMES
+
+%right OPERATOR_FORALL OPERATOR_EXISTS
+
+%nonassoc EXCLAMATION QUESTION LAMBDA CARET
+%nonassoc TILDE
+%pure
+%start tptp
+%verbose
+%%
+
+(*  Title:      HOL/TPTP/TPTP_Parser/tptp.yacc
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+ Parser for TPTP languages. Latest version of the language spec can
+ be obtained from http://www.cs.miami.edu/~tptp/TPTP/SyntaxBNF.html
+*)
+
+annotations : COMMA general_term optional_info (( SOME (general_term, optional_info) ))
+            |                                  (( NONE ))
+
+optional_info : COMMA useful_info (( useful_info ))
+              |                   (( [] ))
+
+useful_info : general_list (( general_list ))
+
+general_list : LBRKT general_terms RBRKT (( general_terms ))
+             | LBRKT RBRKT               (( [] ))
+
+general_terms : general_term COMMA general_terms (( general_term :: general_terms ))
+              | general_term                     (( [general_term] ))
+
+general_term : general_data                    (( General_Data general_data ))
+             | general_data COLON general_term (( General_Term (general_data, general_term) ))
+             | general_list                    (( General_List general_list ))
+
+atomic_word : LOWER_WORD    (( LOWER_WORD ))
+            | SINGLE_QUOTED (( SINGLE_QUOTED ))
+            | THF           (( "thf" ))
+            | TFF           (( "tff" ))
+            | FOF           (( "fof" ))
+            | CNF           (( "cnf" ))
+            | INCLUDE       (( "include" ))
+
+variable_ : UPPER_WORD  (( UPPER_WORD ))
+
+general_function: atomic_word LPAREN general_terms RPAREN (( Application (atomic_word, general_terms) ))
+
+general_data : atomic_word       (( Atomic_Word atomic_word ))
+             | general_function  (( general_function ))
+             | variable_         (( V variable_ ))
+             | number            (( Number number ))
+             | DISTINCT_OBJECT   (( Distinct_Object DISTINCT_OBJECT ))
+             | formula_data      (( formula_data ))
+
+number : integer          (( (Int_num, integer) ))
+       | REAL             (( (Real_num, REAL) ))
+       | RATIONAL         (( (Rat_num, RATIONAL) ))
+
+integer: UNSIGNED_INTEGER (( UNSIGNED_INTEGER ))
+       | SIGNED_INTEGER   (( SIGNED_INTEGER ))
+
+file_name : SINGLE_QUOTED (( SINGLE_QUOTED ))
+
+formula_data : DTHF LPAREN thf_formula RPAREN (( Formula_Data (THF, thf_formula) ))
+             | DTFF LPAREN tff_formula RPAREN (( Formula_Data (TFF, tff_formula) ))
+             | DFOF LPAREN fof_formula RPAREN (( Formula_Data (FOF, fof_formula) ))
+             | DCNF LPAREN cnf_formula RPAREN (( Formula_Data (CNF, cnf_formula) ))
+             | DFOT LPAREN term RPAREN (( Term_Data term ))
+
+system_type : ATOMIC_SYSTEM_WORD (( ATOMIC_SYSTEM_WORD ))
+
+defined_type : ATOMIC_DEFINED_WORD ((
+  case ATOMIC_DEFINED_WORD of
+    "$i" => Type_Ind
+  | "$o" => Type_Bool
+  | "$iType" => Type_Ind
+  | "$oType" => Type_Bool
+  | "$int" => Type_Int
+  | "$real" => Type_Real
+  | "$rat" => Type_Rat
+  | "$tType" => Type_Type
+  | thing => raise UNRECOGNISED_SYMBOL ("defined_type", thing)
+))
+
+functor_ : atomic_word (( Uninterpreted atomic_word ))
+
+arguments : term                  (( [term] ))
+          | term COMMA arguments  (( term :: arguments ))
+
+system_functor : ATOMIC_SYSTEM_WORD (( System ATOMIC_SYSTEM_WORD ))
+system_constant : system_functor (( system_functor ))
+system_term : system_constant                         (( (system_constant, []) ))
+            | system_functor LPAREN arguments RPAREN  (( (system_functor, arguments) ))
+
+defined_functor : ATOMIC_DEFINED_WORD ((
+  case ATOMIC_DEFINED_WORD of
+    "$sum" => Interpreted_ExtraLogic Sum
+  | "$difference" => Interpreted_ExtraLogic Difference
+  | "$product" => Interpreted_ExtraLogic Product
+  | "$quotient" => Interpreted_ExtraLogic Quotient
+  | "$quotient_e" => Interpreted_ExtraLogic Quotient_E
+  | "$quotient_t" => Interpreted_ExtraLogic Quotient_T
+  | "$quotient_f" => Interpreted_ExtraLogic Quotient_F
+  | "$remainder_e" => Interpreted_ExtraLogic Remainder_E
+  | "$remainder_t" => Interpreted_ExtraLogic Remainder_T
+  | "$remainder_f" => Interpreted_ExtraLogic Remainder_F
+  | "$floor" => Interpreted_ExtraLogic Floor
+  | "$ceiling" => Interpreted_ExtraLogic Ceiling
+  | "$truncate" => Interpreted_ExtraLogic Truncate
+  | "$round" => Interpreted_ExtraLogic Round
+  | "$to_int" => Interpreted_ExtraLogic To_Int
+  | "$to_rat" => Interpreted_ExtraLogic To_Rat
+  | "$to_real" => Interpreted_ExtraLogic To_Real
+  | "$uminus" => Interpreted_ExtraLogic UMinus
+
+  | "$i" => TypeSymbol Type_Ind
+  | "$o" => TypeSymbol Type_Bool
+  | "$iType" => TypeSymbol Type_Ind
+  | "$oType" => TypeSymbol Type_Bool
+  | "$int" => TypeSymbol Type_Int
+  | "$real" => TypeSymbol Type_Real
+  | "$rat" => TypeSymbol Type_Rat
+  | "$tType" => TypeSymbol Type_Type
+
+  | "$true" => Interpreted_Logic True
+  | "$false" => Interpreted_Logic False
+
+  | "$less" => Interpreted_ExtraLogic Less
+  | "$lesseq" => Interpreted_ExtraLogic LessEq
+  | "$greatereq" => Interpreted_ExtraLogic GreaterEq
+  | "$greater" => Interpreted_ExtraLogic Greater
+  | "$evaleq" => Interpreted_ExtraLogic EvalEq
+
+  | "$is_int" => Interpreted_ExtraLogic Is_Int
+  | "$is_rat" => Interpreted_ExtraLogic Is_Rat
+
+  | thing => raise UNRECOGNISED_SYMBOL ("defined_functor", thing)
+))
+
+defined_constant : defined_functor (( defined_functor ))
+
+defined_plain_term : defined_constant                        (( (defined_constant, []) ))
+                   | defined_functor LPAREN arguments RPAREN (( (defined_functor, arguments) ))
+defined_atomic_term : defined_plain_term (( Term_Func defined_plain_term ))
+defined_atom : number          (( Term_Num number ))
+             | DISTINCT_OBJECT (( Term_Distinct_Object DISTINCT_OBJECT ))
+defined_term : defined_atom        (( defined_atom ))
+             | defined_atomic_term (( defined_atomic_term ))
+constant : functor_ (( functor_ ))
+plain_term : constant                          (( (constant, []) ))
+           | functor_ LPAREN arguments RPAREN  (( (functor_, arguments) ))
+function_term : plain_term    (( Term_Func plain_term ))
+              | defined_term  (( defined_term ))
+              | system_term   (( Term_Func system_term ))
+
+conditional_term : ITE_T LPAREN tff_logic_formula COMMA term COMMA term RPAREN ((
+  Term_Conditional (tff_logic_formula, term1, term2)
+))
+
+term : function_term     (( function_term ))
+     | variable_         (( Term_Var variable_ ))
+     | conditional_term  (( conditional_term ))
+
+system_atomic_formula : system_term  (( Pred system_term ))
+infix_equality : EQUALS    (( Interpreted_Logic Equals ))
+infix_inequality : NEQUALS (( Interpreted_Logic NEquals ))
+defined_infix_pred : infix_equality  (( infix_equality ))
+defined_infix_formula : term defined_infix_pred term ((Pred (defined_infix_pred, [term1, term2])))
+defined_prop : ATOMIC_DEFINED_WORD ((
+  case ATOMIC_DEFINED_WORD of
+    "$true"  => "$true"
+  | "$false" => "$false"
+  | thing => raise UNRECOGNISED_SYMBOL ("defined_prop", thing)
+))
+defined_pred : ATOMIC_DEFINED_WORD ((
+  case ATOMIC_DEFINED_WORD of
+    "$distinct"  => "$distinct"
+  | "$ite_f" => "$ite_f"
+  | "$less" => "$less"
+  | "$lesseq" => "$lesseq"
+  | "$greater" => "$greater"
+  | "$greatereq" => "$greatereq"
+  | "$is_int" => "$is_int"
+  | "$is_rat" => "$is_rat"
+  | thing => raise UNRECOGNISED_SYMBOL ("defined_pred", thing)
+))
+defined_plain_formula : defined_plain_term (( Pred defined_plain_term ))
+defined_atomic_formula : defined_plain_formula (( defined_plain_formula ))
+                       | defined_infix_formula (( defined_infix_formula ))
+plain_atomic_formula : plain_term (( Pred plain_term ))
+atomic_formula : plain_atomic_formula   (( plain_atomic_formula ))
+               | defined_atomic_formula (( defined_atomic_formula ))
+               | system_atomic_formula  (( system_atomic_formula ))
+
+assoc_connective : VLINE      (( Interpreted_Logic Or ))
+                 | AMPERSAND  (( Interpreted_Logic And ))
+binary_connective : IFF       (( Interpreted_Logic Iff ))
+                  | IMPLIES   (( Interpreted_Logic If ))
+                  | IF        (( Interpreted_Logic Fi ))
+                  | XOR       (( Interpreted_Logic Xor ))
+                  | NOR       (( Interpreted_Logic Nor ))
+                  | NAND      (( Interpreted_Logic Nand ))
+
+fol_quantifier : EXCLAMATION  (( Forall ))
+               | QUESTION     (( Exists ))
+thf_unary_connective : unary_connective (( unary_connective ))
+                     | OPERATOR_FORALL  (( Interpreted_Logic Op_Forall ))
+                     | OPERATOR_EXISTS  (( Interpreted_Logic Op_Exists ))
+thf_pair_connective : infix_equality    (( infix_equality ))
+                    | infix_inequality  (( infix_inequality ))
+                    | binary_connective (( binary_connective ))
+thf_quantifier : fol_quantifier (( fol_quantifier ))
+               | CARET          (( Lambda ))
+               | DEP_PROD       (( Dep_Prod ))
+               | DEP_SUM        (( Dep_Sum ))
+               | INDEF_CHOICE   (( Epsilon ))
+               | DEFIN_CHOICE   (( Iota ))
+fol_infix_unary : term infix_inequality term (( Pred (infix_inequality, [term1, term2]) ))
+thf_conn_term : thf_pair_connective   (( thf_pair_connective ))
+              | assoc_connective      (( assoc_connective ))
+              | thf_unary_connective  (( thf_unary_connective ))
+literal : atomic_formula        (( atomic_formula ))
+        | TILDE atomic_formula  (( Fmla (Interpreted_Logic Not, [atomic_formula]) ))
+        | fol_infix_unary       (( fol_infix_unary ))
+disjunction : literal                    (( literal ))
+            | disjunction VLINE literal  (( Fmla (Interpreted_Logic Or, [disjunction, literal]) ))
+cnf_formula : LPAREN disjunction RPAREN  (( disjunction ))
+            | disjunction                (( disjunction ))
+fof_tuple_list : fof_logic_formula                      (( [fof_logic_formula] ))
+               | fof_logic_formula COMMA fof_tuple_list (( fof_logic_formula :: fof_tuple_list ))
+fof_tuple : LBRKT RBRKT                 (( [] ))
+          | LBRKT fof_tuple_list RBRKT  (( fof_tuple_list ))
+fof_sequent : fof_tuple GENTZEN_ARROW fof_tuple (( Sequent (fof_tuple1, fof_tuple2) ))
+            | LPAREN fof_sequent RPAREN         (( fof_sequent ))
+unary_connective : TILDE (( Interpreted_Logic Not ))
+fof_unary_formula : unary_connective fof_unitary_formula (( Fmla (unary_connective, [fof_unitary_formula]) ))
+                  | fol_infix_unary                      (( fol_infix_unary ))
+fof_variable_list : variable_                          (( [variable_] ))
+                  | variable_ COMMA fof_variable_list  (( variable_ :: fof_variable_list ))
+fof_quantified_formula : fol_quantifier LBRKT fof_variable_list RBRKT COLON fof_unitary_formula ((
+  Quant (fol_quantifier, map (fn v => (v, NONE)) fof_variable_list, fof_unitary_formula)
+))
+fof_unitary_formula : fof_quantified_formula (( fof_quantified_formula ))
+                    | fof_unary_formula      (( fof_unary_formula ))
+                    | atomic_formula         (( atomic_formula ))
+                    | LPAREN fof_logic_formula RPAREN (( fof_logic_formula ))
+fof_and_formula : fof_unitary_formula AMPERSAND fof_unitary_formula (( Fmla (Interpreted_Logic And, [fof_unitary_formula1, fof_unitary_formula2]) ))
+                | fof_and_formula AMPERSAND fof_unitary_formula     (( Fmla (Interpreted_Logic And, [fof_and_formula, fof_unitary_formula]) ))
+fof_or_formula : fof_unitary_formula VLINE fof_unitary_formula  (( Fmla (Interpreted_Logic Or, [fof_unitary_formula1, fof_unitary_formula2]) ))
+               | fof_or_formula VLINE fof_unitary_formula       (( Fmla (Interpreted_Logic Or, [fof_or_formula, fof_unitary_formula]) ))
+fof_binary_assoc : fof_or_formula  (( fof_or_formula ))
+                 | fof_and_formula (( fof_and_formula ))
+fof_binary_nonassoc : fof_unitary_formula binary_connective fof_unitary_formula ((
+  Fmla (binary_connective, [fof_unitary_formula1, fof_unitary_formula2] )
+))
+fof_binary_formula : fof_binary_nonassoc (( fof_binary_nonassoc ))
+                   | fof_binary_assoc    (( fof_binary_assoc ))
+fof_logic_formula : fof_binary_formula   (( fof_binary_formula ))
+                  | fof_unitary_formula  (( fof_unitary_formula ))
+fof_formula : fof_logic_formula  (( fof_logic_formula ))
+            | fof_sequent        (( fof_sequent ))
+
+
+tff_tuple : LBRKT RBRKT    (( [] ))
+          | LBRKT tff_tuple_list RBRKT (( tff_tuple_list ))
+tff_tuple_list : tff_logic_formula COMMA tff_tuple_list (( tff_logic_formula :: tff_tuple_list ))
+               | tff_logic_formula                      (( [tff_logic_formula] ))
+tff_sequent : tff_tuple GENTZEN_ARROW tff_tuple (( Sequent (tff_tuple1, tff_tuple2) ))
+            | LPAREN tff_sequent RPAREN         (( tff_sequent ))
+tff_conditional : ITE_F LPAREN tff_logic_formula COMMA tff_logic_formula COMMA tff_logic_formula RPAREN ((
+  Conditional (tff_logic_formula1, tff_logic_formula2, tff_logic_formula3)
+))
+tff_defined_var : variable_ LET tff_logic_formula (( Let_fmla ((variable_, NONE), tff_logic_formula) ))
+                | variable_ LET_TERM term (( Let_term ((variable_, NONE), term) ))
+                | LPAREN tff_defined_var RPAREN (( tff_defined_var ))
+tff_let_list : tff_defined_var                    (( [tff_defined_var] ))
+             | tff_defined_var COMMA tff_let_list (( tff_defined_var :: tff_let_list ))
+tptp_let : LET LBRKT tff_let_list RBRKT COLON tff_unitary_formula ((
+  Let (tff_let_list, tff_unitary_formula)
+))
+tff_xprod_type : tff_atomic_type TIMES tff_atomic_type (( Prod_type(tff_atomic_type1, tff_atomic_type2) ))
+               | tff_xprod_type TIMES tff_atomic_type  (( Prod_type(tff_xprod_type, tff_atomic_type) ))
+               | LPAREN tff_xprod_type RPAREN          (( tff_xprod_type ))
+tff_mapping_type : tff_unitary_type ARROW tff_atomic_type (( Fn_type(tff_unitary_type, tff_atomic_type) ))
+                 | LPAREN tff_mapping_type RPAREN         (( tff_mapping_type ))
+tff_atomic_type : atomic_word   (( Atom_type atomic_word ))
+                | defined_type  (( Defined_type defined_type ))
+tff_unitary_type : tff_atomic_type               (( tff_atomic_type ))
+                 | LPAREN tff_xprod_type RPAREN  (( tff_xprod_type ))
+tff_top_level_type : tff_atomic_type   (( tff_atomic_type ))
+                   | tff_mapping_type  (( tff_mapping_type ))
+tff_untyped_atom : functor_       (( (functor_, NONE) ))
+                 | system_functor (( (system_functor, NONE) ))
+tff_typed_atom : tff_untyped_atom COLON tff_top_level_type (( (fst tff_untyped_atom, SOME tff_top_level_type) ))
+               | LPAREN tff_typed_atom RPAREN              (( tff_typed_atom ))
+
+tff_unary_formula : unary_connective tff_unitary_formula (( Fmla (unary_connective, [tff_unitary_formula]) ))
+                  | fol_infix_unary                      (( fol_infix_unary ))
+tff_typed_variable : variable_ COLON tff_atomic_type (( (variable_, SOME tff_atomic_type) ))
+tff_variable : tff_typed_variable (( tff_typed_variable ))
+             | variable_          (( (variable_, NONE) ))
+tff_variable_list : tff_variable                         (( [tff_variable] ))
+                  | tff_variable COMMA tff_variable_list (( tff_variable :: tff_variable_list ))
+tff_quantified_formula : fol_quantifier LBRKT tff_variable_list RBRKT COLON tff_unitary_formula ((
+  Quant (fol_quantifier, tff_variable_list, tff_unitary_formula)
+))
+tff_unitary_formula : tff_quantified_formula           (( tff_quantified_formula ))
+                    | tff_unary_formula                (( tff_unary_formula ))
+                    | atomic_formula                   (( atomic_formula ))
+                    | tptp_let                         (( tptp_let ))
+                    | variable_                        (( Pred (Uninterpreted variable_, []) ))
+                    | tff_conditional                  (( tff_conditional ))
+                    | LPAREN tff_logic_formula RPAREN  (( tff_logic_formula ))
+tff_and_formula : tff_unitary_formula AMPERSAND tff_unitary_formula (( Fmla (Interpreted_Logic And, [tff_unitary_formula1, tff_unitary_formula2]) ))
+                | tff_and_formula AMPERSAND tff_unitary_formula     (( Fmla (Interpreted_Logic And, [tff_and_formula, tff_unitary_formula]) ))
+tff_or_formula : tff_unitary_formula VLINE tff_unitary_formula      (( Fmla (Interpreted_Logic Or, [tff_unitary_formula1, tff_unitary_formula2]) ))
+               | tff_or_formula VLINE tff_unitary_formula           (( Fmla (Interpreted_Logic Or, [tff_or_formula, tff_unitary_formula]) ))
+tff_binary_assoc : tff_or_formula  (( tff_or_formula ))
+                 | tff_and_formula (( tff_and_formula ))
+tff_binary_nonassoc : tff_unitary_formula binary_connective tff_unitary_formula (( Fmla (binary_connective, [tff_unitary_formula1, tff_unitary_formula2]) ))
+tff_binary_formula : tff_binary_nonassoc (( tff_binary_nonassoc ))
+                   | tff_binary_assoc    (( tff_binary_assoc ))
+tff_logic_formula : tff_binary_formula   (( tff_binary_formula ))
+                  | tff_unitary_formula  (( tff_unitary_formula ))
+tff_formula : tff_logic_formula  (( tff_logic_formula ))
+            | tff_typed_atom     (( Atom (TFF_Typed_Atom tff_typed_atom) ))
+            | tff_sequent        (( tff_sequent ))
+
+thf_tuple : LBRKT RBRKT                (( [] ))
+          | LBRKT thf_tuple_list RBRKT (( thf_tuple_list ))
+thf_tuple_list : thf_logic_formula                      (( [thf_logic_formula] ))
+               | thf_logic_formula COMMA thf_tuple_list (( thf_logic_formula :: thf_tuple_list ))
+thf_sequent : thf_tuple GENTZEN_ARROW thf_tuple  (( Sequent(thf_tuple1, thf_tuple2) ))
+            | LPAREN thf_sequent RPAREN          (( thf_sequent ))
+thf_conditional : ITE_F LPAREN thf_logic_formula COMMA thf_logic_formula COMMA thf_logic_formula RPAREN ((
+  Conditional (thf_logic_formula1, thf_logic_formula2, thf_logic_formula3)
+))
+thf_defined_var : thf_variable LET thf_logic_formula (( Let_fmla (thf_variable, thf_logic_formula) ))
+                | LPAREN thf_defined_var RPAREN      (( thf_defined_var ))
+thf_let_list : thf_defined_var                    (( [thf_defined_var] ))
+             | thf_defined_var COMMA thf_let_list (( thf_defined_var :: thf_let_list ))
+thf_let : LET LBRKT thf_let_list RBRKT COLON thf_unitary_formula ((
+  Let (thf_let_list, thf_unitary_formula)
+))
+thf_atom : term          (( Atom (THF_Atom_term term) ))
+         | thf_conn_term (( Atom (THF_Atom_conn_term thf_conn_term) ))
+thf_union_type : thf_unitary_type PLUS thf_unitary_type    (( Sum_type(thf_unitary_type1, thf_unitary_type2) ))
+               | thf_union_type PLUS thf_unitary_type      (( Sum_type(thf_union_type, thf_unitary_type) ))
+thf_xprod_type : thf_unitary_type TIMES thf_unitary_type   (( Prod_type(thf_unitary_type1, thf_unitary_type2) ))
+               | thf_xprod_type TIMES thf_unitary_type     (( Prod_type(thf_xprod_type, thf_unitary_type) ))
+thf_mapping_type : thf_unitary_type ARROW thf_unitary_type (( Fn_type(thf_unitary_type1, thf_unitary_type2) ))
+                 | thf_unitary_type ARROW thf_mapping_type (( Fn_type(thf_unitary_type, thf_mapping_type) ))
+thf_binary_type : thf_mapping_type (( thf_mapping_type ))
+                | thf_xprod_type   (( thf_xprod_type ))
+                | thf_union_type   (( thf_union_type ))
+thf_unitary_type : thf_unitary_formula (( Fmla_type thf_unitary_formula ))
+thf_top_level_type : thf_logic_formula (( Fmla_type thf_logic_formula ))
+thf_subtype : constant SUBTYPE constant (( Subtype(constant1, constant2) ))
+thf_typeable_formula : thf_atom                         (( thf_atom ))
+                     | LPAREN thf_logic_formula RPAREN  (( thf_logic_formula ))
+thf_type_formula : thf_typeable_formula COLON thf_top_level_type (( (thf_typeable_formula, thf_top_level_type) ))
+thf_unary_formula : thf_unary_connective LPAREN thf_logic_formula RPAREN ((
+  Fmla (thf_unary_connective, [thf_logic_formula])
+))
+thf_typed_variable : variable_ COLON thf_top_level_type (( (variable_, SOME thf_top_level_type) ))
+thf_variable : thf_typed_variable (( thf_typed_variable ))
+             | variable_          (( (variable_, NONE) ))
+thf_variable_list : thf_variable                          (( [thf_variable] ))
+                  | thf_variable COMMA thf_variable_list  (( thf_variable :: thf_variable_list ))
+thf_quantified_formula : thf_quantifier LBRKT thf_variable_list RBRKT COLON thf_unitary_formula ((
+  Quant (thf_quantifier, thf_variable_list, thf_unitary_formula)
+))
+thf_unitary_formula : thf_quantified_formula (( thf_quantified_formula ))
+                    | thf_unary_formula      (( thf_unary_formula ))
+                    | thf_atom               (( thf_atom ))
+                    | thf_let                (( thf_let ))
+                    | thf_conditional        (( thf_conditional ))
+                    | LPAREN thf_logic_formula RPAREN  (( thf_logic_formula ))
+thf_apply_formula : thf_unitary_formula AT_SIGN thf_unitary_formula (( Fmla (Interpreted_ExtraLogic Apply, [thf_unitary_formula1, thf_unitary_formula2]) ))
+                  | thf_apply_formula AT_SIGN thf_unitary_formula   (( Fmla (Interpreted_ExtraLogic Apply, [thf_apply_formula, thf_unitary_formula]) ))
+thf_and_formula : thf_unitary_formula AMPERSAND thf_unitary_formula (( Fmla (Interpreted_Logic And, [thf_unitary_formula1, thf_unitary_formula2]) ))
+                | thf_and_formula AMPERSAND thf_unitary_formula     (( Fmla (Interpreted_Logic And, [thf_and_formula, thf_unitary_formula]) ))
+thf_or_formula : thf_unitary_formula VLINE thf_unitary_formula (( Fmla (Interpreted_Logic Or, [thf_unitary_formula1, thf_unitary_formula2]) ))
+               | thf_or_formula VLINE thf_unitary_formula      (( Fmla (Interpreted_Logic Or, [thf_or_formula, thf_unitary_formula]) ))
+thf_binary_tuple : thf_or_formula    (( thf_or_formula ))
+                 | thf_and_formula   (( thf_and_formula ))
+                 | thf_apply_formula (( thf_apply_formula ))
+thf_binary_pair : thf_unitary_formula thf_pair_connective thf_unitary_formula ((
+  Fmla (thf_pair_connective, [thf_unitary_formula1, thf_unitary_formula2])
+))
+thf_binary_formula : thf_binary_pair   (( thf_binary_pair ))
+                   | thf_binary_tuple  (( thf_binary_tuple ))
+                   | thf_binary_type   (( THF_type thf_binary_type ))
+thf_logic_formula : thf_binary_formula   (( thf_binary_formula ))
+                  | thf_unitary_formula  (( thf_unitary_formula ))
+                  | thf_type_formula     (( THF_typing thf_type_formula ))
+                  | thf_subtype          (( THF_type thf_subtype ))
+thf_formula : thf_logic_formula (( thf_logic_formula ))
+            | thf_sequent       (( thf_sequent ))
+
+formula_role : LOWER_WORD (( classify_role LOWER_WORD ))
+
+thf_annotated : THF LPAREN name COMMA formula_role COMMA thf_formula annotations RPAREN PERIOD ((
+  Annotated_Formula ((file_name, THFleft + 1, THFright + 1),
+   THF, name, formula_role, thf_formula, annotations)
+))
+
+tff_annotated : TFF LPAREN name COMMA formula_role COMMA tff_formula annotations RPAREN PERIOD ((
+  Annotated_Formula ((file_name, TFFleft + 1, TFFright + 1),
+   TFF, name, formula_role, tff_formula, annotations)
+))
+
+fof_annotated : FOF LPAREN name COMMA formula_role COMMA fof_formula annotations RPAREN PERIOD ((
+  Annotated_Formula ((file_name, FOFleft + 1, FOFright + 1),
+   FOF, name, formula_role, fof_formula, annotations)
+))
+
+cnf_annotated : CNF LPAREN name COMMA formula_role COMMA cnf_formula annotations RPAREN PERIOD ((
+  Annotated_Formula ((file_name, CNFleft + 1, CNFright + 1),
+   CNF, name, formula_role, cnf_formula, annotations)
+))
+
+annotated_formula : cnf_annotated (( cnf_annotated ))
+                  | fof_annotated (( fof_annotated ))
+                  | tff_annotated (( tff_annotated ))
+                  | thf_annotated (( thf_annotated ))
+
+include_ : INCLUDE LPAREN file_name formula_selection RPAREN PERIOD ((
+  Include (file_name, formula_selection)
+))
+
+formula_selection : COMMA LBRKT name_list RBRKT   (( name_list  ))
+                  |                               (( [] ))
+
+name_list : name COMMA name_list   (( name :: name_list ))
+          | name                   (( [name] ))
+
+name : atomic_word (( atomic_word ))
+     | integer     (( integer ))
+
+tptp_input : annotated_formula (( annotated_formula ))
+           | include_           (( include_ ))
+
+tptp_file : tptp_input tptp_file (( tptp_input :: tptp_file ))
+          | COMMENT tptp_file    (( tptp_file ))
+          |                      (( [] ))
+
+tptp : tptp_file (( tptp_file ))
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_interpret.ML	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,924 @@
+(*  Title:      HOL/TPTP/?/tptp_interpret.ML
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+Interprets TPTP problems in Isabelle/HOL.
+*)
+
+signature TPTP_INTERPRET =
+sig
+  (*signature extension: typing information for variables and constants
+  (note that the former isn't really part of the signature, but it's bundled
+  along for more configurability -- though it's probably useless and could be
+  dropped in the future.*)
+  type const_map = (string * term) list
+  type var_types = (string * typ option) list
+
+  (*mapping from THF types to Isabelle/HOL types. This map works over all
+  base types (i.e. THF functions must be interpreted as Isabelle/HOL functions).
+  The map must be total wrt THF types. Later on there could be a configuration
+  option to make a map extensible.*)
+  type type_map = (TPTP_Syntax.tptp_type * typ) list
+
+  (*inference info, when available, consists of the name of the inference rule
+  and the names of the inferences involved in the reasoning step.*)
+  type inference_info = (string * string list) option
+
+  (*a parsed annotated formula is represented as a triple consisting of
+  the formula's label, its role, and a constant function to its Isabelle/HOL
+  term meaning*)
+  type tptp_formula_meaning =
+    string * TPTP_Syntax.role * TPTP_Syntax.tptp_formula * term * inference_info
+
+  (*In general, the meaning of a TPTP statement (which, through the Include
+  directive, may include the meaning of an entire TPTP file, is an extended
+  Isabelle/HOL theory, an explicit map from constants to their Isabelle/HOL
+  counterparts and their types, the meaning of any TPTP formulas encountered
+  while interpreting that statement, and a map from THF to Isabelle/HOL types
+  (these types would have been added to the theory returned in the first position
+  of the tuple). The last value is NONE when the function which produced the
+  whole tptp_general_meaning value was given a type_map argument -- since
+  this means that the type_map is already known, and it has not been changed.*)
+  type tptp_general_meaning =
+    (type_map * const_map * tptp_formula_meaning list)
+
+  (*if problem_name is given then it is used to qualify types & constants*)
+  type config =
+    {(*init_type_map : type_map with_origin,
+     init_const_map : type_map with_origin,*)
+     cautious : bool,
+     problem_name : TPTP_Problem_Name.problem_name option
+     (*dont_build_maps : bool,
+     extend_given_type_map : bool,
+     extend_given_const_map : bool,
+     generative_type_interpretation : bool,
+     generative_const_interpretation : bool*)}
+
+  (*map types form THF to Isabelle/HOL*)
+  val interpret_type : config -> theory -> type_map ->
+    TPTP_Syntax.tptp_type -> typ
+
+  (*interpret a TPTP line: return an updated theory including the
+  types & constants which were specified in that formula, a map from
+  constant names to their types, and a map from constant names to Isabelle/HOL
+  constants; and a list possible formulas resulting from that line.
+  Note that type/constant declarations do not result in any formulas being
+  returned. A typical TPTP line might update the theory, or return a single
+  formula. The sole exception is the "include" directive which may update the
+  theory and also return a list of formulas from the included file.
+  Arguments:
+    cautious = indicates whether additional checks are made to check
+      that all used types have been declared.
+    type_map = mapping of THF-types to Isabelle/HOL types. This argument may be
+      given to force a specific mapping: this is usually done for using an
+      imported THF problem in Isar.
+    const_map = as previous, but for constants.
+    path_prefix = path where THF problems etc are located (to help the "include"
+      directive find the files.
+    thy = theory where interpreted info will be stored.
+  *)
+
+  (*map terms form TPTP to Isabelle/HOL*)
+  val interpret_term : bool -> config -> TPTP_Syntax.language -> theory ->
+    const_map -> var_types -> type_map -> TPTP_Syntax.tptp_term ->
+    term * theory
+
+  val interpret_formula : config -> TPTP_Syntax.language -> const_map ->
+    var_types -> type_map -> TPTP_Syntax.tptp_formula -> theory ->
+    term * theory
+
+  val interpret_line : config -> type_map -> const_map ->
+    Path.T -> TPTP_Syntax.tptp_line -> theory -> tptp_general_meaning * theory
+
+  (*Like "interpret_line" above, but works over a whole parsed problem.
+    Arguments:
+      new_basic_types = indicates whether interpretations of $i and $o
+        types are to be added to the type map (unless it is Given).
+        This is "true" if we are running this over a fresh THF problem, but
+        "false" if we are calling this _during_ the interpretation of a THF file
+        (i.e. when interpreting an "include" directive.
+      config = config
+      path_prefix = " "
+      contents = parsed TPTP syntax
+      type_map = see "interpret_line"
+      const_map = " "
+      thy = " "
+  *)
+  val interpret_problem : bool -> config -> Path.T ->
+    TPTP_Syntax.tptp_line list -> type_map -> const_map -> theory ->
+    tptp_general_meaning * theory
+
+  (*Like "interpret_problem" above, but it's given a filename rather than
+  a parsed problem.*)
+  val interpret_file : bool -> Path.T -> Path.T -> type_map -> const_map ->
+    theory -> tptp_general_meaning * theory
+
+  (*General exception for this signature.*)
+  exception MISINTERPRET_FORMULA of string * TPTP_Syntax.tptp_formula
+  exception MISINTERPRET_SYMBOL of string * TPTP_Syntax.symbol
+  exception MISINTERPRET_TERM of string * TPTP_Syntax.tptp_term
+  exception MISINTERPRET_TYPE of string * TPTP_Syntax.tptp_type
+
+  (*Generates a fresh Isabelle/HOL type for interpreting a THF type in a theory.*)
+  val declare_type : config -> (TPTP_Syntax.tptp_type * string) -> type_map ->
+    theory -> type_map * theory
+
+  (*Returns the list of all files included in a directory and its
+  subdirectories. This is only used for testing the parser/interpreter against
+  all THF problems.*)
+  val get_file_list : Path.T -> Path.T list
+
+  type manifest = TPTP_Problem_Name.problem_name * tptp_general_meaning
+
+  val get_manifests : theory -> manifest list
+
+  val import_file : bool -> Path.T -> Path.T -> type_map -> const_map ->
+    theory -> theory
+
+  val extract_inference_info : (TPTP_Syntax.general_term * 'a list) option ->
+                               (string * string list) option
+end
+
+structure TPTP_Interpret : TPTP_INTERPRET =
+struct
+
+open TPTP_Syntax
+exception MISINTERPRET_FORMULA of string * TPTP_Syntax.tptp_formula
+exception MISINTERPRET_SYMBOL of string * TPTP_Syntax.symbol
+exception MISINTERPRET_TERM of string * TPTP_Syntax.tptp_term
+exception MISINTERPRET_TYPE of string * TPTP_Syntax.tptp_type
+
+(* General stuff *)
+
+type config =
+  {(*init_type_map : type_map with_origin,
+   init_const_map : type_map with_origin,*)
+   cautious : bool,
+   problem_name : TPTP_Problem_Name.problem_name option
+   (*dont_build_maps : bool,
+   extend_given_type_map : bool,
+   extend_given_const_map : bool,
+   generative_type_interpretation : bool,
+   generative_const_interpretation : bool*)}
+
+(* Interpretation *)
+
+(** Signatures and other type abbrevations **)
+
+type const_map = (string * term) list
+type var_types = (string * typ option) list
+type type_map = (TPTP_Syntax.tptp_type * typ) list
+type inference_info = (string * string list) option
+type tptp_formula_meaning =
+  string * TPTP_Syntax.role * TPTP_Syntax.tptp_formula * term * inference_info
+type tptp_general_meaning =
+  (type_map * const_map * tptp_formula_meaning list)
+
+type manifest = TPTP_Problem_Name.problem_name * tptp_general_meaning
+structure TPTP_Data = Theory_Data
+  (type T = manifest list
+   val empty = []
+   val extend = I
+   val merge = Library.merge (op =))
+val get_manifests = TPTP_Data.get
+
+
+(** Adding types to a signature **)
+
+fun type_exists thy ty_name =
+  Sign.declared_tyname thy (Sign.full_bname thy ty_name)
+
+val IND_TYPE = "ind"
+
+fun mk_binding config ident =
+  let
+    val pre_binding = Binding.name ident
+  in
+    case #problem_name config of
+      NONE => pre_binding
+    | SOME prob =>
+        Binding.qualify
+         false
+         (TPTP_Problem_Name.mangle_problem_name prob)
+         pre_binding
+  end
+
+(*Returns updated theory and the name of the final type's name -- i.e. if the
+original name is already taken then the function looks for an available
+alternative. It also returns an updated type_map if one has been passed to it.*)
+fun declare_type (config : config) (thf_type, type_name) ty_map thy =
+  if type_exists thy type_name then
+    raise MISINTERPRET_TYPE ("Type already exists", Atom_type type_name)
+  else
+    let
+      val binding = mk_binding config type_name
+      val final_fqn = Sign.full_name thy binding
+      val thy' = Typedecl.typedecl_global (mk_binding config type_name, [], NoSyn) thy |> snd
+      val typ_map_entry = (thf_type, (Type (final_fqn, [])))
+      val ty_map' = typ_map_entry :: ty_map
+    in (ty_map', thy') end
+
+
+(** Adding constants to signature **)
+
+fun full_name thy config const_name =
+  Sign.full_name thy (mk_binding config const_name)
+
+fun const_exists config thy const_name =
+  Sign.declared_const thy (full_name thy config const_name)
+
+fun declare_constant config const_name ty thy =
+  if const_exists config thy const_name then
+    raise MISINTERPRET_TERM
+     ("Const already declared", Term_Func (Uninterpreted const_name, []))
+  else
+    Theory.specify_const
+     ((mk_binding config const_name, ty), NoSyn) thy
+
+
+(** Interpretation **)
+
+(*Types in THF are encoded as formulas. This function translates them to type form.*)
+(*FIXME possibly incomplete hack*)
+fun fmlatype_to_type (Atom (THF_Atom_term (Term_Func (TypeSymbol typ, [])))) =
+      Defined_type typ
+  | fmlatype_to_type (Atom (THF_Atom_term (Term_Func (Uninterpreted str, [])))) =
+      Atom_type str
+  | fmlatype_to_type (THF_type (Fn_type (ty1, ty2))) =
+      let
+        val ty1' = case ty1 of
+            Fn_type _ => fmlatype_to_type (THF_type ty1)
+          | Fmla_type ty1 => fmlatype_to_type ty1
+        val ty2' = case ty2 of
+            Fn_type _ => fmlatype_to_type (THF_type ty2)
+          | Fmla_type ty2 => fmlatype_to_type ty2
+      in Fn_type (ty1', ty2') end
+
+fun interpret_type config thy type_map thf_ty =
+  let
+    fun lookup_type ty_map thf_ty =
+      (case AList.lookup (op =) ty_map thf_ty of
+          NONE =>
+            raise MISINTERPRET_TYPE
+              ("Could not find the interpretation of this TPTP type in the " ^
+               "mapping to Isabelle/HOL", thf_ty)
+        | SOME ty => ty)
+  in
+    case thf_ty of (*FIXME make tail*)
+       Prod_type (thf_ty1, thf_ty2) =>
+         Type ("Product_Type.prod",
+          [interpret_type config thy type_map thf_ty1,
+           interpret_type config thy type_map thf_ty2])
+     | Fn_type (thf_ty1, thf_ty2) =>
+         Type ("fun",
+          [interpret_type config thy type_map thf_ty1,
+           interpret_type config thy type_map thf_ty2])
+     | Atom_type _ =>
+         lookup_type type_map thf_ty
+     | Defined_type tptp_base_type =>
+         (case tptp_base_type of
+            Type_Ind =>
+              lookup_type type_map thf_ty
+          | Type_Bool => HOLogic.boolT
+          | Type_Type =>
+              raise MISINTERPRET_TYPE
+               ("No type interpretation", thf_ty)
+          | Type_Int => Type ("Int.int", [])
+         (*FIXME these types are currently unsupported, so they're treated as
+         individuals*)
+          | Type_Rat =>
+              interpret_type config thy type_map (Defined_type Type_Ind)
+          | Type_Real =>
+              interpret_type config thy type_map (Defined_type Type_Ind)
+          )
+     | Sum_type _ =>
+         raise MISINTERPRET_TYPE (*FIXME*)
+          ("No type interpretation (sum type)", thf_ty)
+     | Fmla_type tptp_ty =>
+        fmlatype_to_type tptp_ty
+        |> interpret_type config thy type_map
+     | Subtype _ =>
+         raise MISINTERPRET_TYPE (*FIXME*)
+          ("No type interpretation (subtype)", thf_ty)
+  end
+
+fun the_const config thy language const_map_payload str =
+  if language = THF then
+    (case AList.lookup (op =) const_map_payload str of
+        NONE => raise MISINTERPRET_TERM
+          ("Could not find the interpretation of this constant in the " ^
+          "mapping to Isabelle/HOL", Term_Func (Uninterpreted str, []))
+      | SOME t => t)
+  else
+    if const_exists config thy str then
+       Sign.mk_const thy ((Sign.full_name thy (mk_binding config str)), [])
+    else raise MISINTERPRET_TERM
+          ("Could not find the interpretation of this constant in the " ^
+          "mapping to Isabelle/HOL: ", Term_Func (Uninterpreted str, []))
+
+(*Eta-expands Isabelle/HOL binary function.
+ "str" is the name of a polymorphic constant in Isabelle/HOL*)
+fun mk_fun str =
+  (Const (str, Term.dummyT) $ Bound 1 $ Bound 0)
+   |> (Term.absdummy Term.dummyT #> Term.absdummy Term.dummyT)
+(*As above, but swaps the function's arguments*)
+fun mk_swapped_fun str =
+  (Const (str, Term.dummyT) $ Bound 0 $ Bound 1)
+   |> (Term.absdummy Term.dummyT #> Term.absdummy Term.dummyT)
+
+fun dummy_THF () =
+  HOLogic.choice_const Term.dummyT $ Term.absdummy Term.dummyT @{term "True"}
+
+fun interpret_symbol config thy language const_map symb =
+  case symb of
+     Uninterpreted n =>
+       (*Constant would have been added to the map at the time of
+       declaration*)
+       the_const config thy language const_map n
+   | Interpreted_ExtraLogic interpreted_symbol =>
+       (case interpreted_symbol of (*FIXME incomplete,
+                                     and doesn't work with $i*)
+          Sum => mk_fun @{const_name Groups.plus}
+        | UMinus =>
+            (Const (@{const_name Groups.uminus}, Term.dummyT) $ Bound 0)
+             |> Term.absdummy Term.dummyT
+        | Difference => mk_fun @{const_name Groups.minus}
+        | Product => mk_fun @{const_name Groups.times}
+        | Quotient => mk_fun @{const_name Fields.divide}
+        | Less => mk_fun @{const_name Orderings.less}
+        | LessEq => mk_fun @{const_name Orderings.less_eq}
+        | Greater => mk_swapped_fun @{const_name Orderings.less}
+          (*FIXME would be nicer to expand "greater"
+            and "greater_eq" abbreviations*)
+        | GreaterEq => mk_swapped_fun @{const_name Orderings.less_eq}
+        (* FIXME todo
+        | Quotient_E | Quotient_T | Quotient_F | Remainder_E | Remainder_T
+        | Remainder_F | Floor | Ceiling | Truncate | Round | To_Int | To_Rat
+        | To_Real | EvalEq | Is_Int | Is_Rat*)
+        | Apply => raise MISINTERPRET_SYMBOL ("Malformed term?", symb)
+        | _ => dummy_THF ()
+        )
+   | Interpreted_Logic logic_symbol =>
+       (case logic_symbol of
+          Equals => HOLogic.eq_const Term.dummyT
+        | NEquals =>
+           HOLogic.mk_not (HOLogic.eq_const Term.dummyT $ Bound 1 $ Bound 0)
+           |> (Term.absdummy Term.dummyT #> Term.absdummy Term.dummyT)
+        | Or => HOLogic.disj
+        | And => HOLogic.conj
+        | Iff => HOLogic.eq_const HOLogic.boolT
+        | If => HOLogic.imp
+        | Fi => @{term "\<lambda> x. \<lambda> y. y \<longrightarrow> x"}
+        | Xor =>
+            @{term
+              "\<lambda> x. \<lambda> y. \<not> (x \<longleftrightarrow> y)"}
+        | Nor => @{term "\<lambda> x. \<lambda> y. \<not> (x \<or> y)"}
+        | Nand => @{term "\<lambda> x. \<lambda> y. \<not> (x \<and> y)"}
+        | Not => HOLogic.Not
+        | Op_Forall => HOLogic.all_const Term.dummyT
+        | Op_Exists => HOLogic.exists_const Term.dummyT
+        | True => @{term "True"}
+        | False => @{term "False"}
+        )
+   | TypeSymbol _ =>
+       raise MISINTERPRET_SYMBOL
+        ("Cannot have TypeSymbol in term", symb)
+   | System str =>
+       raise MISINTERPRET_SYMBOL
+        ("How to interpret system terms?", symb)
+
+(*Apply a function to a list of arguments*)
+val mapply = fold (fn x => fn y => y $ x)
+(*As above, but for products*)
+fun mtimes thy =
+  fold (fn x => fn y =>
+    Sign.mk_const thy
+    ("Product_Type.Pair",[Term.dummyT, Term.dummyT]) $ y $ x)
+
+(*Adds constants to signature in FOL. "formula_level" means that the constants
+are to be interpreted as predicates, otherwise as functions over individuals.*)
+fun FO_const_types config formula_level type_map tptp_t thy =
+  let
+    val ind_type = interpret_type config thy type_map (Defined_type Type_Ind)
+    val value_type =
+      if formula_level then
+        interpret_type config thy type_map (Defined_type Type_Bool)
+      else ind_type
+  in case tptp_t of
+      Term_Func (symb, tptp_ts) =>
+        let
+          val thy' = fold (FO_const_types config false type_map) tptp_ts thy
+          val ty = fold (fn ty1 => fn ty2 => Type ("fun", [ty1, ty2]))
+            (map (fn _ => ind_type) tptp_ts) value_type
+        in
+          case symb of
+             Uninterpreted const_name =>
+                 if const_exists config thy' const_name then thy'
+                 else snd (declare_constant config const_name ty thy')
+           | _ => thy'
+        end
+     | _ => thy
+  end
+
+(*like FO_const_types but works over symbols*)
+fun symb_const_types config type_map formula_level const_name n_args thy =
+  let
+    val ind_type = interpret_type config thy type_map (Defined_type Type_Ind)
+    val value_type =
+      if formula_level then
+        interpret_type config thy type_map (Defined_type Type_Bool)
+      else interpret_type config thy type_map (Defined_type Type_Ind)
+    fun mk_i_fun b n acc =
+      if n = 0 then acc b
+      else
+        mk_i_fun b (n - 1) (fn ty => Type ("fun", [ind_type, acc ty]))
+  in
+    if const_exists config thy const_name then thy
+    else
+      snd (declare_constant config const_name
+           (mk_i_fun value_type n_args I) thy)
+  end
+
+(*Next batch of functions give info about Isabelle/HOL types*)
+fun is_fun (Type ("fun", _)) = true
+  | is_fun _ = false
+fun is_prod (Type ("Product_Type.prod", _)) = true
+  | is_prod _ = false
+fun dom_type (Type ("fun", [ty1, _])) = ty1
+fun is_prod_typed thy config symb =
+  let
+    fun symb_type const_name =
+      Sign.the_const_type thy (full_name thy config const_name)
+  in
+    case symb of
+       Uninterpreted const_name =>
+         if is_fun (symb_type const_name) then
+           symb_type const_name |> dom_type |> is_prod
+         else false
+     | _ => false
+   end
+
+
+(*
+ Information needed to be carried around:
+ - constant mapping: maps constant names to Isabelle terms with fully-qualified
+    names. This is fixed, and it is determined by declaration lines earlier
+    in the script (i.e. constants must be declared before appearing in terms.
+ - type context: maps bound variables to their Isabelle type. This is discarded
+    after each call of interpret_term since variables' can't be bound across
+    terms.
+*)
+fun interpret_term formula_level config language thy const_map var_types type_map tptp_t =
+  case tptp_t of
+    Term_Func (symb, tptp_ts) =>
+       let
+         val thy' = FO_const_types config formula_level type_map tptp_t thy
+         fun typeof_const const_name = Sign.the_const_type thy'
+             (Sign.full_name thy' (mk_binding config const_name))
+       in
+         case symb of
+           Interpreted_ExtraLogic Apply =>
+             (*apply the head of the argument list to the tail*)
+             (mapply
+               (map (interpret_term false config language thy' const_map
+                var_types type_map #> fst) (tl tptp_ts))
+               (interpret_term formula_level config language thy' const_map
+                var_types type_map (hd tptp_ts) |> fst),
+              thy')
+           | _ =>
+              (*apply symb to tptp_ts*)
+              if is_prod_typed thy' config symb then
+                (interpret_symbol config thy' language const_map symb $
+                  mtimes thy'
+                  (map (interpret_term false config language thy' const_map
+                   var_types type_map #> fst) (tl tptp_ts))
+                  ((interpret_term false config language thy' const_map
+                   var_types type_map #> fst) (hd tptp_ts)),
+                 thy')
+              else
+                (mapply
+                  (map (interpret_term false config language thy' const_map
+                   var_types type_map #> fst) tptp_ts)
+                  (interpret_symbol config thy' language const_map symb),
+                 thy')
+       end
+  | Term_Var n =>
+     (if language = THF orelse language = TFF then
+        (case AList.lookup (op =) var_types n of
+           SOME ty =>
+             (case ty of
+                SOME ty => Free (n, ty)
+              | NONE => Free (n, Term.dummyT) (*to infer the variable's type*)
+             )
+         | NONE =>
+             raise MISINTERPRET_TERM ("Could not type variable", tptp_t))
+      else (*variables range over individuals*)
+        Free (n, interpret_type config thy type_map (Defined_type Type_Ind)),
+      thy)
+  | Term_Conditional (tptp_formula, tptp_t1, tptp_t2) =>
+      (mk_fun @{const_name If}, thy)
+  | Term_Num (number_kind, num) =>
+      let
+        val num_term =
+          if number_kind = Int_num then
+            HOLogic.mk_number @{typ "int"} (the (Int.fromString num))
+          else dummy_THF () (*FIXME: not yet supporting rationals and reals*)
+      in (num_term, thy) end
+  | Term_Distinct_Object str =>
+      let
+        fun alphanumerate c =
+          let
+            val c' = ord c
+            val out_of_range =
+              ((c' > 64) andalso (c' < 91)) orelse ((c' > 96)
+               andalso (c' < 123)) orelse ((c' > 47) andalso (c' < 58))
+          in
+            if (not out_of_range) andalso (not (c = "_")) then
+              "_nom_" ^ Int.toString (ord c) ^ "_"
+            else c
+          end
+        val mangle_name = raw_explode #> map alphanumerate #> implode
+      in declare_constant config (mangle_name str)
+          (interpret_type config thy type_map (Defined_type Type_Ind)) thy
+      end
+
+(*Given a list of "'a option", then applies a function to each element if that
+element is SOME value, otherwise it leaves it as NONE.*)
+fun map_opt f xs =
+  fold
+  (fn x => fn y =>
+    (if Option.isSome x then
+       SOME (f (Option.valOf x))
+     else NONE) :: y
+  ) xs []
+
+fun interpret_formula config language const_map var_types type_map tptp_fmla thy =
+  case tptp_fmla of
+      Pred app =>
+        interpret_term true config language thy const_map
+         var_types type_map (Term_Func app)
+    | Fmla (symbol, tptp_formulas) =>
+       (case symbol of
+          Interpreted_ExtraLogic Apply =>
+            let
+              val (args, thy') = (fold_map (interpret_formula config language
+               const_map var_types type_map) (tl tptp_formulas) thy)
+              val (func, thy') = interpret_formula config language const_map
+               var_types type_map (hd tptp_formulas) thy'
+            in (mapply args func, thy') end
+        | Uninterpreted const_name =>
+            let
+              val (args, thy') = (fold_map (interpret_formula config language
+               const_map var_types type_map) tptp_formulas thy)
+              val thy' = symb_const_types config type_map true const_name
+               (length tptp_formulas) thy'
+            in
+              (if is_prod_typed thy' config symbol then
+                 mtimes thy' args (interpret_symbol config thy' language
+                  const_map symbol)
+               else
+                mapply args
+                 (interpret_symbol config thy' language const_map symbol),
+              thy')
+            end
+        | _ =>
+            let
+              val (args, thy') =
+                fold_map
+                 (interpret_formula config language
+                  const_map var_types type_map)
+                 tptp_formulas thy
+            in
+              (if is_prod_typed thy' config symbol then
+                 mtimes thy' args (interpret_symbol config thy' language
+                  const_map symbol)
+               else
+                 mapply args
+                  (interpret_symbol config thy' language const_map symbol),
+               thy')
+            end)
+    | Sequent _ =>
+        raise MISINTERPRET_FORMULA ("Sequent", tptp_fmla) (*FIXME unsupported*)
+    | Quant (quantifier, bindings, tptp_formula) =>
+        let
+          val (bound_vars, bound_var_types) = ListPair.unzip bindings
+          val bound_var_types' : typ option list =
+            (map_opt : (tptp_type -> typ) -> (tptp_type option list) -> typ option list)
+            (interpret_type config thy type_map) bound_var_types
+          val var_types' =
+            ListPair.zip (bound_vars, List.rev bound_var_types')
+            |> List.rev
+          fun fold_bind f =
+            fold
+              (fn (n, ty) => fn t =>
+                 case ty of
+                   NONE =>
+                     f (n,
+                        if language = THF then Term.dummyT
+                        else
+                          interpret_type config thy type_map
+                           (Defined_type Type_Ind),
+                        t)
+                 | SOME ty => f (n, ty, t)
+              )
+              var_types'
+        in case quantifier of
+          Forall =>
+            interpret_formula config language const_map (var_types' @ var_types)
+             type_map tptp_formula thy
+            |>> fold_bind HOLogic.mk_all
+        | Exists =>
+            interpret_formula config language const_map (var_types' @ var_types)
+             type_map tptp_formula thy
+            |>> fold_bind HOLogic.mk_exists
+        | Lambda =>
+            interpret_formula config language const_map (var_types' @ var_types)
+             type_map tptp_formula thy
+            |>> fold_bind (fn (n, ty, rest) => lambda (Free (n, ty)) rest)
+        | Epsilon =>
+            let val (t, thy') =
+              interpret_formula config language const_map var_types type_map
+               (Quant (Lambda, bindings, tptp_formula)) thy
+            in ((HOLogic.choice_const Term.dummyT) $ t, thy') end
+        | Iota =>
+            let val (t, thy') =
+              interpret_formula config language const_map var_types type_map
+               (Quant (Lambda, bindings, tptp_formula)) thy
+            in (Const (@{const_name The}, Term.dummyT) $ t, thy') end
+        | Dep_Prod =>
+            raise MISINTERPRET_FORMULA ("Unsupported", tptp_fmla)
+        | Dep_Sum =>
+            raise MISINTERPRET_FORMULA ("Unsupported", tptp_fmla)
+        end
+  (*FIXME unsupported
+    | Conditional of tptp_formula * tptp_formula * tptp_formula
+    | Let of tptp_let list * tptp_formula
+
+    and tptp_let =
+        Let_fmla of (string * tptp_type option) * tptp_formula (*both TFF and THF*)
+      | Let_term of (string * tptp_type option) * tptp_term  (*only TFF*)
+  *)
+    | Atom tptp_atom =>
+        (case tptp_atom of
+          TFF_Typed_Atom (symbol, tptp_type_opt) =>
+            (*FIXME ignoring type info*)
+            (interpret_symbol config thy language const_map symbol, thy)
+        | THF_Atom_term tptp_term =>
+            interpret_term true config language thy const_map var_types
+             type_map tptp_term
+        | THF_Atom_conn_term symbol =>
+            (interpret_symbol config thy language const_map symbol, thy))
+    | THF_type _ => (*FIXME unsupported*)
+         raise MISINTERPRET_FORMULA
+          ("Cannot interpret types as formulas", tptp_fmla)
+    | THF_typing (tptp_formula, _) => (*FIXME ignoring type info*)
+        interpret_formula config language
+         const_map var_types type_map tptp_formula thy
+
+
+(*Extract name of inference rule, and the inferences it relies on*)
+(*This is tuned to work with LEO-II, and is brittle to upstream
+  changes of the proof protocol.*)
+fun extract_inference_info annot =
+  let
+    fun get_line_id (General_Data (Number (Int_num, num))) = [num]
+      | get_line_id (General_Data (Atomic_Word name)) = [name]
+      | get_line_id (General_Term (Number (Int_num, num), _ (*e.g. a bind*))) = [num]
+      | get_line_id _ = []
+        (*e.g. General_Data (Application ("theory", [General_Data
+          (Atomic_Word "equality")])) -- which would come from E through LEO-II*)
+  in
+    case annot of
+      NONE => NONE
+    | SOME annot =>
+      (case annot of
+        (General_Data (Application ("inference",
+        [General_Data (Atomic_Word inference_name),
+         _,
+         General_List related_lines])), _) =>
+          (SOME (inference_name, map get_line_id related_lines |> List.concat))
+      | _ => NONE)
+  end
+
+
+(*Extract the type from a typing*)
+local
+  fun extract_type tptp_type =
+    case tptp_type of
+        Fmla_type typ => fmlatype_to_type typ
+      | _ => tptp_type
+in
+  fun typeof_typing (THF_typing (_, tptp_type)) = extract_type tptp_type
+  fun typeof_tff_typing (Atom (TFF_Typed_Atom (_, SOME tptp_type))) =
+    extract_type tptp_type
+end
+fun nameof_typing
+  (THF_typing
+     ((Atom (THF_Atom_term (Term_Func (Uninterpreted str, [])))), _)) = str
+fun nameof_tff_typing (Atom (TFF_Typed_Atom (Uninterpreted str, _))) = str
+
+fun interpret_line config type_map const_map path_prefix line thy =
+  case line of
+     Include (quoted_path, _) => (*FIXME currently everything is imported*)
+       interpret_file'
+        false
+        config
+        path_prefix
+        (Path.append
+          path_prefix
+          (quoted_path
+           |> unenclose
+           |> Path.explode))
+        type_map
+        const_map
+        thy
+   | Annotated_Formula (pos, lang, label, role, tptp_formula, annotation_opt) =>
+       case role of
+         Role_Type =>
+           let
+             val thy_name = Context.theory_name thy
+             val (tptp_ty, name) =
+               if lang = THF then
+                 (typeof_typing tptp_formula (*assuming tptp_formula is a typing*),
+                  nameof_typing tptp_formula (*and that the LHS is a (atom) name*))
+               else
+                 (typeof_tff_typing tptp_formula,
+                  nameof_tff_typing tptp_formula)
+           in
+             case tptp_ty of
+               Defined_type Type_Type => (*add new type*)
+                 (*generate an Isabelle/HOL type to interpret this THF type,
+                 and add it to both the Isabelle/HOL theory and to the type_map*)
+                  let
+                    val (type_map', thy') =
+                      declare_type
+                       config
+                       (Atom_type name, name)
+                       type_map
+                       thy
+                  in ((type_map',
+                       const_map,
+                       []),
+                      thy')
+                  end
+
+             | _ => (*declaration of constant*)
+                (*Here we populate the map from constants to the Isabelle/HOL
+                terms they map to (which in turn contain their types).*)
+                let
+                  val ty = interpret_type config thy type_map tptp_ty
+                  (*make sure that the theory contains all the types appearing
+                  in this constant's signature. Exception is thrown if encounter
+                  an undeclared types.*)
+                  val (t, thy') =
+                    let
+                      fun analyse_type thy thf_ty =
+                         if #cautious config then
+                          case thf_ty of
+                             Fn_type (thf_ty1, thf_ty2) =>
+                               (analyse_type thy thf_ty1;
+                               analyse_type thy thf_ty2)
+                           | Atom_type ty_name =>
+                               if type_exists thy ty_name then ()
+                               else
+                                raise MISINTERPRET_TYPE
+                                 ("This type, in signature of " ^
+                                    name ^ " has not been declared.",
+                                   Atom_type ty_name)
+                           | _ => ()
+                         else () (*skip test if we're not being cautious.*)
+                    in
+                      analyse_type thy tptp_ty;
+                      declare_constant config name ty thy
+                    end
+                  (*register a mapping from name to t. Constants' type
+                  declarations should occur at most once, so it's justified to
+                  use "::". Note that here we use a constant's name rather
+                  than the possibly- new one, since all references in the
+                  theory will be to this name.*)
+                  val const_map' = ((name, t) :: const_map)
+                in ((type_map,(*type_map unchanged, since a constant's
+                                declaration shouldn't include any new types.*)
+                     const_map',(*typing table of constant was extended*)
+                     []),(*no formulas were interpreted*)
+                    thy')(*theory was extended with new constant*)
+                end
+           end
+
+       | _ => (*i.e. the AF is not a type declaration*)
+           let
+             (*gather interpreted term, and the map of types occurring in that term*)
+             val (t, thy') =
+               interpret_formula config lang
+                const_map [] type_map tptp_formula thy
+               |>> HOLogic.mk_Trueprop
+             (*type_maps grow monotonically, so return the newest value (type_mapped);
+             there's no need to unify it with the type_map parameter.*)
+           in
+            ((type_map, const_map,
+              [(label, role, tptp_formula,
+                Syntax.check_term (Proof_Context.init_global thy') t,
+                extract_inference_info annotation_opt)]), thy')
+           end
+
+and interpret_problem new_basic_types config path_prefix lines type_map const_map thy =
+      let
+        val thy_name = Context.theory_name thy
+        (*Add interpretation of $o and generate an Isabelle/HOL type to
+        interpret $i, unless we've been given a mapping of types.*)
+        val (thy', type_map') =
+          if not new_basic_types then (thy, type_map)
+          else
+            let
+              val (type_map', thy') =
+                declare_type config
+                 (Defined_type Type_Ind, IND_TYPE) type_map thy
+            in
+              (thy', type_map')
+            end
+      in
+        fold (fn line =>
+               fn ((type_map, const_map, lines), thy) =>
+                 let
+                   (*process the line, ignoring the type-context for variables*)
+                   val ((type_map', const_map', l'), thy') =
+                     interpret_line config type_map const_map path_prefix line thy
+                 in
+                   ((type_map',
+                     const_map',
+                     l' @ lines),(*append is sufficient, union would be overkill*)
+                    thy')
+                 end)
+             lines (*lines of the problem file*)
+             ((type_map', const_map, []), thy') (*initial values*)
+      end
+
+and interpret_file' new_basic_types config path_prefix file_name =
+  let
+    val file_name' = Path.expand file_name
+  in
+    if Path.is_absolute file_name' then
+      Path.implode file_name
+      |> TPTP_Parser.parse_file
+      |> interpret_problem new_basic_types config path_prefix
+    else error "Could not determine absolute path to file."
+  end
+
+and interpret_file cautious path_prefix file_name =
+  let
+    val config =
+      {cautious = cautious,
+       problem_name =
+        SOME (TPTP_Problem_Name.parse_problem_name ((Path.base #> Path.implode)
+         file_name))
+      }
+    handle _(*FIXME*) =>
+      {cautious = cautious,
+       problem_name = NONE
+      }
+  in interpret_file' true config path_prefix file_name end
+
+fun import_file cautious path_prefix file_name type_map const_map thy =
+  let
+    val prob_name =
+      TPTP_Problem_Name.parse_problem_name (Path.implode (Path.base file_name))
+      handle _(*FIXME*) => TPTP_Problem_Name.Nonstandard (Path.implode file_name)
+    val (result, thy') =
+      interpret_file cautious path_prefix file_name type_map const_map thy
+  in TPTP_Data.map (cons ((prob_name, result))) thy' end
+
+val _ =
+  Outer_Syntax.improper_command "import_tptp" "import TPTP problem"
+   Keyword.diag (*FIXME why diag?*)
+   (Parse.string >>
+    (fn file_name =>
+      Toplevel.theory (fn thy =>
+        import_file true (Path.explode file_name |> Path.dir)
+         (Path.explode file_name) [] [] thy
+        )))
+
+
+(*Used for debugging. Returns all files contained within a directory or its
+subdirectories. Follows symbolic links, filters away directories.*)
+fun get_file_list path =
+  let
+    fun check_file_entry f rest =
+      let
+        (*NOTE needed since don't have File.is_link and File.read_link*)
+        val f_str = Path.implode f
+      in
+        if File.is_dir f then
+          rest (*filter out directories*)
+        else if OS.FileSys.isLink f_str then
+          (*follow links -- NOTE this breaks if links are relative paths*)
+          check_file_entry (Path.explode (OS.FileSys.readLink f_str)) rest
+        else f :: rest
+      end
+  in
+    File.read_dir path
+    |> map
+        (Path.explode
+        #> Path.append path)
+    |> (fn l => fold check_file_entry l [])
+  end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_lexyacc.ML	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,5594 @@
+
+(******************************************************************)
+(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
+(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
+(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
+(******************************************************************)
+
+(*
+  This file is produced from the parser generated by ML-Yacc from the
+  source files tptp.lex and tptp.yacc.
+*)
+signature TPTP_TOKENS =
+sig
+type ('a,'b) token
+type svalue
+val ITE_T:  'a * 'a -> (svalue,'a) token
+val ITE_F:  'a * 'a -> (svalue,'a) token
+val CNF:  'a * 'a -> (svalue,'a) token
+val FOF:  'a * 'a -> (svalue,'a) token
+val TFF:  'a * 'a -> (svalue,'a) token
+val THF:  'a * 'a -> (svalue,'a) token
+val LET_TERM:  'a * 'a -> (svalue,'a) token
+val SUBTYPE:  'a * 'a -> (svalue,'a) token
+val ATOMIC_SYSTEM_WORD: (string) *  'a * 'a -> (svalue,'a) token
+val ATOMIC_DEFINED_WORD: (string) *  'a * 'a -> (svalue,'a) token
+val DEP_PROD:  'a * 'a -> (svalue,'a) token
+val DEP_SUM:  'a * 'a -> (svalue,'a) token
+val GENTZEN_ARROW:  'a * 'a -> (svalue,'a) token
+val TIMES:  'a * 'a -> (svalue,'a) token
+val PLUS:  'a * 'a -> (svalue,'a) token
+val OPERATOR_EXISTS:  'a * 'a -> (svalue,'a) token
+val OPERATOR_FORALL:  'a * 'a -> (svalue,'a) token
+val DEFIN_CHOICE:  'a * 'a -> (svalue,'a) token
+val INDEF_CHOICE:  'a * 'a -> (svalue,'a) token
+val DUD:  'a * 'a -> (svalue,'a) token
+val DISTINCT_OBJECT: (string) *  'a * 'a -> (svalue,'a) token
+val COMMENT: (string) *  'a * 'a -> (svalue,'a) token
+val LOWER_WORD: (string) *  'a * 'a -> (svalue,'a) token
+val UPPER_WORD: (string) *  'a * 'a -> (svalue,'a) token
+val SINGLE_QUOTED: (string) *  'a * 'a -> (svalue,'a) token
+val DOT_DECIMAL: (string) *  'a * 'a -> (svalue,'a) token
+val UNSIGNED_INTEGER: (string) *  'a * 'a -> (svalue,'a) token
+val SIGNED_INTEGER: (string) *  'a * 'a -> (svalue,'a) token
+val RATIONAL: (string) *  'a * 'a -> (svalue,'a) token
+val REAL: (string) *  'a * 'a -> (svalue,'a) token
+val DTFF:  'a * 'a -> (svalue,'a) token
+val DFOT:  'a * 'a -> (svalue,'a) token
+val DCNF:  'a * 'a -> (svalue,'a) token
+val DFOF:  'a * 'a -> (svalue,'a) token
+val DTHF:  'a * 'a -> (svalue,'a) token
+val EOF:  'a * 'a -> (svalue,'a) token
+val VLINE:  'a * 'a -> (svalue,'a) token
+val TOK_TYPE:  'a * 'a -> (svalue,'a) token
+val TOK_TRUE:  'a * 'a -> (svalue,'a) token
+val TOK_RAT:  'a * 'a -> (svalue,'a) token
+val TOK_REAL:  'a * 'a -> (svalue,'a) token
+val TOK_INT:  'a * 'a -> (svalue,'a) token
+val TOK_O:  'a * 'a -> (svalue,'a) token
+val TOK_I:  'a * 'a -> (svalue,'a) token
+val TOK_FALSE:  'a * 'a -> (svalue,'a) token
+val TILDE:  'a * 'a -> (svalue,'a) token
+val RPAREN:  'a * 'a -> (svalue,'a) token
+val RBRKT:  'a * 'a -> (svalue,'a) token
+val QUESTION:  'a * 'a -> (svalue,'a) token
+val PPLUS:  'a * 'a -> (svalue,'a) token
+val PERIOD:  'a * 'a -> (svalue,'a) token
+val NOR:  'a * 'a -> (svalue,'a) token
+val XOR:  'a * 'a -> (svalue,'a) token
+val NEQUALS:  'a * 'a -> (svalue,'a) token
+val NAND:  'a * 'a -> (svalue,'a) token
+val MMINUS:  'a * 'a -> (svalue,'a) token
+val MAP_TO:  'a * 'a -> (svalue,'a) token
+val LPAREN:  'a * 'a -> (svalue,'a) token
+val LBRKT:  'a * 'a -> (svalue,'a) token
+val LAMBDA:  'a * 'a -> (svalue,'a) token
+val INCLUDE:  'a * 'a -> (svalue,'a) token
+val IMPLIES:  'a * 'a -> (svalue,'a) token
+val IFF:  'a * 'a -> (svalue,'a) token
+val IF:  'a * 'a -> (svalue,'a) token
+val ARROW:  'a * 'a -> (svalue,'a) token
+val LET:  'a * 'a -> (svalue,'a) token
+val EXCLAMATION:  'a * 'a -> (svalue,'a) token
+val EQUALS:  'a * 'a -> (svalue,'a) token
+val COMMA:  'a * 'a -> (svalue,'a) token
+val COLON:  'a * 'a -> (svalue,'a) token
+val CARET:  'a * 'a -> (svalue,'a) token
+val AT_SIGN:  'a * 'a -> (svalue,'a) token
+val AMPERSAND:  'a * 'a -> (svalue,'a) token
+end
+signature TPTP_LRVALS=
+sig
+structure Tokens : TPTP_TOKENS
+structure ParserData:PARSER_DATA
+sharing type ParserData.Token.token = Tokens.token
+sharing type ParserData.svalue = Tokens.svalue
+end
+functor TPTPLexFun(structure Tokens: TPTP_TOKENS)=
+   struct
+    structure UserDeclarations =
+      struct
+(*  Title:      HOL/TPTP/TPTP_Parser/tptp.lex
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+ Notes:
+ * Omit %full in definitions to restrict alphabet to ascii.
+ * Could include %posarg to ensure that start counting character positions from
+   0, but it would punish performance.
+ * %s AF F COMMENT; -- could improve by making stateful.
+
+ Acknowledgements:
+ * Geoff Sutcliffe for help with TPTP.
+ * Timothy Bourke for his tips on getting ML-Yacc working with Poly/ML.
+ * An early version of this was ported from the specification shipped with
+   Leo-II, written by Frank Theiss.
+ * Some boilerplate bits were taken from the ml-yacc/ml-lex manual by Roger Price.
+ * Jasmin Blanchette and Makarius Wenzel for help with Isabelle integration.
+*)
+
+structure T = Tokens
+type pos = int             (* Position in file *)
+type lineNo = int
+type svalue = T.svalue
+type ('a,'b) token = ('a,'b) T.token
+type lexresult = (svalue,pos) token
+type lexarg = string
+type arg = lexarg
+val col = Unsynchronized.ref 0;
+val linep = Unsynchronized.ref 1;         (* Line pointer *)
+val eolpos = Unsynchronized.ref 0;
+
+val badCh : string * string * int * int -> unit = fn
+    (file_name, bad, line, col) =>
+    TextIO.output(TextIO.stdOut, file_name ^ "["
+          ^ Int.toString line ^ "." ^ Int.toString col
+          ^ "] Invalid character \"" ^ bad ^ "\"\n");
+
+val eof = fn file_name =>
+  let
+    val result = T.EOF (!linep,!col);
+    val _ = linep := 0;
+  in result end
+(*here could check whether file ended prematurely:
+ see if have open brackets, or if we're in some state other than INITIAL*)
+
+val count_commentlines : string -> unit = fn str =>
+  let
+    val str' = String.explode str
+    val newlines = List.filter (fn x => x = #"\n") str'
+  in linep := (!linep) + (List.length newlines) end
+
+end (* end of user routines *)
+exception LexError (* raised if illegal leaf action tried *)
+structure Internal =
+	struct
+
+datatype yyfinstate = N of int
+type statedata = {fin : yyfinstate list, trans: string}
+(* transition & final state table *)
+val tab = let
+val s = [ 
+ (0, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (1, 
+"\000\000\000\000\000\000\000\000\000\134\136\000\000\135\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\134\130\124\000\102\090\089\083\082\081\080\078\077\072\070\057\
+\\048\048\048\048\048\048\048\048\048\048\045\000\039\037\036\033\
+\\030\029\029\029\029\029\029\029\029\029\029\029\029\029\029\029\
+\\029\029\029\029\029\029\029\029\029\029\029\028\000\027\026\000\
+\\000\007\007\023\007\007\020\007\007\013\007\007\007\007\007\007\
+\\007\007\007\007\008\007\007\007\007\007\007\000\006\000\003\000\
+\\000"
+),
+ (3, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\005\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\004\000\000\000\
+\\000"
+),
+ (7, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (8, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\007\007\011\007\009\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (9, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\007\007\010\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (11, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\007\007\012\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (13, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\014\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (14, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\015\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (15, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\007\007\007\007\007\007\007\007\016\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (16, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\017\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (17, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\018\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (18, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\007\019\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (20, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\021\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (21, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\007\007\022\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (23, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\024\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (24, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\000\
+\\000\007\007\007\007\007\007\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\007\
+\\000\007\007\007\007\007\025\007\007\007\007\007\007\007\007\007\
+\\007\007\007\007\007\007\007\007\007\007\007\000\000\000\000\000\
+\\000"
+),
+ (29, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\029\029\029\029\029\029\029\029\029\029\000\000\000\000\000\000\
+\\000\029\029\029\029\029\029\029\029\029\029\029\029\029\029\029\
+\\029\029\029\029\029\029\029\029\029\029\029\000\000\000\000\029\
+\\000\029\029\029\029\029\029\029\029\029\029\029\029\029\029\029\
+\\029\029\029\029\029\029\029\029\029\029\029\000\000\000\000\000\
+\\000"
+),
+ (30, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\032\000\031\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (33, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\035\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\034\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (37, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\038\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (39, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\044\042\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\040\000\
+\\000"
+),
+ (40, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\041\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (42, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\043\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (45, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\047\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\046\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (48, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\051\049\
+\\048\048\048\048\048\048\048\048\048\048\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (49, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\050\050\050\050\050\050\050\050\050\050\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (51, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\052\052\052\052\052\052\052\052\052\052\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (52, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\052\052\052\052\052\052\052\052\052\052\000\000\000\000\000\000\
+\\000\000\000\000\000\053\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\053\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (53, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\055\000\055\000\000\
+\\054\054\054\054\054\054\054\054\054\054\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (54, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\054\054\054\054\054\054\054\054\054\054\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (55, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\056\056\056\056\056\056\056\056\056\056\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (57, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\058\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (58, 
+"\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\059\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058"
+),
+ (59, 
+"\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\059\058\058\058\058\060\
+\\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\058\
+\\058"
+),
+ (60, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\064\000\000\000\000\000\000\000\000\000\061\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (61, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\062\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (62, 
+"\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\063\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062"
+),
+ (63, 
+"\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\063\062\062\062\062\060\
+\\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\062\
+\\062"
+),
+ (64, 
+"\064\064\064\064\064\064\064\064\064\064\000\064\064\064\064\064\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\065\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\
+\\064"
+),
+ (65, 
+"\064\064\064\064\064\064\064\064\064\064\000\064\064\064\064\064\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\
+\\064\064\064\064\064\064\064\064\064\064\066\064\064\064\064\065\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\
+\\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\064\
+\\064"
+),
+ (66, 
+"\066\066\066\066\066\066\066\066\066\066\062\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\069\066\066\066\066\067\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066"
+),
+ (67, 
+"\066\066\066\066\066\066\066\066\066\066\062\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\068\066\066\066\066\067\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066"
+),
+ (69, 
+"\066\066\066\066\066\066\066\066\066\066\062\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\069\066\066\066\066\065\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\066\
+\\066"
+),
+ (70, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\071\071\071\071\071\071\071\071\071\071\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (72, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\075\000\000\
+\\074\074\074\074\074\074\074\074\074\074\000\000\000\000\073\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (74, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\051\049\
+\\074\074\074\074\074\074\074\074\074\074\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (75, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\076\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (78, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\079\000\000\000\000\
+\\074\074\074\074\074\074\074\074\074\074\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (83, 
+"\000\000\000\000\000\000\000\000\000\084\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\084\000\000\000\000\000\000\000\084\084\000\084\084\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\088\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\000\
+\\000"
+),
+ (84, 
+"\000\000\000\000\000\000\000\000\000\084\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\084\000\000\000\000\000\000\087\084\084\000\084\084\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\085\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\000\
+\\000"
+),
+ (85, 
+"\000\000\000\000\000\000\000\000\000\084\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\084\000\000\000\000\000\000\086\084\084\000\084\084\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\085\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\
+\\084\084\084\084\084\084\084\084\084\084\084\084\084\084\084\000\
+\\000"
+),
+ (90, 
+"\090\090\090\090\090\090\090\090\090\090\000\090\090\090\090\090\
+\\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\
+\\090\090\090\090\090\101\090\090\090\090\090\090\090\090\090\091\
+\\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\
+\\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\
+\\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\
+\\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\
+\\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\
+\\090"
+),
+ (91, 
+"\090\090\090\090\090\090\090\090\090\090\000\090\090\090\090\090\
+\\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\
+\\090\090\090\090\090\101\090\090\090\090\092\090\090\090\090\091\
+\\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\
+\\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\
+\\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\
+\\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\
+\\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\090\
+\\090"
+),
+ (92, 
+"\092\092\092\092\092\092\092\092\092\092\062\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\096\092\092\092\092\095\092\092\092\092\093\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092"
+),
+ (93, 
+"\092\092\092\092\092\092\092\092\092\092\062\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\096\092\092\092\092\094\092\092\092\092\093\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092"
+),
+ (95, 
+"\092\092\092\092\092\092\092\092\092\092\062\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\096\092\092\092\092\095\092\092\092\092\091\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\092\
+\\092"
+),
+ (96, 
+"\096\096\096\096\096\096\096\096\096\096\062\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\099\096\096\096\096\097\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096"
+),
+ (97, 
+"\096\096\096\096\096\096\096\096\096\096\062\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\098\096\096\096\096\097\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096"
+),
+ (99, 
+"\096\096\096\096\096\096\096\096\096\096\062\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\099\096\096\096\096\100\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\096\
+\\096"
+),
+ (100, 
+"\101\101\101\101\101\101\101\101\101\101\000\101\101\101\101\101\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\
+\\101\101\101\101\101\101\101\101\101\101\096\101\101\101\101\100\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\
+\\101"
+),
+ (101, 
+"\101\101\101\101\101\101\101\101\101\101\000\101\101\101\101\101\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\100\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\
+\\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\101\
+\\101"
+),
+ (102, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\122\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\103\103\119\103\103\115\103\103\109\103\103\103\103\103\103\
+\\103\103\103\103\104\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (103, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\000\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\103\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (104, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\000\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\103\
+\\000\103\103\103\103\103\107\103\105\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (105, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\000\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\103\
+\\000\103\103\103\103\103\106\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (107, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\000\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\103\
+\\000\103\103\103\103\103\108\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (109, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\000\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\103\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\110\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (110, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\000\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\103\
+\\000\103\103\103\103\111\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (111, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\000\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\112\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (112, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\000\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\103\
+\\000\103\103\103\103\103\114\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\113\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (115, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\000\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\103\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\116\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (116, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\000\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\103\
+\\000\103\103\103\103\103\118\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\117\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (119, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\000\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\103\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\120\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (120, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\000\
+\\000\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\103\
+\\000\103\103\103\103\103\121\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
+),
+ (122, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\123\123\123\123\123\123\123\123\123\123\123\123\123\123\123\
+\\123\123\123\123\123\123\123\123\123\123\123\000\000\000\000\000\
+\\000"
+),
+ (123, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\123\123\123\123\123\123\123\123\123\123\000\000\000\000\000\000\
+\\000\123\123\123\123\123\123\123\123\123\123\123\123\123\123\123\
+\\123\123\123\123\123\123\123\123\123\123\123\000\000\000\000\123\
+\\000\123\123\123\123\123\123\123\123\123\123\123\123\123\123\123\
+\\123\123\123\123\123\123\123\123\123\123\123\000\000\000\000\000\
+\\000"
+),
+ (124, 
+"\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\000\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\129\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125"
+),
+ (125, 
+"\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\128\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\126\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125"
+),
+ (126, 
+"\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\127\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\126\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\125\
+\\125"
+),
+ (130, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\133\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\132\131\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (134, 
+"\000\000\000\000\000\000\000\000\000\134\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\134\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+ (135, 
+"\000\000\000\000\000\000\000\000\000\000\136\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000"
+),
+(0, "")]
+fun f x = x 
+val s = map f (rev (tl (rev s))) 
+exception LexHackingError 
+fun look ((j,x)::r, i: int) = if i = j then x else look(r, i) 
+  | look ([], i) = raise LexHackingError
+fun g {fin=x, trans=i} = {fin=x, trans=look(s,i)} 
+in Vector.fromList(map g 
+[{fin = [], trans = 0},
+{fin = [(N 2)], trans = 1},
+{fin = [(N 2)], trans = 1},
+{fin = [(N 84)], trans = 3},
+{fin = [(N 71)], trans = 0},
+{fin = [(N 61)], trans = 0},
+{fin = [(N 86)], trans = 0},
+{fin = [(N 251)], trans = 7},
+{fin = [(N 251)], trans = 8},
+{fin = [(N 251)], trans = 9},
+{fin = [(N 186),(N 251)], trans = 7},
+{fin = [(N 251)], trans = 11},
+{fin = [(N 198),(N 251)], trans = 7},
+{fin = [(N 251)], trans = 13},
+{fin = [(N 251)], trans = 14},
+{fin = [(N 251)], trans = 15},
+{fin = [(N 251)], trans = 16},
+{fin = [(N 251)], trans = 17},
+{fin = [(N 251)], trans = 18},
+{fin = [(N 206),(N 251)], trans = 7},
+{fin = [(N 251)], trans = 20},
+{fin = [(N 251)], trans = 21},
+{fin = [(N 190),(N 251)], trans = 7},
+{fin = [(N 251)], trans = 23},
+{fin = [(N 251)], trans = 24},
+{fin = [(N 194),(N 251)], trans = 7},
+{fin = [(N 25)], trans = 0},
+{fin = [(N 80)], trans = 0},
+{fin = [(N 50)], trans = 0},
+{fin = [(N 157)], trans = 29},
+{fin = [(N 23)], trans = 30},
+{fin = [(N 15)], trans = 0},
+{fin = [(N 12)], trans = 0},
+{fin = [(N 78)], trans = 33},
+{fin = [(N 21)], trans = 0},
+{fin = [(N 283)], trans = 0},
+{fin = [(N 38)], trans = 0},
+{fin = [(N 31)], trans = 37},
+{fin = [(N 48)], trans = 0},
+{fin = [], trans = 39},
+{fin = [], trans = 40},
+{fin = [(N 68)], trans = 0},
+{fin = [(N 41)], trans = 42},
+{fin = [(N 45)], trans = 0},
+{fin = [(N 277)], trans = 0},
+{fin = [(N 27)], trans = 45},
+{fin = [(N 36)], trans = 0},
+{fin = [(N 286)], trans = 0},
+{fin = [(N 126)], trans = 48},
+{fin = [], trans = 49},
+{fin = [(N 104)], trans = 49},
+{fin = [], trans = 51},
+{fin = [(N 119)], trans = 52},
+{fin = [], trans = 53},
+{fin = [(N 119)], trans = 54},
+{fin = [], trans = 55},
+{fin = [(N 119)], trans = 55},
+{fin = [], trans = 57},
+{fin = [], trans = 58},
+{fin = [], trans = 59},
+{fin = [(N 182)], trans = 60},
+{fin = [], trans = 61},
+{fin = [], trans = 62},
+{fin = [], trans = 63},
+{fin = [(N 182)], trans = 64},
+{fin = [(N 182)], trans = 65},
+{fin = [(N 182)], trans = 66},
+{fin = [(N 182)], trans = 67},
+{fin = [(N 182)], trans = 66},
+{fin = [(N 182)], trans = 69},
+{fin = [(N 73)], trans = 70},
+{fin = [(N 130)], trans = 70},
+{fin = [], trans = 72},
+{fin = [(N 55)], trans = 0},
+{fin = [(N 123)], trans = 74},
+{fin = [(N 58)], trans = 75},
+{fin = [(N 274)], trans = 0},
+{fin = [(N 29)], trans = 0},
+{fin = [(N 268)], trans = 78},
+{fin = [(N 76)], trans = 0},
+{fin = [(N 270)], trans = 0},
+{fin = [(N 82)], trans = 0},
+{fin = [(N 52)], trans = 0},
+{fin = [], trans = 83},
+{fin = [], trans = 84},
+{fin = [], trans = 85},
+{fin = [(N 151)], trans = 84},
+{fin = [(N 151)], trans = 0},
+{fin = [], trans = 85},
+{fin = [(N 9)], trans = 0},
+{fin = [(N 182)], trans = 90},
+{fin = [(N 182)], trans = 91},
+{fin = [(N 182)], trans = 92},
+{fin = [(N 182)], trans = 93},
+{fin = [(N 182)], trans = 92},
+{fin = [(N 182)], trans = 95},
+{fin = [(N 182)], trans = 96},
+{fin = [(N 182)], trans = 97},
+{fin = [(N 182)], trans = 96},
+{fin = [(N 182)], trans = 99},
+{fin = [(N 182)], trans = 100},
+{fin = [(N 182)], trans = 101},
+{fin = [], trans = 102},
+{fin = [(N 266)], trans = 103},
+{fin = [(N 266)], trans = 104},
+{fin = [(N 266)], trans = 105},
+{fin = [(N 211),(N 266)], trans = 103},
+{fin = [(N 266)], trans = 107},
+{fin = [(N 231),(N 266)], trans = 103},
+{fin = [(N 266)], trans = 109},
+{fin = [(N 266)], trans = 110},
+{fin = [(N 266)], trans = 111},
+{fin = [(N 266)], trans = 112},
+{fin = [(N 245),(N 266)], trans = 103},
+{fin = [(N 238),(N 266)], trans = 103},
+{fin = [(N 266)], trans = 115},
+{fin = [(N 266)], trans = 116},
+{fin = [(N 226),(N 266)], trans = 103},
+{fin = [(N 216),(N 266)], trans = 103},
+{fin = [(N 266)], trans = 119},
+{fin = [(N 266)], trans = 120},
+{fin = [(N 221),(N 266)], trans = 103},
+{fin = [], trans = 122},
+{fin = [(N 259)], trans = 123},
+{fin = [], trans = 124},
+{fin = [], trans = 125},
+{fin = [], trans = 126},
+{fin = [(N 95)], trans = 125},
+{fin = [(N 95)], trans = 0},
+{fin = [], trans = 126},
+{fin = [(N 33)], trans = 130},
+{fin = [(N 280)], trans = 0},
+{fin = [(N 64)], trans = 0},
+{fin = [(N 18)], trans = 0},
+{fin = [(N 2)], trans = 134},
+{fin = [(N 7)], trans = 135},
+{fin = [(N 7)], trans = 0}])
+end
+structure StartStates =
+	struct
+	datatype yystartstate = STARTSTATE of int
+
+(* start state definitions *)
+
+val INITIAL = STARTSTATE 1;
+
+end
+type result = UserDeclarations.lexresult
+	exception LexerError (* raised if illegal leaf action tried *)
+end
+
+fun makeLexer yyinput =
+let	val yygone0=1
+	val yyb = Unsynchronized.ref "\n" 		(* buffer *)
+	val yybl = Unsynchronized.ref 1		(*buffer length *)
+	val yybufpos = Unsynchronized.ref 1		(* location of next character to use *)
+	val yygone = Unsynchronized.ref yygone0	(* position in file of beginning of buffer *)
+	val yydone = Unsynchronized.ref false		(* eof found yet? *)
+	val yybegin = Unsynchronized.ref 1		(*Current 'start state' for lexer *)
+
+	val YYBEGIN = fn (Internal.StartStates.STARTSTATE x) =>
+		 yybegin := x
+
+fun lex (yyarg as (file_name:string)) =
+let fun continue() : Internal.result = 
+  let fun scan (s,AcceptingLeaves : Internal.yyfinstate list list,l,i0) =
+	let fun action (i,nil) = raise LexError
+	| action (i,nil::l) = action (i-1,l)
+	| action (i,(node::acts)::l) =
+		case node of
+		    Internal.N yyk => 
+			(let fun yymktext() = substring(!yyb,i0,i-i0)
+			     val yypos = i0+ !yygone
+			open UserDeclarations Internal.StartStates
+ in (yybufpos := i; case yyk of 
+
+			(* Application actions *)
+
+  104 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.RATIONAL(yytext,!linep,!col) end
+| 119 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.REAL(yytext,!linep,!col) end
+| 12 => (col:=yypos-(!eolpos); T.INDEF_CHOICE(!linep,!col))
+| 123 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.SIGNED_INTEGER(yytext,!linep,!col) end
+| 126 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.UNSIGNED_INTEGER(yytext,!linep,!col) end
+| 130 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.DOT_DECIMAL(yytext,!linep,!col) end
+| 15 => (col:=yypos-(!eolpos); T.DEFIN_CHOICE(!linep,!col))
+| 151 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.SINGLE_QUOTED(yytext,!linep,!col) end
+| 157 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.UPPER_WORD(yytext,!linep,!col) end
+| 18 => (col:=yypos-(!eolpos); T.OPERATOR_FORALL(!linep,!col))
+| 182 => let val yytext=yymktext() in col:=yypos-(!eolpos); count_commentlines yytext;T.COMMENT(yytext,!linep,!col) end
+| 186 => (col:=yypos-(!eolpos); T.THF(!linep,!col))
+| 190 => (col:=yypos-(!eolpos); T.FOF(!linep,!col))
+| 194 => (col:=yypos-(!eolpos); T.CNF(!linep,!col))
+| 198 => (col:=yypos-(!eolpos); T.TFF(!linep,!col))
+| 2 => let val yytext=yymktext() in col:=(!col)+size yytext; continue ()  end
+| 206 => (col:=yypos-(!eolpos); T.INCLUDE(!linep,!col))
+| 21 => (col:=yypos-(!eolpos); T.OPERATOR_EXISTS(!linep,!col))
+| 211 => (col:=yypos-(!eolpos); T.DTHF(!linep,!col))
+| 216 => (col:=yypos-(!eolpos); T.DFOF(!linep,!col))
+| 221 => (col:=yypos-(!eolpos); T.DCNF(!linep,!col))
+| 226 => (col:=yypos-(!eolpos); T.DFOT(!linep,!col))
+| 23 => (col:=yypos-(!eolpos); T.AT_SIGN(!linep,!col))
+| 231 => (col:=yypos-(!eolpos); T.DTFF(!linep,!col))
+| 238 => (col:=yypos-(!eolpos); T.ITE_F(!linep,!col))
+| 245 => (col:=yypos-(!eolpos); T.ITE_T(!linep,!col))
+| 25 => (col:=yypos-(!eolpos); T.CARET(!linep,!col))
+| 251 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.LOWER_WORD(yytext,!linep,!col) end
+| 259 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.ATOMIC_SYSTEM_WORD(yytext,!linep,!col) end
+| 266 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.ATOMIC_DEFINED_WORD(yytext,!linep,!col) end
+| 268 => (col:=yypos-(!eolpos); T.PLUS(!linep,!col))
+| 27 => (col:=yypos-(!eolpos); T.COLON(!linep,!col))
+| 270 => (col:=yypos-(!eolpos); T.TIMES(!linep,!col))
+| 274 => (col:=yypos-(!eolpos); T.GENTZEN_ARROW(!linep,!col))
+| 277 => (col:=yypos-(!eolpos); T.SUBTYPE(!linep,!col))
+| 280 => (col:=yypos-(!eolpos); T.DEP_PROD(!linep,!col))
+| 283 => (col:=yypos-(!eolpos); T.DEP_SUM(!linep,!col))
+| 286 => (col:=yypos-(!eolpos); T.LET_TERM(!linep,!col))
+| 29 => (col:=yypos-(!eolpos); T.COMMA(!linep,!col))
+| 31 => (col:=yypos-(!eolpos); T.EQUALS(!linep,!col))
+| 33 => (col:=yypos-(!eolpos); T.EXCLAMATION(!linep,!col))
+| 36 => (col:=yypos-(!eolpos); T.LET(!linep,!col))
+| 38 => (col:=yypos-(!eolpos); T.ARROW(!linep,!col))
+| 41 => (col:=yypos-(!eolpos); T.IF(!linep,!col))
+| 45 => (col:=yypos-(!eolpos); T.IFF(!linep,!col))
+| 48 => (col:=yypos-(!eolpos); T.IMPLIES(!linep,!col))
+| 50 => (col:=yypos-(!eolpos); T.LBRKT(!linep,!col))
+| 52 => (col:=yypos-(!eolpos); T.LPAREN(!linep,!col))
+| 55 => (col:=yypos-(!eolpos); T.MAP_TO(!linep,!col))
+| 58 => (col:=yypos-(!eolpos); T.MMINUS(!linep,!col))
+| 61 => (col:=yypos-(!eolpos); T.NAND(!linep,!col))
+| 64 => (col:=yypos-(!eolpos); T.NEQUALS(!linep,!col))
+| 68 => (col:=yypos-(!eolpos); T.XOR(!linep,!col))
+| 7 => let val yytext=yymktext() in linep:=(!linep)+1;
+                   eolpos:=yypos+size yytext; continue () end
+| 71 => (col:=yypos-(!eolpos); T.NOR(!linep,!col))
+| 73 => (col:=yypos-(!eolpos); T.PERIOD(!linep,!col))
+| 76 => (col:=yypos-(!eolpos); T.PPLUS(!linep,!col))
+| 78 => (col:=yypos-(!eolpos); T.QUESTION(!linep,!col))
+| 80 => (col:=yypos-(!eolpos); T.RBRKT(!linep,!col))
+| 82 => (col:=yypos-(!eolpos); T.RPAREN(!linep,!col))
+| 84 => (col:=yypos-(!eolpos); T.TILDE(!linep,!col))
+| 86 => (col:=yypos-(!eolpos); T.VLINE(!linep,!col))
+| 9 => (col:=yypos-(!eolpos); T.AMPERSAND(!linep,!col))
+| 95 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.DISTINCT_OBJECT(yytext,!linep,!col) end
+| _ => raise Internal.LexerError
+
+		) end )
+
+	val {fin,trans} = Vector.sub(Internal.tab, s)
+	val NewAcceptingLeaves = fin::AcceptingLeaves
+	in if l = !yybl then
+	     if trans = #trans(Vector.sub(Internal.tab,0))
+	       then action(l,NewAcceptingLeaves
+) else	    let val newchars= if !yydone then "" else yyinput 1024
+	    in if (size newchars)=0
+		  then (yydone := true;
+		        if (l=i0) then UserDeclarations.eof yyarg
+		                  else action(l,NewAcceptingLeaves))
+		  else (if i0=l then yyb := newchars
+		     else yyb := substring(!yyb,i0,l-i0)^newchars;
+		     yygone := !yygone+i0;
+		     yybl := size (!yyb);
+		     scan (s,AcceptingLeaves,l-i0,0))
+	    end
+	  else let val NewChar = Char.ord(CharVector.sub(!yyb,l))
+		val NewChar = if NewChar<128 then NewChar else 128
+		val NewState = Char.ord(CharVector.sub(trans,NewChar))
+		in if NewState=0 then action(l,NewAcceptingLeaves)
+		else scan(NewState,NewAcceptingLeaves,l+1,i0)
+	end
+	end
+(*
+	val start= if substring(!yyb,!yybufpos-1,1)="\n"
+then !yybegin+1 else !yybegin
+*)
+	in scan(!yybegin (* start *),nil,!yybufpos,!yybufpos)
+    end
+in continue end
+  in lex
+  end
+end
+functor TPTPLrValsFun(structure Token : TOKEN)
+ : sig structure ParserData : PARSER_DATA
+       structure Tokens : TPTP_TOKENS
+   end
+ = 
+struct
+structure ParserData=
+struct
+structure Header = 
+struct
+open TPTP_Syntax
+
+exception UNRECOGNISED_SYMBOL of string * string
+
+exception UNRECOGNISED_ROLE of string
+fun classify_role role =
+  case role of
+    "axiom" => Role_Axiom
+  | "hypothesis" => Role_Hypothesis
+  | "definition" => Role_Definition
+  | "assumption" => Role_Assumption
+  | "lemma" => Role_Lemma
+  | "theorem" => Role_Theorem
+  | "conjecture" => Role_Conjecture
+  | "negated_conjecture" => Role_Negated_Conjecture
+  | "plain" => Role_Plain
+  | "fi_domain" => Role_Fi_Domain
+  | "fi_functors" => Role_Fi_Functors
+  | "fi_predicates" => Role_Fi_Predicates
+  | "type" => Role_Type
+  | "unknown" => Role_Unknown
+  | thing => raise (UNRECOGNISED_ROLE thing)
+
+
+end
+structure LrTable = Token.LrTable
+structure Token = Token
+local open LrTable in 
+val table=let val actionRows =
+"\
+\\001\000\001\000\032\002\004\000\155\002\005\000\032\002\006\000\032\002\
+\\010\000\032\002\011\000\032\002\012\000\032\002\016\000\212\000\
+\\019\000\032\002\020\000\032\002\021\000\032\002\022\000\032\002\
+\\027\000\032\002\037\000\032\002\000\000\
+\\001\000\001\000\044\002\004\000\154\002\005\000\044\002\006\000\044\002\
+\\010\000\044\002\011\000\044\002\012\000\044\002\016\000\217\000\
+\\019\000\044\002\020\000\044\002\021\000\044\002\022\000\044\002\
+\\027\000\044\002\037\000\044\002\000\000\
+\\001\000\001\000\054\002\005\000\054\002\006\000\049\002\010\000\054\002\
+\\011\000\054\002\012\000\054\002\019\000\054\002\020\000\049\002\
+\\021\000\054\002\022\000\054\002\026\000\054\002\027\000\054\002\
+\\037\000\054\002\000\000\
+\\001\000\001\000\061\002\005\000\061\002\006\000\039\002\010\000\061\002\
+\\011\000\061\002\012\000\061\002\019\000\061\002\020\000\039\002\
+\\021\000\061\002\022\000\061\002\026\000\061\002\027\000\061\002\
+\\037\000\061\002\000\000\
+\\001\000\001\000\064\002\005\000\064\002\006\000\047\002\010\000\064\002\
+\\011\000\064\002\012\000\064\002\019\000\064\002\020\000\047\002\
+\\021\000\064\002\022\000\064\002\026\000\064\002\027\000\064\002\
+\\037\000\064\002\000\000\
+\\001\000\001\000\170\002\005\000\170\002\006\000\052\002\010\000\170\002\
+\\011\000\170\002\012\000\170\002\019\000\170\002\020\000\052\002\
+\\021\000\170\002\022\000\170\002\026\000\170\002\027\000\170\002\
+\\037\000\170\002\000\000\
+\\001\000\001\000\225\002\002\000\225\002\004\000\213\002\005\000\225\002\
+\\006\000\225\002\008\000\225\002\009\000\225\002\010\000\225\002\
+\\011\000\225\002\012\000\225\002\019\000\225\002\020\000\225\002\
+\\021\000\225\002\022\000\225\002\026\000\225\002\027\000\225\002\
+\\037\000\225\002\059\000\225\002\060\000\225\002\000\000\
+\\001\000\001\000\228\002\002\000\228\002\004\000\214\002\005\000\228\002\
+\\006\000\228\002\008\000\228\002\009\000\228\002\010\000\228\002\
+\\011\000\228\002\012\000\228\002\019\000\228\002\020\000\228\002\
+\\021\000\228\002\022\000\228\002\026\000\228\002\027\000\228\002\
+\\037\000\228\002\059\000\228\002\060\000\228\002\000\000\
+\\001\000\001\000\206\000\003\000\205\000\006\000\204\000\007\000\119\000\
+\\008\000\203\000\010\000\202\000\011\000\201\000\012\000\200\000\
+\\013\000\035\000\015\000\199\000\016\000\198\000\019\000\197\000\
+\\020\000\196\000\021\000\195\000\022\000\194\000\025\000\116\000\
+\\028\000\115\000\037\000\193\000\044\000\096\000\045\000\095\000\
+\\046\000\034\000\047\000\033\000\049\000\032\000\050\000\094\000\
+\\051\000\031\000\053\000\093\000\055\000\192\000\056\000\191\000\
+\\057\000\190\000\058\000\189\000\062\000\188\000\063\000\187\000\
+\\064\000\092\000\065\000\091\000\068\000\030\000\069\000\029\000\
+\\070\000\028\000\071\000\027\000\072\000\186\000\073\000\090\000\000\000\
+\\001\000\001\000\206\000\003\000\205\000\006\000\204\000\007\000\119\000\
+\\008\000\203\000\010\000\202\000\011\000\201\000\012\000\200\000\
+\\013\000\035\000\016\000\024\001\019\000\197\000\020\000\196\000\
+\\021\000\195\000\022\000\194\000\025\000\116\000\026\000\023\001\
+\\028\000\115\000\037\000\193\000\044\000\096\000\045\000\095\000\
+\\046\000\034\000\047\000\033\000\049\000\032\000\050\000\094\000\
+\\051\000\031\000\053\000\093\000\055\000\192\000\056\000\191\000\
+\\057\000\190\000\058\000\189\000\062\000\188\000\063\000\187\000\
+\\064\000\092\000\065\000\091\000\068\000\030\000\069\000\029\000\
+\\070\000\028\000\071\000\027\000\072\000\186\000\073\000\090\000\000\000\
+\\001\000\001\000\206\000\003\000\205\000\006\000\204\000\007\000\119\000\
+\\008\000\203\000\010\000\202\000\011\000\201\000\012\000\200\000\
+\\013\000\035\000\016\000\024\001\019\000\197\000\020\000\196\000\
+\\021\000\195\000\022\000\194\000\025\000\116\000\028\000\115\000\
+\\037\000\193\000\044\000\096\000\045\000\095\000\046\000\034\000\
+\\047\000\033\000\049\000\032\000\050\000\094\000\051\000\031\000\
+\\053\000\093\000\055\000\192\000\056\000\191\000\057\000\190\000\
+\\058\000\189\000\062\000\188\000\063\000\187\000\064\000\092\000\
+\\065\000\091\000\068\000\030\000\069\000\029\000\070\000\028\000\
+\\071\000\027\000\072\000\186\000\073\000\090\000\000\000\
+\\001\000\001\000\206\000\003\000\205\000\006\000\204\000\007\000\119\000\
+\\008\000\203\000\010\000\202\000\011\000\201\000\012\000\200\000\
+\\013\000\035\000\016\000\097\001\019\000\197\000\020\000\196\000\
+\\021\000\195\000\022\000\194\000\025\000\116\000\028\000\115\000\
+\\037\000\193\000\044\000\096\000\045\000\095\000\046\000\034\000\
+\\047\000\033\000\049\000\032\000\050\000\094\000\051\000\031\000\
+\\053\000\093\000\055\000\192\000\056\000\191\000\057\000\190\000\
+\\058\000\189\000\062\000\188\000\063\000\187\000\064\000\092\000\
+\\065\000\091\000\068\000\030\000\069\000\029\000\070\000\028\000\
+\\071\000\027\000\072\000\186\000\073\000\090\000\000\000\
+\\001\000\001\000\007\001\002\000\006\001\005\000\243\002\006\000\204\000\
+\\008\000\243\002\009\000\210\002\010\000\202\000\011\000\201\000\
+\\012\000\200\000\019\000\197\000\020\000\196\000\021\000\195\000\
+\\022\000\194\000\026\000\243\002\027\000\243\002\037\000\005\001\
+\\059\000\210\002\060\000\210\002\000\000\
+\\001\000\004\000\243\000\000\000\
+\\001\000\004\000\008\001\000\000\
+\\001\000\004\000\193\001\000\000\
+\\001\000\004\000\201\001\000\000\
+\\001\000\004\000\205\001\000\000\
+\\001\000\004\000\211\001\000\000\
+\\001\000\004\000\216\001\000\000\
+\\001\000\005\000\152\002\009\000\150\002\027\000\152\002\000\000\
+\\001\000\005\000\041\000\000\000\
+\\001\000\005\000\042\000\000\000\
+\\001\000\005\000\043\000\000\000\
+\\001\000\005\000\044\000\000\000\
+\\001\000\005\000\054\000\000\000\
+\\001\000\005\000\055\000\000\000\
+\\001\000\005\000\056\000\000\000\
+\\001\000\005\000\057\000\000\000\
+\\001\000\005\000\147\001\000\000\
+\\001\000\005\000\161\001\000\000\
+\\001\000\005\000\174\001\000\000\
+\\001\000\005\000\226\001\000\000\
+\\001\000\005\000\232\001\000\000\
+\\001\000\005\000\235\001\000\000\
+\\001\000\006\000\204\000\000\000\
+\\001\000\006\000\204\000\020\000\196\000\000\000\
+\\001\000\007\000\119\000\008\000\146\000\013\000\035\000\015\000\145\000\
+\\016\000\144\000\025\000\116\000\028\000\115\000\044\000\096\000\
+\\045\000\095\000\046\000\034\000\047\000\033\000\049\000\032\000\
+\\050\000\094\000\051\000\031\000\053\000\093\000\064\000\092\000\
+\\065\000\091\000\068\000\030\000\069\000\029\000\070\000\028\000\
+\\071\000\027\000\072\000\143\000\073\000\090\000\000\000\
+\\001\000\007\000\119\000\008\000\146\000\013\000\035\000\016\000\247\000\
+\\025\000\116\000\026\000\254\000\028\000\115\000\044\000\096\000\
+\\045\000\095\000\046\000\034\000\047\000\033\000\049\000\032\000\
+\\050\000\094\000\051\000\031\000\053\000\093\000\064\000\092\000\
+\\065\000\091\000\068\000\030\000\069\000\029\000\070\000\028\000\
+\\071\000\027\000\072\000\143\000\073\000\090\000\000\000\
+\\001\000\007\000\119\000\008\000\146\000\013\000\035\000\016\000\247\000\
+\\025\000\116\000\028\000\115\000\044\000\096\000\045\000\095\000\
+\\046\000\034\000\047\000\033\000\049\000\032\000\050\000\094\000\
+\\051\000\031\000\053\000\093\000\064\000\092\000\065\000\091\000\
+\\068\000\030\000\069\000\029\000\070\000\028\000\071\000\027\000\
+\\072\000\143\000\073\000\090\000\000\000\
+\\001\000\007\000\119\000\013\000\035\000\015\000\118\000\016\000\117\000\
+\\025\000\116\000\028\000\115\000\044\000\096\000\045\000\095\000\
+\\046\000\034\000\047\000\033\000\049\000\032\000\050\000\094\000\
+\\051\000\031\000\053\000\093\000\064\000\092\000\065\000\091\000\
+\\068\000\030\000\069\000\029\000\070\000\028\000\071\000\027\000\
+\\073\000\090\000\000\000\
+\\001\000\007\000\119\000\013\000\035\000\016\000\231\000\025\000\116\000\
+\\026\000\236\000\028\000\115\000\044\000\096\000\045\000\095\000\
+\\046\000\034\000\047\000\033\000\049\000\032\000\050\000\094\000\
+\\051\000\031\000\053\000\093\000\064\000\092\000\065\000\091\000\
+\\068\000\030\000\069\000\029\000\070\000\028\000\071\000\027\000\
+\\073\000\090\000\000\000\
+\\001\000\007\000\119\000\013\000\035\000\016\000\231\000\025\000\116\000\
+\\028\000\115\000\044\000\096\000\045\000\095\000\046\000\034\000\
+\\047\000\033\000\049\000\032\000\050\000\094\000\051\000\031\000\
+\\053\000\093\000\064\000\092\000\065\000\091\000\068\000\030\000\
+\\069\000\029\000\070\000\028\000\071\000\027\000\073\000\090\000\000\000\
+\\001\000\008\000\166\001\067\000\165\001\000\000\
+\\001\000\008\000\176\001\000\000\
+\\001\000\009\000\151\002\027\000\145\002\060\000\145\002\000\000\
+\\001\000\009\000\011\001\059\000\010\001\060\000\009\001\000\000\
+\\001\000\009\000\153\001\000\000\
+\\001\000\013\000\035\000\015\000\042\001\026\000\142\001\039\000\041\001\
+\\040\000\040\001\041\000\039\001\042\000\038\001\043\000\037\001\
+\\044\000\096\000\045\000\095\000\046\000\034\000\047\000\033\000\
+\\049\000\032\000\050\000\094\000\051\000\031\000\053\000\036\001\
+\\068\000\030\000\069\000\029\000\070\000\028\000\071\000\027\000\000\000\
+\\001\000\013\000\035\000\015\000\042\001\039\000\041\001\040\000\040\001\
+\\041\000\039\001\042\000\038\001\043\000\037\001\044\000\096\000\
+\\045\000\095\000\046\000\034\000\047\000\033\000\049\000\032\000\
+\\050\000\094\000\051\000\031\000\053\000\036\001\068\000\030\000\
+\\069\000\029\000\070\000\028\000\071\000\027\000\000\000\
+\\001\000\013\000\035\000\016\000\098\000\028\000\097\000\044\000\096\000\
+\\045\000\095\000\046\000\034\000\047\000\033\000\049\000\032\000\
+\\050\000\094\000\051\000\031\000\053\000\093\000\064\000\092\000\
+\\065\000\091\000\068\000\030\000\069\000\029\000\070\000\028\000\
+\\071\000\027\000\073\000\090\000\000\000\
+\\001\000\013\000\035\000\016\000\078\001\049\000\032\000\051\000\031\000\
+\\064\000\077\001\068\000\030\000\069\000\029\000\070\000\028\000\
+\\071\000\027\000\000\000\
+\\001\000\013\000\035\000\016\000\157\001\049\000\032\000\051\000\031\000\
+\\064\000\077\001\068\000\030\000\069\000\029\000\070\000\028\000\
+\\071\000\027\000\000\000\
+\\001\000\013\000\035\000\028\000\097\000\044\000\096\000\045\000\095\000\
+\\046\000\034\000\047\000\033\000\049\000\032\000\050\000\094\000\
+\\051\000\031\000\053\000\093\000\064\000\092\000\065\000\091\000\
+\\068\000\030\000\069\000\029\000\070\000\028\000\071\000\027\000\
+\\073\000\090\000\000\000\
+\\001\000\013\000\035\000\044\000\096\000\045\000\095\000\046\000\034\000\
+\\047\000\033\000\049\000\032\000\050\000\094\000\051\000\031\000\
+\\053\000\093\000\064\000\092\000\065\000\091\000\068\000\030\000\
+\\069\000\029\000\070\000\028\000\071\000\027\000\073\000\090\000\000\000\
+\\001\000\013\000\035\000\046\000\034\000\047\000\033\000\049\000\032\000\
+\\051\000\031\000\068\000\030\000\069\000\029\000\070\000\028\000\
+\\071\000\027\000\000\000\
+\\001\000\013\000\035\000\049\000\032\000\051\000\031\000\064\000\077\001\
+\\068\000\030\000\069\000\029\000\070\000\028\000\071\000\027\000\000\000\
+\\001\000\013\000\035\000\049\000\032\000\051\000\031\000\068\000\030\000\
+\\069\000\029\000\070\000\028\000\071\000\027\000\000\000\
+\\001\000\015\000\053\000\000\000\
+\\001\000\015\000\118\000\000\000\
+\\001\000\015\000\145\000\000\000\
+\\001\000\015\000\199\000\000\000\
+\\001\000\015\000\229\000\000\000\
+\\001\000\015\000\245\000\000\000\
+\\001\000\015\000\255\000\000\000\
+\\001\000\015\000\015\001\000\000\
+\\001\000\015\000\025\001\000\000\
+\\001\000\015\000\042\001\000\000\
+\\001\000\016\000\018\000\000\000\
+\\001\000\016\000\019\000\000\000\
+\\001\000\016\000\020\000\000\000\
+\\001\000\016\000\021\000\000\000\
+\\001\000\016\000\023\000\000\000\
+\\001\000\016\000\218\000\000\000\
+\\001\000\016\000\248\000\000\000\
+\\001\000\016\000\018\001\000\000\
+\\001\000\016\000\093\001\050\000\094\000\000\000\
+\\001\000\016\000\129\001\050\000\094\000\000\000\
+\\001\000\016\000\135\001\000\000\
+\\001\000\016\000\136\001\000\000\
+\\001\000\016\000\137\001\000\000\
+\\001\000\016\000\138\001\000\000\
+\\001\000\016\000\139\001\000\000\
+\\001\000\023\000\058\000\000\000\
+\\001\000\023\000\130\001\000\000\
+\\001\000\023\000\148\001\000\000\
+\\001\000\023\000\152\001\000\000\
+\\001\000\023\000\168\001\000\000\
+\\001\000\026\000\207\000\000\000\
+\\001\000\026\000\064\001\000\000\
+\\001\000\026\000\089\001\000\000\
+\\001\000\026\000\125\001\000\000\
+\\001\000\026\000\149\001\000\000\
+\\001\000\026\000\158\001\000\000\
+\\001\000\026\000\163\001\000\000\
+\\001\000\026\000\170\001\000\000\
+\\001\000\026\000\177\001\000\000\
+\\001\000\026\000\190\001\000\000\
+\\001\000\027\000\052\000\000\000\
+\\001\000\027\000\027\001\000\000\
+\\001\000\027\000\051\001\037\000\211\000\000\000\
+\\001\000\027\000\052\001\000\000\
+\\001\000\027\000\061\001\000\000\
+\\001\000\027\000\062\001\000\000\
+\\001\000\027\000\065\001\000\000\
+\\001\000\027\000\085\001\000\000\
+\\001\000\027\000\086\001\000\000\
+\\001\000\027\000\087\001\000\000\
+\\001\000\027\000\094\001\000\000\
+\\001\000\027\000\122\001\000\000\
+\\001\000\027\000\123\001\000\000\
+\\001\000\027\000\143\001\000\000\
+\\001\000\027\000\145\001\000\000\
+\\001\000\027\000\146\001\000\000\
+\\001\000\027\000\173\001\000\000\
+\\001\000\027\000\197\001\000\000\
+\\001\000\027\000\199\001\060\000\198\001\000\000\
+\\001\000\027\000\209\001\000\000\
+\\001\000\027\000\210\001\000\000\
+\\001\000\027\000\218\001\000\000\
+\\001\000\027\000\219\001\000\000\
+\\001\000\027\000\220\001\000\000\
+\\001\000\027\000\221\001\000\000\
+\\001\000\027\000\222\001\000\000\
+\\001\000\027\000\223\001\000\000\
+\\001\000\027\000\224\001\000\000\
+\\001\000\027\000\230\001\060\000\198\001\000\000\
+\\001\000\027\000\240\001\000\000\
+\\001\000\027\000\241\001\000\000\
+\\001\000\027\000\242\001\000\000\
+\\001\000\038\000\000\000\000\000\
+\\001\000\049\000\040\000\000\000\
+\\001\000\050\000\094\000\000\000\
+\\001\000\051\000\048\000\000\000\
+\\001\000\061\000\228\000\000\000\
+\\001\000\061\000\244\000\000\000\
+\\001\000\061\000\014\001\000\000\
+\\244\001\000\000\
+\\245\001\005\000\210\000\000\000\
+\\246\001\000\000\
+\\247\001\005\000\134\001\000\000\
+\\248\001\000\000\
+\\249\001\000\000\
+\\250\001\000\000\
+\\251\001\000\000\
+\\252\001\005\000\189\001\000\000\
+\\253\001\004\000\131\001\000\000\
+\\254\001\000\000\
+\\255\001\000\000\
+\\000\002\000\000\
+\\001\002\000\000\
+\\002\002\000\000\
+\\003\002\000\000\
+\\004\002\000\000\
+\\005\002\000\000\
+\\006\002\000\000\
+\\007\002\000\000\
+\\008\002\000\000\
+\\009\002\016\000\132\001\000\000\
+\\010\002\000\000\
+\\011\002\000\000\
+\\012\002\000\000\
+\\013\002\000\000\
+\\014\002\000\000\
+\\015\002\000\000\
+\\016\002\000\000\
+\\017\002\000\000\
+\\018\002\000\000\
+\\019\002\000\000\
+\\020\002\000\000\
+\\021\002\000\000\
+\\022\002\000\000\
+\\023\002\000\000\
+\\024\002\000\000\
+\\025\002\000\000\
+\\027\002\000\000\
+\\028\002\000\000\
+\\029\002\005\000\144\001\000\000\
+\\030\002\000\000\
+\\031\002\000\000\
+\\032\002\016\000\212\000\000\000\
+\\033\002\000\000\
+\\034\002\000\000\
+\\035\002\000\000\
+\\036\002\016\000\213\000\000\000\
+\\037\002\000\000\
+\\038\002\000\000\
+\\039\002\000\000\
+\\040\002\000\000\
+\\041\002\000\000\
+\\042\002\000\000\
+\\043\002\000\000\
+\\044\002\000\000\
+\\044\002\016\000\217\000\000\000\
+\\045\002\000\000\
+\\045\002\066\000\017\001\000\000\
+\\046\002\000\000\
+\\047\002\000\000\
+\\048\002\000\000\
+\\049\002\000\000\
+\\050\002\000\000\
+\\051\002\000\000\
+\\052\002\000\000\
+\\053\002\000\000\
+\\055\002\000\000\
+\\056\002\000\000\
+\\057\002\000\000\
+\\058\002\000\000\
+\\062\002\000\000\
+\\063\002\000\000\
+\\065\002\000\000\
+\\066\002\000\000\
+\\067\002\000\000\
+\\068\002\000\000\
+\\069\002\000\000\
+\\070\002\000\000\
+\\071\002\000\000\
+\\072\002\000\000\
+\\073\002\000\000\
+\\074\002\000\000\
+\\075\002\000\000\
+\\076\002\000\000\
+\\077\002\000\000\
+\\078\002\000\000\
+\\079\002\000\000\
+\\080\002\000\000\
+\\081\002\000\000\
+\\082\002\000\000\
+\\083\002\000\000\
+\\084\002\000\000\
+\\085\002\000\000\
+\\086\002\000\000\
+\\087\002\000\000\
+\\088\002\000\000\
+\\089\002\000\000\
+\\090\002\000\000\
+\\091\002\000\000\
+\\092\002\000\000\
+\\093\002\016\000\016\001\000\000\
+\\094\002\000\000\
+\\095\002\000\000\
+\\096\002\000\000\
+\\097\002\000\000\
+\\098\002\000\000\
+\\099\002\000\000\
+\\100\002\037\000\211\000\000\000\
+\\101\002\005\000\063\001\000\000\
+\\102\002\000\000\
+\\103\002\000\000\
+\\104\002\000\000\
+\\105\002\000\000\
+\\106\002\000\000\
+\\107\002\000\000\
+\\108\002\000\000\
+\\109\002\000\000\
+\\110\002\005\000\150\001\000\000\
+\\111\002\000\000\
+\\112\002\000\000\
+\\113\002\000\000\
+\\114\002\000\000\
+\\115\002\000\000\
+\\116\002\000\000\
+\\117\002\000\000\
+\\118\002\000\000\
+\\119\002\000\000\
+\\120\002\000\000\
+\\121\002\037\000\223\000\000\000\
+\\122\002\001\000\224\000\000\000\
+\\123\002\000\000\
+\\124\002\000\000\
+\\125\002\000\000\
+\\126\002\000\000\
+\\127\002\001\000\227\000\010\000\202\000\011\000\201\000\012\000\200\000\
+\\019\000\197\000\021\000\195\000\022\000\194\000\037\000\226\000\000\000\
+\\128\002\000\000\
+\\129\002\000\000\
+\\130\002\000\000\
+\\131\002\000\000\
+\\132\002\000\000\
+\\133\002\005\000\088\001\000\000\
+\\134\002\000\000\
+\\135\002\000\000\
+\\136\002\000\000\
+\\137\002\000\000\
+\\138\002\000\000\
+\\139\002\000\000\
+\\140\002\005\000\164\001\000\000\
+\\141\002\000\000\
+\\142\002\000\000\
+\\143\002\000\000\
+\\144\002\000\000\
+\\146\002\000\000\
+\\147\002\000\000\
+\\148\002\000\000\
+\\149\002\000\000\
+\\150\002\060\000\196\001\000\000\
+\\151\002\000\000\
+\\153\002\000\000\
+\\156\002\000\000\
+\\157\002\000\000\
+\\158\002\000\000\
+\\159\002\000\000\
+\\160\002\000\000\
+\\161\002\000\000\
+\\162\002\004\000\160\001\000\000\
+\\163\002\005\000\159\001\000\000\
+\\164\002\000\000\
+\\165\002\000\000\
+\\166\002\000\000\
+\\167\002\000\000\
+\\168\002\000\000\
+\\169\002\000\000\
+\\171\002\000\000\
+\\172\002\000\000\
+\\173\002\000\000\
+\\174\002\000\000\
+\\175\002\000\000\
+\\176\002\000\000\
+\\177\002\037\000\238\000\000\000\
+\\178\002\001\000\239\000\000\000\
+\\179\002\000\000\
+\\180\002\000\000\
+\\181\002\000\000\
+\\182\002\000\000\
+\\183\002\001\000\242\000\010\000\202\000\011\000\201\000\012\000\200\000\
+\\019\000\197\000\021\000\195\000\022\000\194\000\037\000\241\000\000\000\
+\\184\002\000\000\
+\\185\002\000\000\
+\\186\002\000\000\
+\\187\002\000\000\
+\\188\002\000\000\
+\\189\002\005\000\124\001\000\000\
+\\190\002\000\000\
+\\191\002\000\000\
+\\192\002\000\000\
+\\193\002\000\000\
+\\194\002\000\000\
+\\195\002\000\000\
+\\196\002\005\000\178\001\000\000\
+\\197\002\000\000\
+\\198\002\000\000\
+\\199\002\000\000\
+\\200\002\000\000\
+\\201\002\000\000\
+\\202\002\000\000\
+\\203\002\000\000\
+\\204\002\000\000\
+\\205\002\009\000\011\001\000\000\
+\\206\002\000\000\
+\\207\002\000\000\
+\\208\002\060\000\012\001\000\000\
+\\209\002\059\000\013\001\000\000\
+\\210\002\000\000\
+\\211\002\000\000\
+\\212\002\000\000\
+\\215\002\000\000\
+\\216\002\000\000\
+\\217\002\000\000\
+\\218\002\000\000\
+\\219\002\004\000\172\001\000\000\
+\\220\002\005\000\171\001\000\000\
+\\221\002\000\000\
+\\222\002\000\000\
+\\223\002\000\000\
+\\224\002\000\000\
+\\225\002\000\000\
+\\226\002\000\000\
+\\227\002\000\000\
+\\228\002\000\000\
+\\229\002\000\000\
+\\230\002\000\000\
+\\231\002\000\000\
+\\232\002\000\000\
+\\233\002\000\000\
+\\234\002\000\000\
+\\235\002\037\000\001\001\000\000\
+\\236\002\001\000\002\001\000\000\
+\\237\002\002\000\003\001\000\000\
+\\238\002\000\000\
+\\239\002\000\000\
+\\240\002\000\000\
+\\241\002\000\000\
+\\242\002\000\000\
+\\244\002\000\000\
+\\245\002\000\000\
+\\246\002\000\000\
+\\247\002\000\000\
+\\248\002\000\000\
+\\249\002\000\000\
+\\250\002\000\000\
+\\251\002\000\000\
+\\252\002\000\000\
+\\253\002\000\000\
+\\254\002\000\000\
+\\255\002\000\000\
+\\000\003\000\000\
+\\001\003\000\000\
+\\002\003\000\000\
+\\003\003\005\000\046\000\000\000\
+\\004\003\000\000\
+\\005\003\005\000\208\000\000\000\
+\\006\003\000\000\
+\\007\003\000\000\
+\\008\003\000\000\
+\\009\003\000\000\
+\\010\003\000\000\
+\\011\003\000\000\
+\\012\003\013\000\016\000\052\000\015\000\068\000\014\000\069\000\013\000\
+\\070\000\012\000\071\000\011\000\000\000\
+\\013\003\000\000\
+\"
+val actionRowNumbers =
+"\149\001\150\001\149\001\146\001\
+\\145\001\137\001\136\001\135\001\
+\\134\001\068\000\069\000\070\000\
+\\071\000\149\001\072\000\147\001\
+\\055\000\055\000\055\000\055\000\
+\\148\001\131\000\144\001\143\001\
+\\021\000\154\000\153\000\152\000\
+\\151\000\149\000\150\000\167\000\
+\\168\000\155\000\022\000\023\000\
+\\024\000\140\001\169\000\133\000\
+\\133\000\133\000\133\000\098\000\
+\\058\000\025\000\129\001\026\000\
+\\027\000\028\000\083\000\055\000\
+\\050\000\040\000\037\000\008\000\
+\\138\001\088\000\142\001\138\000\
+\\245\000\242\000\241\000\239\000\
+\\210\000\211\000\208\000\209\000\
+\\212\000\203\000\201\000\004\000\
+\\194\000\198\000\190\000\191\000\
+\\003\000\185\000\002\000\181\000\
+\\180\000\184\000\036\000\193\000\
+\\164\000\188\000\202\000\176\000\
+\\073\000\179\000\183\000\189\000\
+\\156\000\166\000\165\000\054\000\
+\\053\000\138\000\017\001\015\001\
+\\013\001\014\001\010\001\011\001\
+\\016\001\002\001\003\001\018\001\
+\\134\000\254\000\062\000\042\000\
+\\004\001\252\000\222\000\040\000\
+\\041\000\221\000\138\000\068\001\
+\\066\001\064\001\065\001\061\001\
+\\062\001\067\001\051\001\052\001\
+\\069\001\013\000\054\001\055\001\
+\\070\001\135\000\044\001\063\000\
+\\039\000\053\001\000\000\001\000\
+\\005\000\074\000\037\000\038\000\
+\\064\000\138\000\127\001\124\001\
+\\121\001\122\001\117\001\118\001\
+\\119\001\012\000\105\001\106\001\
+\\125\001\014\000\126\001\046\000\
+\\123\001\091\001\092\001\093\001\
+\\006\000\108\001\109\001\128\001\
+\\136\000\084\001\065\000\236\000\
+\\238\000\229\000\228\000\237\000\
+\\223\000\227\000\226\000\197\000\
+\\195\000\187\000\199\000\083\001\
+\\075\000\231\000\232\000\225\000\
+\\224\000\234\000\233\000\213\000\
+\\219\000\218\000\205\000\220\000\
+\\008\000\009\000\216\000\215\000\
+\\217\000\066\000\204\000\230\000\
+\\214\000\139\001\055\000\099\000\
+\\049\000\053\000\054\000\054\000\
+\\054\000\054\000\206\000\054\000\
+\\039\000\240\000\035\000\100\000\
+\\101\000\042\000\042\000\042\000\
+\\042\000\042\000\059\000\132\000\
+\\253\000\042\000\102\000\103\000\
+\\246\000\089\000\248\000\104\000\
+\\039\000\039\000\039\000\039\000\
+\\039\000\051\000\060\000\132\000\
+\\043\001\039\000\039\000\105\000\
+\\106\000\107\000\022\001\090\000\
+\\019\001\076\000\108\000\011\000\
+\\011\000\011\000\011\000\011\000\
+\\011\000\011\000\010\000\011\000\
+\\011\000\011\000\011\000\011\000\
+\\061\000\132\000\010\000\057\000\
+\\010\000\109\000\110\000\073\001\
+\\091\000\071\001\010\000\077\000\
+\\141\001\084\000\159\000\163\000\
+\\161\000\160\000\146\000\158\000\
+\\140\000\148\000\162\000\078\000\
+\\079\000\080\000\081\000\082\000\
+\\048\000\243\000\111\000\177\000\
+\\112\000\207\000\235\000\113\000\
+\\029\000\244\000\085\000\009\001\
+\\007\001\012\001\008\001\006\001\
+\\250\000\092\000\255\000\005\001\
+\\251\000\042\000\249\000\086\000\
+\\060\001\058\001\063\001\059\001\
+\\057\001\041\001\047\000\020\000\
+\\040\001\037\001\036\001\175\000\
+\\052\000\023\001\093\000\048\001\
+\\046\001\047\001\030\000\056\001\
+\\042\001\024\001\039\000\020\001\
+\\094\000\029\001\043\000\076\000\
+\\087\000\116\001\107\001\010\000\
+\\114\001\112\001\120\001\115\001\
+\\111\001\113\001\095\001\097\001\
+\\094\001\087\001\085\001\089\001\
+\\090\001\088\001\086\001\075\001\
+\\095\000\102\001\100\001\101\001\
+\\114\000\096\001\192\000\031\000\
+\\007\000\076\001\010\000\072\001\
+\\044\000\096\000\080\001\077\000\
+\\133\001\049\000\049\000\137\000\
+\\067\000\037\000\054\000\050\000\
+\\040\000\008\000\145\000\097\000\
+\\143\000\182\000\054\000\186\000\
+\\196\000\054\000\132\001\015\000\
+\\132\000\247\000\131\001\056\000\
+\\038\001\115\000\116\000\052\000\
+\\016\000\132\000\056\000\039\000\
+\\021\001\017\000\076\000\054\000\
+\\039\000\117\000\130\001\118\000\
+\\018\000\132\000\010\000\098\001\
+\\010\000\074\001\010\000\019\000\
+\\077\000\119\000\147\000\120\000\
+\\139\000\141\000\121\000\122\000\
+\\123\000\124\000\125\000\049\000\
+\\142\000\178\000\032\000\042\000\
+\\000\001\034\001\056\000\035\001\
+\\056\000\039\001\126\000\039\000\
+\\049\001\045\001\033\000\039\000\
+\\030\001\027\001\026\001\028\001\
+\\110\001\011\000\103\001\099\001\
+\\034\000\078\001\011\000\081\001\
+\\079\001\157\000\171\000\174\000\
+\\173\000\172\000\170\000\144\000\
+\\054\000\001\001\032\001\033\001\
+\\045\000\050\001\039\000\031\001\
+\\104\001\010\000\082\001\127\000\
+\\128\000\129\000\200\000\025\001\
+\\077\001\130\000"
+val gotoT =
+"\
+\\133\000\008\000\134\000\007\000\135\000\006\000\136\000\005\000\
+\\137\000\004\000\138\000\003\000\139\000\002\000\140\000\001\000\
+\\141\000\241\001\000\000\
+\\000\000\
+\\133\000\008\000\134\000\007\000\135\000\006\000\136\000\005\000\
+\\137\000\004\000\138\000\003\000\139\000\002\000\140\000\015\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\133\000\008\000\134\000\007\000\135\000\006\000\136\000\005\000\
+\\137\000\004\000\138\000\003\000\139\000\002\000\140\000\020\000\000\000\
+\\000\000\
+\\000\000\
+\\002\000\024\000\009\000\023\000\014\000\022\000\000\000\
+\\002\000\034\000\009\000\023\000\014\000\022\000\000\000\
+\\002\000\035\000\009\000\023\000\014\000\022\000\000\000\
+\\002\000\036\000\009\000\023\000\014\000\022\000\000\000\
+\\000\000\
+\\018\000\037\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\004\000\043\000\000\000\
+\\000\000\
+\\132\000\045\000\000\000\
+\\132\000\047\000\000\000\
+\\132\000\048\000\000\000\
+\\132\000\049\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\002\000\058\000\003\000\057\000\009\000\023\000\014\000\022\000\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\063\000\055\000\062\000\057\000\061\000\058\000\060\000\
+\\059\000\059\000\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\061\000\108\000\062\000\107\000\063\000\106\000\065\000\105\000\
+\\066\000\104\000\067\000\103\000\068\000\102\000\069\000\101\000\
+\\070\000\100\000\071\000\099\000\072\000\098\000\073\000\097\000\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\139\000\020\000\082\000\022\000\081\000\023\000\138\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\074\000\133\000\076\000\132\000\077\000\131\000\080\000\130\000\
+\\086\000\129\000\087\000\128\000\088\000\127\000\092\000\126\000\
+\\093\000\125\000\094\000\124\000\095\000\123\000\096\000\122\000\
+\\097\000\121\000\098\000\120\000\099\000\119\000\100\000\118\000\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\101\000\168\000\103\000\167\000\104\000\166\000\
+\\107\000\165\000\108\000\164\000\109\000\163\000\110\000\162\000\
+\\111\000\161\000\112\000\160\000\113\000\159\000\115\000\158\000\
+\\116\000\157\000\117\000\156\000\118\000\155\000\122\000\154\000\
+\\123\000\153\000\124\000\152\000\125\000\151\000\126\000\150\000\
+\\127\000\149\000\128\000\148\000\129\000\147\000\130\000\146\000\
+\\131\000\145\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\001\000\207\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\036\000\214\000\037\000\213\000\038\000\212\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\218\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\217\000\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\063\000\055\000\062\000\057\000\061\000\058\000\219\000\000\000\
+\\001\000\220\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\050\000\223\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\063\000\106\000\065\000\105\000\066\000\228\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\061\000\108\000\062\000\231\000\063\000\106\000\065\000\105\000\
+\\066\000\104\000\067\000\103\000\068\000\102\000\069\000\101\000\
+\\070\000\100\000\071\000\099\000\072\000\230\000\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\060\000\233\000\063\000\106\000\065\000\105\000\066\000\104\000\
+\\067\000\103\000\068\000\102\000\069\000\101\000\070\000\100\000\
+\\071\000\099\000\072\000\232\000\000\000\
+\\000\000\
+\\001\000\235\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\050\000\238\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\244\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\139\000\020\000\082\000\022\000\081\000\023\000\138\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\074\000\133\000\076\000\249\000\077\000\131\000\080\000\130\000\
+\\086\000\129\000\087\000\248\000\088\000\127\000\092\000\126\000\
+\\093\000\125\000\094\000\124\000\095\000\123\000\096\000\122\000\
+\\097\000\121\000\098\000\120\000\099\000\247\000\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\075\000\251\000\077\000\131\000\080\000\130\000\088\000\127\000\
+\\092\000\126\000\093\000\125\000\094\000\124\000\095\000\123\000\
+\\096\000\122\000\097\000\121\000\098\000\120\000\099\000\250\000\000\000\
+\\000\000\
+\\001\000\254\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\036\000\178\000\037\000\177\000\050\000\174\000\053\000\002\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\101\000\168\000\103\000\018\001\104\000\166\000\
+\\107\000\165\000\108\000\164\000\109\000\163\000\110\000\162\000\
+\\111\000\161\000\112\000\160\000\113\000\159\000\115\000\158\000\
+\\116\000\157\000\117\000\156\000\118\000\155\000\122\000\154\000\
+\\123\000\153\000\124\000\152\000\125\000\151\000\126\000\150\000\
+\\127\000\149\000\128\000\148\000\129\000\147\000\130\000\017\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\102\000\020\001\104\000\166\000\107\000\165\000\
+\\108\000\164\000\109\000\163\000\110\000\162\000\111\000\161\000\
+\\112\000\160\000\113\000\159\000\115\000\158\000\116\000\157\000\
+\\117\000\156\000\118\000\155\000\122\000\154\000\123\000\153\000\
+\\124\000\152\000\125\000\151\000\126\000\150\000\127\000\149\000\
+\\128\000\148\000\129\000\147\000\130\000\019\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\002\000\058\000\003\000\024\001\009\000\023\000\014\000\022\000\000\000\
+\\000\000\
+\\006\000\033\001\008\000\032\001\009\000\031\001\010\000\030\001\
+\\011\000\029\001\012\000\028\001\013\000\027\001\014\000\084\000\
+\\016\000\026\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\063\000\055\000\062\000\057\000\041\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\043\001\021\000\042\001\022\000\081\000\
+\\023\000\080\000\024\000\079\000\025\000\182\000\026\000\077\000\
+\\027\000\181\000\028\000\075\000\029\000\074\000\030\000\073\000\
+\\031\000\072\000\032\000\179\000\033\000\070\000\034\000\069\000\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\043\001\021\000\044\001\022\000\081\000\
+\\023\000\080\000\024\000\079\000\025\000\182\000\026\000\077\000\
+\\027\000\181\000\028\000\075\000\029\000\074\000\030\000\073\000\
+\\031\000\072\000\032\000\179\000\033\000\070\000\034\000\069\000\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\045\001\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\046\001\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\043\001\021\000\047\001\022\000\081\000\
+\\023\000\080\000\024\000\079\000\025\000\182\000\026\000\077\000\
+\\027\000\181\000\028\000\075\000\029\000\074\000\030\000\073\000\
+\\031\000\072\000\032\000\179\000\033\000\070\000\034\000\069\000\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\125\000\094\000\124\000\095\000\123\000\096\000\122\000\
+\\097\000\121\000\098\000\120\000\099\000\048\001\000\000\
+\\000\000\
+\\036\000\214\000\038\000\212\000\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\063\000\106\000\065\000\105\000\066\000\051\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\063\000\106\000\065\000\105\000\066\000\052\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\063\000\106\000\065\000\105\000\066\000\053\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\063\000\106\000\065\000\105\000\066\000\054\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\063\000\106\000\065\000\105\000\066\000\055\001\000\000\
+\\061\000\056\001\000\000\
+\\011\000\058\001\064\000\057\001\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\063\000\106\000\065\000\105\000\066\000\104\000\067\000\103\000\
+\\068\000\102\000\069\000\101\000\070\000\100\000\071\000\099\000\
+\\072\000\230\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\064\001\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\065\001\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\066\001\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\067\001\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\068\001\000\000\
+\\009\000\074\001\047\000\073\001\082\000\072\001\083\000\071\001\
+\\084\000\070\001\085\000\069\001\000\000\
+\\074\000\077\001\000\000\
+\\011\000\081\001\089\000\080\001\090\000\079\001\091\000\078\001\000\000\
+\\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\125\000\094\000\124\000\095\000\123\000\096\000\122\000\
+\\097\000\121\000\098\000\120\000\099\000\247\000\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\125\000\094\000\124\000\095\000\123\000\096\000\122\000\
+\\097\000\121\000\098\000\120\000\099\000\082\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\011\000\090\001\078\000\089\001\079\000\088\001\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\118\000\155\000\122\000\154\000\123\000\093\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\118\000\155\000\122\000\154\000\123\000\096\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\118\000\155\000\122\000\154\000\123\000\097\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\118\000\155\000\122\000\154\000\123\000\098\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\118\000\155\000\122\000\154\000\123\000\099\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\118\000\155\000\122\000\154\000\123\000\100\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\118\000\155\000\122\000\154\000\123\000\101\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\164\000\
+\\109\000\163\000\110\000\162\000\111\000\161\000\112\000\160\000\
+\\113\000\159\000\114\000\103\001\115\000\158\000\116\000\157\000\
+\\117\000\156\000\118\000\155\000\122\000\154\000\123\000\153\000\
+\\124\000\152\000\125\000\151\000\126\000\150\000\127\000\149\000\
+\\128\000\148\000\129\000\147\000\130\000\102\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\113\000\105\001\118\000\155\000\122\000\154\000\123\000\104\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\113\000\106\001\118\000\155\000\122\000\154\000\123\000\104\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\111\000\108\001\113\000\107\001\118\000\155\000\122\000\154\000\
+\\123\000\104\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\113\000\109\001\118\000\155\000\122\000\154\000\123\000\104\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\113\000\110\001\118\000\155\000\122\000\154\000\123\000\104\001\000\000\
+\\101\000\111\001\000\000\
+\\011\000\115\001\119\000\114\001\120\000\113\001\121\000\112\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\164\000\
+\\109\000\163\000\110\000\162\000\111\000\161\000\112\000\160\000\
+\\113\000\159\000\115\000\158\000\116\000\157\000\117\000\156\000\
+\\118\000\155\000\122\000\154\000\123\000\153\000\124\000\152\000\
+\\125\000\151\000\126\000\150\000\127\000\149\000\128\000\148\000\
+\\129\000\147\000\130\000\116\001\000\000\
+\\009\000\087\000\019\000\118\001\031\000\117\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\164\000\
+\\109\000\163\000\110\000\162\000\111\000\161\000\112\000\160\000\
+\\113\000\159\000\115\000\158\000\116\000\157\000\117\000\156\000\
+\\118\000\155\000\122\000\154\000\123\000\153\000\124\000\152\000\
+\\125\000\151\000\126\000\150\000\127\000\149\000\128\000\148\000\
+\\129\000\147\000\130\000\119\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\164\000\
+\\109\000\163\000\110\000\162\000\111\000\161\000\112\000\160\000\
+\\113\000\159\000\115\000\158\000\116\000\157\000\117\000\156\000\
+\\118\000\155\000\122\000\154\000\123\000\153\000\124\000\152\000\
+\\125\000\151\000\126\000\150\000\127\000\149\000\128\000\148\000\
+\\129\000\147\000\130\000\017\001\000\000\
+\\011\000\115\001\105\000\126\001\106\000\125\001\119\000\114\001\
+\\120\000\124\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\005\000\131\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\006\000\033\001\007\000\139\001\008\000\138\001\009\000\031\001\
+\\010\000\030\001\011\000\029\001\012\000\028\001\013\000\027\001\
+\\014\000\084\000\016\000\026\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\060\000\149\001\063\000\106\000\065\000\105\000\066\000\104\000\
+\\067\000\103\000\068\000\102\000\069\000\101\000\070\000\100\000\
+\\071\000\099\000\072\000\232\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\074\001\047\000\073\001\081\000\154\001\082\000\153\001\
+\\083\000\152\001\084\000\070\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\075\000\160\001\077\000\131\000\080\000\130\000\088\000\127\000\
+\\092\000\126\000\093\000\125\000\094\000\124\000\095\000\123\000\
+\\096\000\122\000\097\000\121\000\098\000\120\000\099\000\250\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\011\000\090\001\078\000\165\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\164\000\
+\\109\000\163\000\110\000\162\000\111\000\161\000\112\000\160\000\
+\\113\000\159\000\115\000\158\000\116\000\157\000\117\000\156\000\
+\\118\000\155\000\122\000\154\000\123\000\153\000\124\000\152\000\
+\\125\000\151\000\126\000\150\000\127\000\149\000\128\000\148\000\
+\\129\000\147\000\130\000\167\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\102\000\173\001\104\000\166\000\107\000\165\000\
+\\108\000\164\000\109\000\163\000\110\000\162\000\111\000\161\000\
+\\112\000\160\000\113\000\159\000\115\000\158\000\116\000\157\000\
+\\117\000\156\000\118\000\155\000\122\000\154\000\123\000\153\000\
+\\124\000\152\000\125\000\151\000\126\000\150\000\127\000\149\000\
+\\128\000\148\000\129\000\147\000\130\000\019\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\011\000\115\001\105\000\177\001\119\000\114\001\120\000\124\001\000\000\
+\\000\000\
+\\006\000\033\001\008\000\178\001\009\000\031\001\010\000\030\001\
+\\011\000\029\001\012\000\028\001\013\000\027\001\014\000\084\000\
+\\016\000\026\001\000\000\
+\\006\000\033\001\007\000\179\001\008\000\138\001\009\000\031\001\
+\\010\000\030\001\011\000\029\001\012\000\028\001\013\000\027\001\
+\\014\000\084\000\016\000\026\001\000\000\
+\\000\000\
+\\006\000\181\001\017\000\180\001\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\139\000\020\000\082\000\022\000\081\000\023\000\138\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\074\000\133\000\076\000\132\000\077\000\131\000\080\000\130\000\
+\\086\000\129\000\087\000\128\000\088\000\127\000\092\000\126\000\
+\\093\000\125\000\094\000\124\000\095\000\123\000\096\000\122\000\
+\\097\000\121\000\098\000\120\000\099\000\119\000\100\000\182\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\001\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\063\000\055\000\062\000\057\000\061\000\058\000\060\000\
+\\059\000\184\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\061\000\108\000\062\000\107\000\063\000\106\000\065\000\105\000\
+\\066\000\104\000\067\000\103\000\068\000\102\000\069\000\101\000\
+\\070\000\100\000\071\000\099\000\072\000\098\000\073\000\185\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\101\000\168\000\103\000\167\000\104\000\166\000\
+\\107\000\165\000\108\000\164\000\109\000\163\000\110\000\162\000\
+\\111\000\161\000\112\000\160\000\113\000\159\000\115\000\158\000\
+\\116\000\157\000\117\000\156\000\118\000\155\000\122\000\154\000\
+\\123\000\153\000\124\000\152\000\125\000\151\000\126\000\150\000\
+\\127\000\149\000\128\000\148\000\129\000\147\000\130\000\146\000\
+\\131\000\186\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\043\001\021\000\189\001\022\000\081\000\
+\\023\000\080\000\024\000\079\000\025\000\182\000\026\000\077\000\
+\\027\000\181\000\028\000\075\000\029\000\074\000\030\000\073\000\
+\\031\000\072\000\032\000\179\000\033\000\070\000\034\000\069\000\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\190\001\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\000\000\
+\\000\000\
+\\000\000\
+\\011\000\058\001\064\000\192\001\000\000\
+\\000\000\
+\\000\000\
+\\009\000\074\001\047\000\073\001\083\000\193\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\074\001\047\000\073\001\081\000\198\001\082\000\153\001\
+\\083\000\152\001\084\000\070\001\000\000\
+\\000\000\
+\\011\000\081\001\089\000\080\001\090\000\079\001\091\000\200\001\000\000\
+\\009\000\074\001\047\000\073\001\083\000\201\001\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\125\000\094\000\124\000\095\000\123\000\096\000\122\000\
+\\097\000\121\000\098\000\120\000\099\000\202\001\000\000\
+\\000\000\
+\\000\000\
+\\011\000\090\001\078\000\089\001\079\000\204\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\205\001\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\125\000\094\000\124\000\095\000\123\000\096\000\122\000\
+\\097\000\121\000\098\000\120\000\099\000\206\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\011\000\115\001\119\000\114\001\120\000\113\001\121\000\210\001\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\164\000\
+\\109\000\163\000\110\000\162\000\111\000\161\000\112\000\160\000\
+\\113\000\159\000\114\000\211\001\115\000\158\000\116\000\157\000\
+\\117\000\156\000\118\000\155\000\122\000\154\000\123\000\153\000\
+\\124\000\152\000\125\000\151\000\126\000\150\000\127\000\149\000\
+\\128\000\148\000\129\000\147\000\130\000\102\001\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\164\000\
+\\109\000\163\000\110\000\162\000\111\000\161\000\112\000\160\000\
+\\113\000\159\000\115\000\158\000\116\000\157\000\117\000\156\000\
+\\118\000\155\000\122\000\154\000\123\000\153\000\124\000\152\000\
+\\125\000\151\000\126\000\150\000\127\000\149\000\128\000\148\000\
+\\129\000\147\000\130\000\212\001\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\164\000\
+\\109\000\163\000\110\000\162\000\111\000\161\000\112\000\160\000\
+\\113\000\159\000\115\000\158\000\116\000\157\000\117\000\156\000\
+\\118\000\155\000\122\000\154\000\123\000\153\000\124\000\152\000\
+\\125\000\151\000\126\000\150\000\127\000\149\000\128\000\148\000\
+\\129\000\147\000\130\000\213\001\000\000\
+\\000\000\
+\\011\000\115\001\105\000\126\001\106\000\215\001\119\000\114\001\
+\\120\000\124\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\006\000\033\001\007\000\223\001\008\000\138\001\009\000\031\001\
+\\010\000\030\001\011\000\029\001\012\000\028\001\013\000\027\001\
+\\014\000\084\000\016\000\026\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\112\000\046\000\111\000\051\000\110\000\055\000\109\000\
+\\063\000\106\000\065\000\105\000\066\000\225\001\000\000\
+\\000\000\
+\\000\000\
+\\009\000\074\001\047\000\073\001\083\000\226\001\000\000\
+\\000\000\
+\\009\000\074\001\047\000\073\001\083\000\227\001\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\229\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\231\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\118\000\155\000\122\000\154\000\123\000\232\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\094\001\
+\\118\000\155\000\122\000\154\000\123\000\234\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\235\001\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\140\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\082\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\078\000\026\000\077\000\027\000\076\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\072\000\
+\\032\000\071\000\033\000\070\000\034\000\069\000\035\000\068\000\
+\\039\000\067\000\042\000\066\000\043\000\065\000\044\000\064\000\
+\\045\000\137\000\046\000\136\000\051\000\135\000\055\000\134\000\
+\\077\000\131\000\080\000\130\000\088\000\127\000\092\000\126\000\
+\\093\000\125\000\094\000\124\000\095\000\123\000\096\000\122\000\
+\\097\000\121\000\098\000\120\000\099\000\236\001\000\000\
+\\000\000\
+\\000\000\
+\\009\000\087\000\011\000\086\000\012\000\085\000\014\000\084\000\
+\\019\000\083\000\020\000\183\000\022\000\081\000\023\000\080\000\
+\\024\000\079\000\025\000\182\000\026\000\077\000\027\000\181\000\
+\\028\000\075\000\029\000\074\000\030\000\073\000\031\000\180\000\
+\\032\000\179\000\033\000\070\000\034\000\069\000\036\000\178\000\
+\\037\000\177\000\046\000\176\000\049\000\175\000\050\000\174\000\
+\\051\000\173\000\052\000\172\000\053\000\171\000\054\000\170\000\
+\\056\000\169\000\104\000\166\000\107\000\165\000\108\000\164\000\
+\\109\000\163\000\110\000\162\000\111\000\161\000\112\000\160\000\
+\\113\000\159\000\115\000\158\000\116\000\157\000\117\000\156\000\
+\\118\000\155\000\122\000\154\000\123\000\153\000\124\000\152\000\
+\\125\000\151\000\126\000\150\000\127\000\149\000\128\000\148\000\
+\\129\000\147\000\130\000\237\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\"
+val numstates = 498
+val numrules = 282
+val s = Unsynchronized.ref "" and index = Unsynchronized.ref 0
+val string_to_int = fn () => 
+let val i = !index
+in index := i+2; Char.ord(String.sub(!s,i)) + Char.ord(String.sub(!s,i+1)) * 256
+end
+val string_to_list = fn s' =>
+    let val len = String.size s'
+        fun f () =
+           if !index < len then string_to_int() :: f()
+           else nil
+   in index := 0; s := s'; f ()
+   end
+val string_to_pairlist = fn (conv_key,conv_entry) =>
+     let fun f () =
+         case string_to_int()
+         of 0 => EMPTY
+          | n => PAIR(conv_key (n-1),conv_entry (string_to_int()),f())
+     in f
+     end
+val string_to_pairlist_default = fn (conv_key,conv_entry) =>
+    let val conv_row = string_to_pairlist(conv_key,conv_entry)
+    in fn () =>
+       let val default = conv_entry(string_to_int())
+           val row = conv_row()
+       in (row,default)
+       end
+   end
+val string_to_table = fn (convert_row,s') =>
+    let val len = String.size s'
+        fun f ()=
+           if !index < len then convert_row() :: f()
+           else nil
+     in (s := s'; index := 0; f ())
+     end
+local
+  val memo = Array.array(numstates+numrules,ERROR)
+  val _ =let fun g i=(Array.update(memo,i,REDUCE(i-numstates)); g(i+1))
+       fun f i =
+            if i=numstates then g i
+            else (Array.update(memo,i,SHIFT (STATE i)); f (i+1))
+          in f 0 handle Subscript => ()
+          end
+in
+val entry_to_action = fn 0 => ACCEPT | 1 => ERROR | j => Array.sub(memo,(j-2))
+end
+val gotoT=Array.fromList(string_to_table(string_to_pairlist(NT,STATE),gotoT))
+val actionRows=string_to_table(string_to_pairlist_default(T,entry_to_action),actionRows)
+val actionRowNumbers = string_to_list actionRowNumbers
+val actionT = let val actionRowLookUp=
+let val a=Array.fromList(actionRows) in fn i=>Array.sub(a,i) end
+in Array.fromList(map actionRowLookUp actionRowNumbers)
+end
+in LrTable.mkLrTable {actions=actionT,gotos=gotoT,numRules=numrules,
+numStates=numstates,initialState=STATE 0}
+end
+end
+local open Header in
+type pos = int
+type arg = string
+structure MlyValue = 
+struct
+datatype svalue = VOID | ntVOID of unit
+ | ATOMIC_SYSTEM_WORD of  (string) | ATOMIC_DEFINED_WORD of  (string)
+ | DISTINCT_OBJECT of  (string) | COMMENT of  (string)
+ | LOWER_WORD of  (string) | UPPER_WORD of  (string)
+ | SINGLE_QUOTED of  (string) | DOT_DECIMAL of  (string)
+ | UNSIGNED_INTEGER of  (string) | SIGNED_INTEGER of  (string)
+ | RATIONAL of  (string) | REAL of  (string) | tptp of  (tptp_problem)
+ | tptp_file of  (tptp_problem) | tptp_input of  (tptp_line)
+ | include_ of  (tptp_line) | annotated_formula of  (tptp_line)
+ | thf_annotated of  (tptp_line) | tff_annotated of  (tptp_line)
+ | fof_annotated of  (tptp_line) | cnf_annotated of  (tptp_line)
+ | formula_role of  (role) | thf_formula of  (tptp_formula)
+ | thf_logic_formula of  (tptp_formula)
+ | thf_binary_formula of  (tptp_formula)
+ | thf_binary_pair of  (tptp_formula)
+ | thf_binary_tuple of  (tptp_formula)
+ | thf_or_formula of  (tptp_formula)
+ | thf_and_formula of  (tptp_formula)
+ | thf_apply_formula of  (tptp_formula)
+ | thf_unitary_formula of  (tptp_formula)
+ | thf_quantified_formula of  (tptp_formula)
+ | thf_variable_list of  ( ( string * tptp_type option )  list)
+ | thf_variable of  (string*tptp_type option)
+ | thf_typed_variable of  (string*tptp_type option)
+ | thf_unary_formula of  (tptp_formula)
+ | thf_type_formula of  (tptp_formula*tptp_type)
+ | thf_typeable_formula of  (tptp_formula)
+ | thf_subtype of  (tptp_type) | thf_top_level_type of  (tptp_type)
+ | thf_unitary_type of  (tptp_type) | thf_binary_type of  (tptp_type)
+ | thf_mapping_type of  (tptp_type) | thf_xprod_type of  (tptp_type)
+ | thf_union_type of  (tptp_type) | thf_atom of  (tptp_formula)
+ | thf_let of  (tptp_formula) | thf_let_list of  (tptp_let list)
+ | thf_defined_var of  (tptp_let) | thf_conditional of  (tptp_formula)
+ | thf_sequent of  (tptp_formula)
+ | thf_tuple_list of  (tptp_formula list)
+ | thf_tuple of  (tptp_formula list) | tff_formula of  (tptp_formula)
+ | tff_logic_formula of  (tptp_formula)
+ | tff_binary_formula of  (tptp_formula)
+ | tff_binary_nonassoc of  (tptp_formula)
+ | tff_binary_assoc of  (tptp_formula)
+ | tff_or_formula of  (tptp_formula)
+ | tff_and_formula of  (tptp_formula)
+ | tff_unitary_formula of  (tptp_formula)
+ | tff_quantified_formula of  (tptp_formula)
+ | tff_variable_list of  ( ( string * tptp_type option )  list)
+ | tff_variable of  (string*tptp_type option)
+ | tff_typed_variable of  (string*tptp_type option)
+ | tff_unary_formula of  (tptp_formula)
+ | tff_typed_atom of  (symbol*tptp_type option)
+ | tff_untyped_atom of  (symbol*tptp_type option)
+ | tff_top_level_type of  (tptp_type)
+ | tff_unitary_type of  (tptp_type) | tff_atomic_type of  (tptp_type)
+ | tff_mapping_type of  (tptp_type) | tff_xprod_type of  (tptp_type)
+ | tptp_let of  (tptp_formula) | tff_let_list of  (tptp_let list)
+ | tff_defined_var of  (tptp_let) | tff_conditional of  (tptp_formula)
+ | tff_sequent of  (tptp_formula)
+ | tff_tuple_list of  (tptp_formula list)
+ | tff_tuple of  (tptp_formula list) | fof_formula of  (tptp_formula)
+ | fof_logic_formula of  (tptp_formula)
+ | fof_binary_formula of  (tptp_formula)
+ | fof_binary_nonassoc of  (tptp_formula)
+ | fof_binary_assoc of  (tptp_formula)
+ | fof_or_formula of  (tptp_formula)
+ | fof_and_formula of  (tptp_formula)
+ | fof_unitary_formula of  (tptp_formula)
+ | fof_quantified_formula of  (tptp_formula)
+ | fof_variable_list of  (string list)
+ | fof_unary_formula of  (tptp_formula)
+ | fof_sequent of  (tptp_formula) | fof_tuple of  (tptp_formula list)
+ | fof_tuple_list of  (tptp_formula list)
+ | cnf_formula of  (tptp_formula) | disjunction of  (tptp_formula)
+ | literal of  (tptp_formula) | thf_conn_term of  (symbol)
+ | fol_infix_unary of  (tptp_formula)
+ | thf_quantifier of  (quantifier) | thf_pair_connective of  (symbol)
+ | thf_unary_connective of  (symbol) | fol_quantifier of  (quantifier)
+ | binary_connective of  (symbol) | assoc_connective of  (symbol)
+ | system_type of  (string) | defined_type of  (tptp_base_type)
+ | unary_connective of  (symbol) | atomic_formula of  (tptp_formula)
+ | plain_atomic_formula of  (tptp_formula)
+ | defined_atomic_formula of  (tptp_formula)
+ | defined_plain_formula of  (tptp_formula)
+ | defined_pred of  (string) | defined_prop of  (string)
+ | defined_infix_formula of  (tptp_formula)
+ | defined_infix_pred of  (symbol) | infix_inequality of  (symbol)
+ | infix_equality of  (symbol)
+ | system_atomic_formula of  (tptp_formula)
+ | conditional_term of  (tptp_term) | function_term of  (tptp_term)
+ | plain_term of  (symbol*tptp_term list) | constant of  (symbol)
+ | defined_term of  (tptp_term) | defined_atom of  (tptp_term)
+ | defined_atomic_term of  (tptp_term)
+ | defined_plain_term of  (symbol*tptp_term list)
+ | defined_constant of  (symbol)
+ | system_term of  (symbol*tptp_term list)
+ | system_constant of  (symbol) | system_functor of  (symbol)
+ | defined_functor of  (symbol) | arguments of  (tptp_term list)
+ | term of  (tptp_term) | functor_ of  (symbol)
+ | file_name of  (string) | useful_info of  (general_list)
+ | general_function of  (general_data) | identifier of  (string)
+ | integer of  (string) | formula_data of  (general_data)
+ | number of  (number_kind*string) | variable_ of  (string)
+ | general_data of  (general_data) | atomic_word of  (string)
+ | general_term of  (general_term)
+ | general_terms of  (general_term list)
+ | general_list of  (general_list)
+ | optional_info of  (general_term list)
+ | formula_selection of  (string list) | name_list of  (string list)
+ | name of  (string) | annotations of  (annotation option)
+end
+type svalue = MlyValue.svalue
+type result = tptp_problem
+end
+structure EC=
+struct
+open LrTable
+infix 5 $$
+fun x $$ y = y::x
+val is_keyword =
+fn _ => false
+val preferred_change : (term list * term list) list = 
+nil
+val noShift = 
+fn (T 37) => true | _ => false
+val showTerminal =
+fn (T 0) => "AMPERSAND"
+  | (T 1) => "AT_SIGN"
+  | (T 2) => "CARET"
+  | (T 3) => "COLON"
+  | (T 4) => "COMMA"
+  | (T 5) => "EQUALS"
+  | (T 6) => "EXCLAMATION"
+  | (T 7) => "LET"
+  | (T 8) => "ARROW"
+  | (T 9) => "IF"
+  | (T 10) => "IFF"
+  | (T 11) => "IMPLIES"
+  | (T 12) => "INCLUDE"
+  | (T 13) => "LAMBDA"
+  | (T 14) => "LBRKT"
+  | (T 15) => "LPAREN"
+  | (T 16) => "MAP_TO"
+  | (T 17) => "MMINUS"
+  | (T 18) => "NAND"
+  | (T 19) => "NEQUALS"
+  | (T 20) => "XOR"
+  | (T 21) => "NOR"
+  | (T 22) => "PERIOD"
+  | (T 23) => "PPLUS"
+  | (T 24) => "QUESTION"
+  | (T 25) => "RBRKT"
+  | (T 26) => "RPAREN"
+  | (T 27) => "TILDE"
+  | (T 28) => "TOK_FALSE"
+  | (T 29) => "TOK_I"
+  | (T 30) => "TOK_O"
+  | (T 31) => "TOK_INT"
+  | (T 32) => "TOK_REAL"
+  | (T 33) => "TOK_RAT"
+  | (T 34) => "TOK_TRUE"
+  | (T 35) => "TOK_TYPE"
+  | (T 36) => "VLINE"
+  | (T 37) => "EOF"
+  | (T 38) => "DTHF"
+  | (T 39) => "DFOF"
+  | (T 40) => "DCNF"
+  | (T 41) => "DFOT"
+  | (T 42) => "DTFF"
+  | (T 43) => "REAL"
+  | (T 44) => "RATIONAL"
+  | (T 45) => "SIGNED_INTEGER"
+  | (T 46) => "UNSIGNED_INTEGER"
+  | (T 47) => "DOT_DECIMAL"
+  | (T 48) => "SINGLE_QUOTED"
+  | (T 49) => "UPPER_WORD"
+  | (T 50) => "LOWER_WORD"
+  | (T 51) => "COMMENT"
+  | (T 52) => "DISTINCT_OBJECT"
+  | (T 53) => "DUD"
+  | (T 54) => "INDEF_CHOICE"
+  | (T 55) => "DEFIN_CHOICE"
+  | (T 56) => "OPERATOR_FORALL"
+  | (T 57) => "OPERATOR_EXISTS"
+  | (T 58) => "PLUS"
+  | (T 59) => "TIMES"
+  | (T 60) => "GENTZEN_ARROW"
+  | (T 61) => "DEP_SUM"
+  | (T 62) => "DEP_PROD"
+  | (T 63) => "ATOMIC_DEFINED_WORD"
+  | (T 64) => "ATOMIC_SYSTEM_WORD"
+  | (T 65) => "SUBTYPE"
+  | (T 66) => "LET_TERM"
+  | (T 67) => "THF"
+  | (T 68) => "TFF"
+  | (T 69) => "FOF"
+  | (T 70) => "CNF"
+  | (T 71) => "ITE_F"
+  | (T 72) => "ITE_T"
+  | _ => "bogus-term"
+local open Header in
+val errtermvalue=
+fn _ => MlyValue.VOID
+end
+val terms : term list = nil
+ $$ (T 72) $$ (T 71) $$ (T 70) $$ (T 69) $$ (T 68) $$ (T 67) $$ (T 66)
+ $$ (T 65) $$ (T 62) $$ (T 61) $$ (T 60) $$ (T 59) $$ (T 58) $$ (T 57)
+ $$ (T 56) $$ (T 55) $$ (T 54) $$ (T 53) $$ (T 42) $$ (T 41) $$ (T 40)
+ $$ (T 39) $$ (T 38) $$ (T 37) $$ (T 36) $$ (T 35) $$ (T 34) $$ (T 33)
+ $$ (T 32) $$ (T 31) $$ (T 30) $$ (T 29) $$ (T 28) $$ (T 27) $$ (T 26)
+ $$ (T 25) $$ (T 24) $$ (T 23) $$ (T 22) $$ (T 21) $$ (T 20) $$ (T 19)
+ $$ (T 18) $$ (T 17) $$ (T 16) $$ (T 15) $$ (T 14) $$ (T 13) $$ (T 12)
+ $$ (T 11) $$ (T 10) $$ (T 9) $$ (T 8) $$ (T 7) $$ (T 6) $$ (T 5) $$ 
+(T 4) $$ (T 3) $$ (T 2) $$ (T 1) $$ (T 0)end
+structure Actions =
+struct 
+exception mlyAction of int
+local open Header in
+val actions = 
+fn (i392,defaultPos,stack,
+    (file_name):arg) =>
+case (i392,stack)
+of  ( 0, ( ( _, ( MlyValue.optional_info optional_info, _, 
+optional_info1right)) :: ( _, ( MlyValue.general_term general_term, _,
+ _)) :: ( _, ( _, COMMA1left, _)) :: rest671)) => let val  result = 
+MlyValue.annotations (( SOME (general_term, optional_info) ))
+ in ( LrTable.NT 0, ( result, COMMA1left, optional_info1right), 
+rest671)
+end
+|  ( 1, ( rest671)) => let val  result = MlyValue.annotations (
+( NONE ))
+ in ( LrTable.NT 0, ( result, defaultPos, defaultPos), rest671)
+end
+|  ( 2, ( ( _, ( MlyValue.useful_info useful_info, _, 
+useful_info1right)) :: ( _, ( _, COMMA1left, _)) :: rest671)) => let
+ val  result = MlyValue.optional_info (( useful_info ))
+ in ( LrTable.NT 4, ( result, COMMA1left, useful_info1right), rest671)
+
+end
+|  ( 3, ( rest671)) => let val  result = MlyValue.optional_info (
+( [] ))
+ in ( LrTable.NT 4, ( result, defaultPos, defaultPos), rest671)
+end
+|  ( 4, ( ( _, ( MlyValue.general_list general_list, general_list1left
+, general_list1right)) :: rest671)) => let val  result = 
+MlyValue.useful_info (( general_list ))
+ in ( LrTable.NT 16, ( result, general_list1left, general_list1right),
+ rest671)
+end
+|  ( 5, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( MlyValue.general_terms 
+general_terms, _, _)) :: ( _, ( _, LBRKT1left, _)) :: rest671)) => let
+ val  result = MlyValue.general_list (( general_terms ))
+ in ( LrTable.NT 5, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 6, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( _, LBRKT1left, _)) :: 
+rest671)) => let val  result = MlyValue.general_list (( [] ))
+ in ( LrTable.NT 5, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 7, ( ( _, ( MlyValue.general_terms general_terms, _, 
+general_terms1right)) :: _ :: ( _, ( MlyValue.general_term 
+general_term, general_term1left, _)) :: rest671)) => let val  result =
+ MlyValue.general_terms (( general_term :: general_terms ))
+ in ( LrTable.NT 6, ( result, general_term1left, general_terms1right),
+ rest671)
+end
+|  ( 8, ( ( _, ( MlyValue.general_term general_term, general_term1left
+, general_term1right)) :: rest671)) => let val  result = 
+MlyValue.general_terms (( [general_term] ))
+ in ( LrTable.NT 6, ( result, general_term1left, general_term1right), 
+rest671)
+end
+|  ( 9, ( ( _, ( MlyValue.general_data general_data, general_data1left
+, general_data1right)) :: rest671)) => let val  result = 
+MlyValue.general_term (( General_Data general_data ))
+ in ( LrTable.NT 7, ( result, general_data1left, general_data1right), 
+rest671)
+end
+|  ( 10, ( ( _, ( MlyValue.general_term general_term, _, 
+general_term1right)) :: _ :: ( _, ( MlyValue.general_data general_data
+, general_data1left, _)) :: rest671)) => let val  result = 
+MlyValue.general_term (( General_Term (general_data, general_term) ))
+ in ( LrTable.NT 7, ( result, general_data1left, general_term1right), 
+rest671)
+end
+|  ( 11, ( ( _, ( MlyValue.general_list general_list, 
+general_list1left, general_list1right)) :: rest671)) => let val  
+result = MlyValue.general_term (( General_List general_list ))
+ in ( LrTable.NT 7, ( result, general_list1left, general_list1right), 
+rest671)
+end
+|  ( 12, ( ( _, ( MlyValue.LOWER_WORD LOWER_WORD, LOWER_WORD1left, 
+LOWER_WORD1right)) :: rest671)) => let val  result = 
+MlyValue.atomic_word (( LOWER_WORD ))
+ in ( LrTable.NT 8, ( result, LOWER_WORD1left, LOWER_WORD1right), 
+rest671)
+end
+|  ( 13, ( ( _, ( MlyValue.SINGLE_QUOTED SINGLE_QUOTED, 
+SINGLE_QUOTED1left, SINGLE_QUOTED1right)) :: rest671)) => let val  
+result = MlyValue.atomic_word (( SINGLE_QUOTED ))
+ in ( LrTable.NT 8, ( result, SINGLE_QUOTED1left, SINGLE_QUOTED1right)
+, rest671)
+end
+|  ( 14, ( ( _, ( _, THF1left, THF1right)) :: rest671)) => let val  
+result = MlyValue.atomic_word (( "thf" ))
+ in ( LrTable.NT 8, ( result, THF1left, THF1right), rest671)
+end
+|  ( 15, ( ( _, ( _, TFF1left, TFF1right)) :: rest671)) => let val  
+result = MlyValue.atomic_word (( "tff" ))
+ in ( LrTable.NT 8, ( result, TFF1left, TFF1right), rest671)
+end
+|  ( 16, ( ( _, ( _, FOF1left, FOF1right)) :: rest671)) => let val  
+result = MlyValue.atomic_word (( "fof" ))
+ in ( LrTable.NT 8, ( result, FOF1left, FOF1right), rest671)
+end
+|  ( 17, ( ( _, ( _, CNF1left, CNF1right)) :: rest671)) => let val  
+result = MlyValue.atomic_word (( "cnf" ))
+ in ( LrTable.NT 8, ( result, CNF1left, CNF1right), rest671)
+end
+|  ( 18, ( ( _, ( _, INCLUDE1left, INCLUDE1right)) :: rest671)) => let
+ val  result = MlyValue.atomic_word (( "include" ))
+ in ( LrTable.NT 8, ( result, INCLUDE1left, INCLUDE1right), rest671)
+
+end
+|  ( 19, ( ( _, ( MlyValue.UPPER_WORD UPPER_WORD, UPPER_WORD1left, 
+UPPER_WORD1right)) :: rest671)) => let val  result = 
+MlyValue.variable_ (( UPPER_WORD ))
+ in ( LrTable.NT 10, ( result, UPPER_WORD1left, UPPER_WORD1right), 
+rest671)
+end
+|  ( 20, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.general_terms general_terms, _, _)) :: _ :: ( _, ( 
+MlyValue.atomic_word atomic_word, atomic_word1left, _)) :: rest671))
+ => let val  result = MlyValue.general_function (
+( Application (atomic_word, general_terms) ))
+ in ( LrTable.NT 15, ( result, atomic_word1left, RPAREN1right), 
+rest671)
+end
+|  ( 21, ( ( _, ( MlyValue.atomic_word atomic_word, atomic_word1left, 
+atomic_word1right)) :: rest671)) => let val  result = 
+MlyValue.general_data (( Atomic_Word atomic_word ))
+ in ( LrTable.NT 9, ( result, atomic_word1left, atomic_word1right), 
+rest671)
+end
+|  ( 22, ( ( _, ( MlyValue.general_function general_function, 
+general_function1left, general_function1right)) :: rest671)) => let
+ val  result = MlyValue.general_data (( general_function ))
+ in ( LrTable.NT 9, ( result, general_function1left, 
+general_function1right), rest671)
+end
+|  ( 23, ( ( _, ( MlyValue.variable_ variable_, variable_1left, 
+variable_1right)) :: rest671)) => let val  result = 
+MlyValue.general_data (( V variable_ ))
+ in ( LrTable.NT 9, ( result, variable_1left, variable_1right), 
+rest671)
+end
+|  ( 24, ( ( _, ( MlyValue.number number, number1left, number1right))
+ :: rest671)) => let val  result = MlyValue.general_data (
+( Number number ))
+ in ( LrTable.NT 9, ( result, number1left, number1right), rest671)
+end
+|  ( 25, ( ( _, ( MlyValue.DISTINCT_OBJECT DISTINCT_OBJECT, 
+DISTINCT_OBJECT1left, DISTINCT_OBJECT1right)) :: rest671)) => let val 
+ result = MlyValue.general_data (( Distinct_Object DISTINCT_OBJECT ))
+ in ( LrTable.NT 9, ( result, DISTINCT_OBJECT1left, 
+DISTINCT_OBJECT1right), rest671)
+end
+|  ( 26, ( ( _, ( MlyValue.formula_data formula_data, 
+formula_data1left, formula_data1right)) :: rest671)) => let val  
+result = MlyValue.general_data (( formula_data ))
+ in ( LrTable.NT 9, ( result, formula_data1left, formula_data1right), 
+rest671)
+end
+|  ( 27, ( ( _, ( MlyValue.integer integer, integer1left, 
+integer1right)) :: rest671)) => let val  result = MlyValue.number (
+( (Int_num, integer) ))
+ in ( LrTable.NT 11, ( result, integer1left, integer1right), rest671)
+
+end
+|  ( 28, ( ( _, ( MlyValue.REAL REAL, REAL1left, REAL1right)) :: 
+rest671)) => let val  result = MlyValue.number (( (Real_num, REAL) ))
+ in ( LrTable.NT 11, ( result, REAL1left, REAL1right), rest671)
+end
+|  ( 29, ( ( _, ( MlyValue.RATIONAL RATIONAL, RATIONAL1left, 
+RATIONAL1right)) :: rest671)) => let val  result = MlyValue.number (
+( (Rat_num, RATIONAL) ))
+ in ( LrTable.NT 11, ( result, RATIONAL1left, RATIONAL1right), rest671
+)
+end
+|  ( 30, ( ( _, ( MlyValue.UNSIGNED_INTEGER UNSIGNED_INTEGER, 
+UNSIGNED_INTEGER1left, UNSIGNED_INTEGER1right)) :: rest671)) => let
+ val  result = MlyValue.integer (( UNSIGNED_INTEGER ))
+ in ( LrTable.NT 13, ( result, UNSIGNED_INTEGER1left, 
+UNSIGNED_INTEGER1right), rest671)
+end
+|  ( 31, ( ( _, ( MlyValue.SIGNED_INTEGER SIGNED_INTEGER, 
+SIGNED_INTEGER1left, SIGNED_INTEGER1right)) :: rest671)) => let val  
+result = MlyValue.integer (( SIGNED_INTEGER ))
+ in ( LrTable.NT 13, ( result, SIGNED_INTEGER1left, 
+SIGNED_INTEGER1right), rest671)
+end
+|  ( 32, ( ( _, ( MlyValue.SINGLE_QUOTED SINGLE_QUOTED, 
+SINGLE_QUOTED1left, SINGLE_QUOTED1right)) :: rest671)) => let val  
+result = MlyValue.file_name (( SINGLE_QUOTED ))
+ in ( LrTable.NT 17, ( result, SINGLE_QUOTED1left, SINGLE_QUOTED1right
+), rest671)
+end
+|  ( 33, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.thf_formula 
+thf_formula, _, _)) :: _ :: ( _, ( _, DTHF1left, _)) :: rest671)) =>
+ let val  result = MlyValue.formula_data (
+( Formula_Data (THF, thf_formula) ))
+ in ( LrTable.NT 12, ( result, DTHF1left, RPAREN1right), rest671)
+end
+|  ( 34, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.tff_formula 
+tff_formula, _, _)) :: _ :: ( _, ( _, DTFF1left, _)) :: rest671)) =>
+ let val  result = MlyValue.formula_data (
+( Formula_Data (TFF, tff_formula) ))
+ in ( LrTable.NT 12, ( result, DTFF1left, RPAREN1right), rest671)
+end
+|  ( 35, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.fof_formula 
+fof_formula, _, _)) :: _ :: ( _, ( _, DFOF1left, _)) :: rest671)) =>
+ let val  result = MlyValue.formula_data (
+( Formula_Data (FOF, fof_formula) ))
+ in ( LrTable.NT 12, ( result, DFOF1left, RPAREN1right), rest671)
+end
+|  ( 36, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.cnf_formula 
+cnf_formula, _, _)) :: _ :: ( _, ( _, DCNF1left, _)) :: rest671)) =>
+ let val  result = MlyValue.formula_data (
+( Formula_Data (CNF, cnf_formula) ))
+ in ( LrTable.NT 12, ( result, DCNF1left, RPAREN1right), rest671)
+end
+|  ( 37, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.term term, _
+, _)) :: _ :: ( _, ( _, DFOT1left, _)) :: rest671)) => let val  result
+ = MlyValue.formula_data (( Term_Data term ))
+ in ( LrTable.NT 12, ( result, DFOT1left, RPAREN1right), rest671)
+end
+|  ( 38, ( ( _, ( MlyValue.ATOMIC_SYSTEM_WORD ATOMIC_SYSTEM_WORD, 
+ATOMIC_SYSTEM_WORD1left, ATOMIC_SYSTEM_WORD1right)) :: rest671)) =>
+ let val  result = MlyValue.system_type (( ATOMIC_SYSTEM_WORD ))
+ in ( LrTable.NT 47, ( result, ATOMIC_SYSTEM_WORD1left, 
+ATOMIC_SYSTEM_WORD1right), rest671)
+end
+|  ( 39, ( ( _, ( MlyValue.ATOMIC_DEFINED_WORD ATOMIC_DEFINED_WORD, 
+ATOMIC_DEFINED_WORD1left, ATOMIC_DEFINED_WORD1right)) :: rest671)) =>
+ let val  result = MlyValue.defined_type (
+(
+  case ATOMIC_DEFINED_WORD of
+    "$i" => Type_Ind
+  | "$o" => Type_Bool
+  | "$iType" => Type_Ind
+  | "$oType" => Type_Bool
+  | "$int" => Type_Int
+  | "$real" => Type_Real
+  | "$rat" => Type_Rat
+  | "$tType" => Type_Type
+  | thing => raise UNRECOGNISED_SYMBOL ("defined_type", thing)
+)
+)
+ in ( LrTable.NT 46, ( result, ATOMIC_DEFINED_WORD1left, 
+ATOMIC_DEFINED_WORD1right), rest671)
+end
+|  ( 40, ( ( _, ( MlyValue.atomic_word atomic_word, atomic_word1left, 
+atomic_word1right)) :: rest671)) => let val  result = 
+MlyValue.functor_ (( Uninterpreted atomic_word ))
+ in ( LrTable.NT 18, ( result, atomic_word1left, atomic_word1right), 
+rest671)
+end
+|  ( 41, ( ( _, ( MlyValue.term term, term1left, term1right)) :: 
+rest671)) => let val  result = MlyValue.arguments (( [term] ))
+ in ( LrTable.NT 20, ( result, term1left, term1right), rest671)
+end
+|  ( 42, ( ( _, ( MlyValue.arguments arguments, _, arguments1right))
+ :: _ :: ( _, ( MlyValue.term term, term1left, _)) :: rest671)) => let
+ val  result = MlyValue.arguments (( term :: arguments ))
+ in ( LrTable.NT 20, ( result, term1left, arguments1right), rest671)
+
+end
+|  ( 43, ( ( _, ( MlyValue.ATOMIC_SYSTEM_WORD ATOMIC_SYSTEM_WORD, 
+ATOMIC_SYSTEM_WORD1left, ATOMIC_SYSTEM_WORD1right)) :: rest671)) =>
+ let val  result = MlyValue.system_functor (
+( System ATOMIC_SYSTEM_WORD ))
+ in ( LrTable.NT 22, ( result, ATOMIC_SYSTEM_WORD1left, 
+ATOMIC_SYSTEM_WORD1right), rest671)
+end
+|  ( 44, ( ( _, ( MlyValue.system_functor system_functor, 
+system_functor1left, system_functor1right)) :: rest671)) => let val  
+result = MlyValue.system_constant (( system_functor ))
+ in ( LrTable.NT 23, ( result, system_functor1left, 
+system_functor1right), rest671)
+end
+|  ( 45, ( ( _, ( MlyValue.system_constant system_constant, 
+system_constant1left, system_constant1right)) :: rest671)) => let val 
+ result = MlyValue.system_term (( (system_constant, []) ))
+ in ( LrTable.NT 24, ( result, system_constant1left, 
+system_constant1right), rest671)
+end
+|  ( 46, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.arguments 
+arguments, _, _)) :: _ :: ( _, ( MlyValue.system_functor 
+system_functor, system_functor1left, _)) :: rest671)) => let val  
+result = MlyValue.system_term (( (system_functor, arguments) ))
+ in ( LrTable.NT 24, ( result, system_functor1left, RPAREN1right), 
+rest671)
+end
+|  ( 47, ( ( _, ( MlyValue.ATOMIC_DEFINED_WORD ATOMIC_DEFINED_WORD, 
+ATOMIC_DEFINED_WORD1left, ATOMIC_DEFINED_WORD1right)) :: rest671)) =>
+ let val  result = MlyValue.defined_functor (
+(
+  case ATOMIC_DEFINED_WORD of
+    "$sum" => Interpreted_ExtraLogic Sum
+  | "$difference" => Interpreted_ExtraLogic Difference
+  | "$product" => Interpreted_ExtraLogic Product
+  | "$quotient" => Interpreted_ExtraLogic Quotient
+  | "$quotient_e" => Interpreted_ExtraLogic Quotient_E
+  | "$quotient_t" => Interpreted_ExtraLogic Quotient_T
+  | "$quotient_f" => Interpreted_ExtraLogic Quotient_F
+  | "$remainder_e" => Interpreted_ExtraLogic Remainder_E
+  | "$remainder_t" => Interpreted_ExtraLogic Remainder_T
+  | "$remainder_f" => Interpreted_ExtraLogic Remainder_F
+  | "$floor" => Interpreted_ExtraLogic Floor
+  | "$ceiling" => Interpreted_ExtraLogic Ceiling
+  | "$truncate" => Interpreted_ExtraLogic Truncate
+  | "$round" => Interpreted_ExtraLogic Round
+  | "$to_int" => Interpreted_ExtraLogic To_Int
+  | "$to_rat" => Interpreted_ExtraLogic To_Rat
+  | "$to_real" => Interpreted_ExtraLogic To_Real
+  | "$uminus" => Interpreted_ExtraLogic UMinus
+
+  | "$i" => TypeSymbol Type_Ind
+  | "$o" => TypeSymbol Type_Bool
+  | "$iType" => TypeSymbol Type_Ind
+  | "$oType" => TypeSymbol Type_Bool
+  | "$int" => TypeSymbol Type_Int
+  | "$real" => TypeSymbol Type_Real
+  | "$rat" => TypeSymbol Type_Rat
+  | "$tType" => TypeSymbol Type_Type
+
+  | "$true" => Interpreted_Logic True
+  | "$false" => Interpreted_Logic False
+
+  | "$less" => Interpreted_ExtraLogic Less
+  | "$lesseq" => Interpreted_ExtraLogic LessEq
+  | "$greatereq" => Interpreted_ExtraLogic GreaterEq
+  | "$greater" => Interpreted_ExtraLogic Greater
+  | "$evaleq" => Interpreted_ExtraLogic EvalEq
+
+  | "$is_int" => Interpreted_ExtraLogic Is_Int
+  | "$is_rat" => Interpreted_ExtraLogic Is_Rat
+
+  | thing => raise UNRECOGNISED_SYMBOL ("defined_functor", thing)
+)
+)
+ in ( LrTable.NT 21, ( result, ATOMIC_DEFINED_WORD1left, 
+ATOMIC_DEFINED_WORD1right), rest671)
+end
+|  ( 48, ( ( _, ( MlyValue.defined_functor defined_functor, 
+defined_functor1left, defined_functor1right)) :: rest671)) => let val 
+ result = MlyValue.defined_constant (( defined_functor ))
+ in ( LrTable.NT 25, ( result, defined_functor1left, 
+defined_functor1right), rest671)
+end
+|  ( 49, ( ( _, ( MlyValue.defined_constant defined_constant, 
+defined_constant1left, defined_constant1right)) :: rest671)) => let
+ val  result = MlyValue.defined_plain_term (( (defined_constant, []) )
+)
+ in ( LrTable.NT 26, ( result, defined_constant1left, 
+defined_constant1right), rest671)
+end
+|  ( 50, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.arguments 
+arguments, _, _)) :: _ :: ( _, ( MlyValue.defined_functor 
+defined_functor, defined_functor1left, _)) :: rest671)) => let val  
+result = MlyValue.defined_plain_term (( (defined_functor, arguments) )
+)
+ in ( LrTable.NT 26, ( result, defined_functor1left, RPAREN1right), 
+rest671)
+end
+|  ( 51, ( ( _, ( MlyValue.defined_plain_term defined_plain_term, 
+defined_plain_term1left, defined_plain_term1right)) :: rest671)) =>
+ let val  result = MlyValue.defined_atomic_term (
+( Term_Func defined_plain_term ))
+ in ( LrTable.NT 27, ( result, defined_plain_term1left, 
+defined_plain_term1right), rest671)
+end
+|  ( 52, ( ( _, ( MlyValue.number number, number1left, number1right))
+ :: rest671)) => let val  result = MlyValue.defined_atom (
+( Term_Num number ))
+ in ( LrTable.NT 28, ( result, number1left, number1right), rest671)
+
+end
+|  ( 53, ( ( _, ( MlyValue.DISTINCT_OBJECT DISTINCT_OBJECT, 
+DISTINCT_OBJECT1left, DISTINCT_OBJECT1right)) :: rest671)) => let val 
+ result = MlyValue.defined_atom (
+( Term_Distinct_Object DISTINCT_OBJECT ))
+ in ( LrTable.NT 28, ( result, DISTINCT_OBJECT1left, 
+DISTINCT_OBJECT1right), rest671)
+end
+|  ( 54, ( ( _, ( MlyValue.defined_atom defined_atom, 
+defined_atom1left, defined_atom1right)) :: rest671)) => let val  
+result = MlyValue.defined_term (( defined_atom ))
+ in ( LrTable.NT 29, ( result, defined_atom1left, defined_atom1right),
+ rest671)
+end
+|  ( 55, ( ( _, ( MlyValue.defined_atomic_term defined_atomic_term, 
+defined_atomic_term1left, defined_atomic_term1right)) :: rest671)) =>
+ let val  result = MlyValue.defined_term (( defined_atomic_term ))
+ in ( LrTable.NT 29, ( result, defined_atomic_term1left, 
+defined_atomic_term1right), rest671)
+end
+|  ( 56, ( ( _, ( MlyValue.functor_ functor_, functor_1left, 
+functor_1right)) :: rest671)) => let val  result = MlyValue.constant (
+( functor_ ))
+ in ( LrTable.NT 30, ( result, functor_1left, functor_1right), rest671
+)
+end
+|  ( 57, ( ( _, ( MlyValue.constant constant, constant1left, 
+constant1right)) :: rest671)) => let val  result = MlyValue.plain_term
+ (( (constant, []) ))
+ in ( LrTable.NT 31, ( result, constant1left, constant1right), rest671
+)
+end
+|  ( 58, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.arguments 
+arguments, _, _)) :: _ :: ( _, ( MlyValue.functor_ functor_, 
+functor_1left, _)) :: rest671)) => let val  result = 
+MlyValue.plain_term (( (functor_, arguments) ))
+ in ( LrTable.NT 31, ( result, functor_1left, RPAREN1right), rest671)
+
+end
+|  ( 59, ( ( _, ( MlyValue.plain_term plain_term, plain_term1left, 
+plain_term1right)) :: rest671)) => let val  result = 
+MlyValue.function_term (( Term_Func plain_term ))
+ in ( LrTable.NT 32, ( result, plain_term1left, plain_term1right), 
+rest671)
+end
+|  ( 60, ( ( _, ( MlyValue.defined_term defined_term, 
+defined_term1left, defined_term1right)) :: rest671)) => let val  
+result = MlyValue.function_term (( defined_term ))
+ in ( LrTable.NT 32, ( result, defined_term1left, defined_term1right),
+ rest671)
+end
+|  ( 61, ( ( _, ( MlyValue.system_term system_term, system_term1left, 
+system_term1right)) :: rest671)) => let val  result = 
+MlyValue.function_term (( Term_Func system_term ))
+ in ( LrTable.NT 32, ( result, system_term1left, system_term1right), 
+rest671)
+end
+|  ( 62, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.term term2,
+ _, _)) :: _ :: ( _, ( MlyValue.term term1, _, _)) :: _ :: ( _, ( 
+MlyValue.tff_logic_formula tff_logic_formula, _, _)) :: _ :: ( _, ( _,
+ ITE_T1left, _)) :: rest671)) => let val  result = 
+MlyValue.conditional_term (
+(
+  Term_Conditional (tff_logic_formula, term1, term2)
+))
+ in ( LrTable.NT 33, ( result, ITE_T1left, RPAREN1right), rest671)
+end
+|  ( 63, ( ( _, ( MlyValue.function_term function_term, 
+function_term1left, function_term1right)) :: rest671)) => let val  
+result = MlyValue.term (( function_term ))
+ in ( LrTable.NT 19, ( result, function_term1left, function_term1right
+), rest671)
+end
+|  ( 64, ( ( _, ( MlyValue.variable_ variable_, variable_1left, 
+variable_1right)) :: rest671)) => let val  result = MlyValue.term (
+( Term_Var variable_ ))
+ in ( LrTable.NT 19, ( result, variable_1left, variable_1right), 
+rest671)
+end
+|  ( 65, ( ( _, ( MlyValue.conditional_term conditional_term, 
+conditional_term1left, conditional_term1right)) :: rest671)) => let
+ val  result = MlyValue.term (( conditional_term ))
+ in ( LrTable.NT 19, ( result, conditional_term1left, 
+conditional_term1right), rest671)
+end
+|  ( 66, ( ( _, ( MlyValue.system_term system_term, system_term1left, 
+system_term1right)) :: rest671)) => let val  result = 
+MlyValue.system_atomic_formula (( Pred system_term ))
+ in ( LrTable.NT 34, ( result, system_term1left, system_term1right), 
+rest671)
+end
+|  ( 67, ( ( _, ( _, EQUALS1left, EQUALS1right)) :: rest671)) => let
+ val  result = MlyValue.infix_equality (( Interpreted_Logic Equals ))
+ in ( LrTable.NT 35, ( result, EQUALS1left, EQUALS1right), rest671)
+
+end
+|  ( 68, ( ( _, ( _, NEQUALS1left, NEQUALS1right)) :: rest671)) => let
+ val  result = MlyValue.infix_inequality (
+( Interpreted_Logic NEquals ))
+ in ( LrTable.NT 36, ( result, NEQUALS1left, NEQUALS1right), rest671)
+
+end
+|  ( 69, ( ( _, ( MlyValue.infix_equality infix_equality, 
+infix_equality1left, infix_equality1right)) :: rest671)) => let val  
+result = MlyValue.defined_infix_pred (( infix_equality ))
+ in ( LrTable.NT 37, ( result, infix_equality1left, 
+infix_equality1right), rest671)
+end
+|  ( 70, ( ( _, ( MlyValue.term term2, _, term2right)) :: ( _, ( 
+MlyValue.defined_infix_pred defined_infix_pred, _, _)) :: ( _, ( 
+MlyValue.term term1, term1left, _)) :: rest671)) => let val  result = 
+MlyValue.defined_infix_formula (
+(Pred (defined_infix_pred, [term1, term2])))
+ in ( LrTable.NT 38, ( result, term1left, term2right), rest671)
+end
+|  ( 71, ( ( _, ( MlyValue.ATOMIC_DEFINED_WORD ATOMIC_DEFINED_WORD, 
+ATOMIC_DEFINED_WORD1left, ATOMIC_DEFINED_WORD1right)) :: rest671)) =>
+ let val  result = MlyValue.defined_prop (
+(
+  case ATOMIC_DEFINED_WORD of
+    "$true"  => "$true"
+  | "$false" => "$false"
+  | thing => raise UNRECOGNISED_SYMBOL ("defined_prop", thing)
+)
+)
+ in ( LrTable.NT 39, ( result, ATOMIC_DEFINED_WORD1left, 
+ATOMIC_DEFINED_WORD1right), rest671)
+end
+|  ( 72, ( ( _, ( MlyValue.ATOMIC_DEFINED_WORD ATOMIC_DEFINED_WORD, 
+ATOMIC_DEFINED_WORD1left, ATOMIC_DEFINED_WORD1right)) :: rest671)) =>
+ let val  result = MlyValue.defined_pred (
+(
+  case ATOMIC_DEFINED_WORD of
+    "$distinct"  => "$distinct"
+  | "$ite_f" => "$ite_f"
+  | "$less" => "$less"
+  | "$lesseq" => "$lesseq"
+  | "$greater" => "$greater"
+  | "$greatereq" => "$greatereq"
+  | "$is_int" => "$is_int"
+  | "$is_rat" => "$is_rat"
+  | thing => raise UNRECOGNISED_SYMBOL ("defined_pred", thing)
+)
+)
+ in ( LrTable.NT 40, ( result, ATOMIC_DEFINED_WORD1left, 
+ATOMIC_DEFINED_WORD1right), rest671)
+end
+|  ( 73, ( ( _, ( MlyValue.defined_plain_term defined_plain_term, 
+defined_plain_term1left, defined_plain_term1right)) :: rest671)) =>
+ let val  result = MlyValue.defined_plain_formula (
+( Pred defined_plain_term ))
+ in ( LrTable.NT 41, ( result, defined_plain_term1left, 
+defined_plain_term1right), rest671)
+end
+|  ( 74, ( ( _, ( MlyValue.defined_plain_formula defined_plain_formula
+, defined_plain_formula1left, defined_plain_formula1right)) :: rest671
+)) => let val  result = MlyValue.defined_atomic_formula (
+( defined_plain_formula ))
+ in ( LrTable.NT 42, ( result, defined_plain_formula1left, 
+defined_plain_formula1right), rest671)
+end
+|  ( 75, ( ( _, ( MlyValue.defined_infix_formula defined_infix_formula
+, defined_infix_formula1left, defined_infix_formula1right)) :: rest671
+)) => let val  result = MlyValue.defined_atomic_formula (
+( defined_infix_formula ))
+ in ( LrTable.NT 42, ( result, defined_infix_formula1left, 
+defined_infix_formula1right), rest671)
+end
+|  ( 76, ( ( _, ( MlyValue.plain_term plain_term, plain_term1left, 
+plain_term1right)) :: rest671)) => let val  result = 
+MlyValue.plain_atomic_formula (( Pred plain_term ))
+ in ( LrTable.NT 43, ( result, plain_term1left, plain_term1right), 
+rest671)
+end
+|  ( 77, ( ( _, ( MlyValue.plain_atomic_formula plain_atomic_formula, 
+plain_atomic_formula1left, plain_atomic_formula1right)) :: rest671))
+ => let val  result = MlyValue.atomic_formula (
+( plain_atomic_formula ))
+ in ( LrTable.NT 44, ( result, plain_atomic_formula1left, 
+plain_atomic_formula1right), rest671)
+end
+|  ( 78, ( ( _, ( MlyValue.defined_atomic_formula 
+defined_atomic_formula, defined_atomic_formula1left, 
+defined_atomic_formula1right)) :: rest671)) => let val  result = 
+MlyValue.atomic_formula (( defined_atomic_formula ))
+ in ( LrTable.NT 44, ( result, defined_atomic_formula1left, 
+defined_atomic_formula1right), rest671)
+end
+|  ( 79, ( ( _, ( MlyValue.system_atomic_formula system_atomic_formula
+, system_atomic_formula1left, system_atomic_formula1right)) :: rest671
+)) => let val  result = MlyValue.atomic_formula (
+( system_atomic_formula ))
+ in ( LrTable.NT 44, ( result, system_atomic_formula1left, 
+system_atomic_formula1right), rest671)
+end
+|  ( 80, ( ( _, ( _, VLINE1left, VLINE1right)) :: rest671)) => let
+ val  result = MlyValue.assoc_connective (( Interpreted_Logic Or ))
+ in ( LrTable.NT 48, ( result, VLINE1left, VLINE1right), rest671)
+end
+|  ( 81, ( ( _, ( _, AMPERSAND1left, AMPERSAND1right)) :: rest671)) =>
+ let val  result = MlyValue.assoc_connective (
+( Interpreted_Logic And ))
+ in ( LrTable.NT 48, ( result, AMPERSAND1left, AMPERSAND1right), 
+rest671)
+end
+|  ( 82, ( ( _, ( _, IFF1left, IFF1right)) :: rest671)) => let val  
+result = MlyValue.binary_connective (( Interpreted_Logic Iff ))
+ in ( LrTable.NT 49, ( result, IFF1left, IFF1right), rest671)
+end
+|  ( 83, ( ( _, ( _, IMPLIES1left, IMPLIES1right)) :: rest671)) => let
+ val  result = MlyValue.binary_connective (( Interpreted_Logic If ))
+ in ( LrTable.NT 49, ( result, IMPLIES1left, IMPLIES1right), rest671)
+
+end
+|  ( 84, ( ( _, ( _, IF1left, IF1right)) :: rest671)) => let val  
+result = MlyValue.binary_connective (( Interpreted_Logic Fi ))
+ in ( LrTable.NT 49, ( result, IF1left, IF1right), rest671)
+end
+|  ( 85, ( ( _, ( _, XOR1left, XOR1right)) :: rest671)) => let val  
+result = MlyValue.binary_connective (( Interpreted_Logic Xor ))
+ in ( LrTable.NT 49, ( result, XOR1left, XOR1right), rest671)
+end
+|  ( 86, ( ( _, ( _, NOR1left, NOR1right)) :: rest671)) => let val  
+result = MlyValue.binary_connective (( Interpreted_Logic Nor ))
+ in ( LrTable.NT 49, ( result, NOR1left, NOR1right), rest671)
+end
+|  ( 87, ( ( _, ( _, NAND1left, NAND1right)) :: rest671)) => let val  
+result = MlyValue.binary_connective (( Interpreted_Logic Nand ))
+ in ( LrTable.NT 49, ( result, NAND1left, NAND1right), rest671)
+end
+|  ( 88, ( ( _, ( _, EXCLAMATION1left, EXCLAMATION1right)) :: rest671)
+) => let val  result = MlyValue.fol_quantifier (( Forall ))
+ in ( LrTable.NT 50, ( result, EXCLAMATION1left, EXCLAMATION1right), 
+rest671)
+end
+|  ( 89, ( ( _, ( _, QUESTION1left, QUESTION1right)) :: rest671)) =>
+ let val  result = MlyValue.fol_quantifier (( Exists ))
+ in ( LrTable.NT 50, ( result, QUESTION1left, QUESTION1right), rest671
+)
+end
+|  ( 90, ( ( _, ( MlyValue.unary_connective unary_connective, 
+unary_connective1left, unary_connective1right)) :: rest671)) => let
+ val  result = MlyValue.thf_unary_connective (( unary_connective ))
+ in ( LrTable.NT 51, ( result, unary_connective1left, 
+unary_connective1right), rest671)
+end
+|  ( 91, ( ( _, ( _, OPERATOR_FORALL1left, OPERATOR_FORALL1right)) :: 
+rest671)) => let val  result = MlyValue.thf_unary_connective (
+( Interpreted_Logic Op_Forall ))
+ in ( LrTable.NT 51, ( result, OPERATOR_FORALL1left, 
+OPERATOR_FORALL1right), rest671)
+end
+|  ( 92, ( ( _, ( _, OPERATOR_EXISTS1left, OPERATOR_EXISTS1right)) :: 
+rest671)) => let val  result = MlyValue.thf_unary_connective (
+( Interpreted_Logic Op_Exists ))
+ in ( LrTable.NT 51, ( result, OPERATOR_EXISTS1left, 
+OPERATOR_EXISTS1right), rest671)
+end
+|  ( 93, ( ( _, ( MlyValue.infix_equality infix_equality, 
+infix_equality1left, infix_equality1right)) :: rest671)) => let val  
+result = MlyValue.thf_pair_connective (( infix_equality ))
+ in ( LrTable.NT 52, ( result, infix_equality1left, 
+infix_equality1right), rest671)
+end
+|  ( 94, ( ( _, ( MlyValue.infix_inequality infix_inequality, 
+infix_inequality1left, infix_inequality1right)) :: rest671)) => let
+ val  result = MlyValue.thf_pair_connective (( infix_inequality ))
+ in ( LrTable.NT 52, ( result, infix_inequality1left, 
+infix_inequality1right), rest671)
+end
+|  ( 95, ( ( _, ( MlyValue.binary_connective binary_connective, 
+binary_connective1left, binary_connective1right)) :: rest671)) => let
+ val  result = MlyValue.thf_pair_connective (( binary_connective ))
+ in ( LrTable.NT 52, ( result, binary_connective1left, 
+binary_connective1right), rest671)
+end
+|  ( 96, ( ( _, ( MlyValue.fol_quantifier fol_quantifier, 
+fol_quantifier1left, fol_quantifier1right)) :: rest671)) => let val  
+result = MlyValue.thf_quantifier (( fol_quantifier ))
+ in ( LrTable.NT 53, ( result, fol_quantifier1left, 
+fol_quantifier1right), rest671)
+end
+|  ( 97, ( ( _, ( _, CARET1left, CARET1right)) :: rest671)) => let
+ val  result = MlyValue.thf_quantifier (( Lambda ))
+ in ( LrTable.NT 53, ( result, CARET1left, CARET1right), rest671)
+end
+|  ( 98, ( ( _, ( _, DEP_PROD1left, DEP_PROD1right)) :: rest671)) =>
+ let val  result = MlyValue.thf_quantifier (( Dep_Prod ))
+ in ( LrTable.NT 53, ( result, DEP_PROD1left, DEP_PROD1right), rest671
+)
+end
+|  ( 99, ( ( _, ( _, DEP_SUM1left, DEP_SUM1right)) :: rest671)) => let
+ val  result = MlyValue.thf_quantifier (( Dep_Sum ))
+ in ( LrTable.NT 53, ( result, DEP_SUM1left, DEP_SUM1right), rest671)
+
+end
+|  ( 100, ( ( _, ( _, INDEF_CHOICE1left, INDEF_CHOICE1right)) :: 
+rest671)) => let val  result = MlyValue.thf_quantifier (( Epsilon ))
+ in ( LrTable.NT 53, ( result, INDEF_CHOICE1left, INDEF_CHOICE1right),
+ rest671)
+end
+|  ( 101, ( ( _, ( _, DEFIN_CHOICE1left, DEFIN_CHOICE1right)) :: 
+rest671)) => let val  result = MlyValue.thf_quantifier (( Iota ))
+ in ( LrTable.NT 53, ( result, DEFIN_CHOICE1left, DEFIN_CHOICE1right),
+ rest671)
+end
+|  ( 102, ( ( _, ( MlyValue.term term2, _, term2right)) :: ( _, ( 
+MlyValue.infix_inequality infix_inequality, _, _)) :: ( _, ( 
+MlyValue.term term1, term1left, _)) :: rest671)) => let val  result = 
+MlyValue.fol_infix_unary (( Pred (infix_inequality, [term1, term2]) ))
+ in ( LrTable.NT 54, ( result, term1left, term2right), rest671)
+end
+|  ( 103, ( ( _, ( MlyValue.thf_pair_connective thf_pair_connective, 
+thf_pair_connective1left, thf_pair_connective1right)) :: rest671)) =>
+ let val  result = MlyValue.thf_conn_term (( thf_pair_connective ))
+ in ( LrTable.NT 55, ( result, thf_pair_connective1left, 
+thf_pair_connective1right), rest671)
+end
+|  ( 104, ( ( _, ( MlyValue.assoc_connective assoc_connective, 
+assoc_connective1left, assoc_connective1right)) :: rest671)) => let
+ val  result = MlyValue.thf_conn_term (( assoc_connective ))
+ in ( LrTable.NT 55, ( result, assoc_connective1left, 
+assoc_connective1right), rest671)
+end
+|  ( 105, ( ( _, ( MlyValue.thf_unary_connective thf_unary_connective,
+ thf_unary_connective1left, thf_unary_connective1right)) :: rest671))
+ => let val  result = MlyValue.thf_conn_term (( thf_unary_connective )
+)
+ in ( LrTable.NT 55, ( result, thf_unary_connective1left, 
+thf_unary_connective1right), rest671)
+end
+|  ( 106, ( ( _, ( MlyValue.atomic_formula atomic_formula, 
+atomic_formula1left, atomic_formula1right)) :: rest671)) => let val  
+result = MlyValue.literal (( atomic_formula ))
+ in ( LrTable.NT 56, ( result, atomic_formula1left, 
+atomic_formula1right), rest671)
+end
+|  ( 107, ( ( _, ( MlyValue.atomic_formula atomic_formula, _, 
+atomic_formula1right)) :: ( _, ( _, TILDE1left, _)) :: rest671)) =>
+ let val  result = MlyValue.literal (
+( Fmla (Interpreted_Logic Not, [atomic_formula]) ))
+ in ( LrTable.NT 56, ( result, TILDE1left, atomic_formula1right), 
+rest671)
+end
+|  ( 108, ( ( _, ( MlyValue.fol_infix_unary fol_infix_unary, 
+fol_infix_unary1left, fol_infix_unary1right)) :: rest671)) => let val 
+ result = MlyValue.literal (( fol_infix_unary ))
+ in ( LrTable.NT 56, ( result, fol_infix_unary1left, 
+fol_infix_unary1right), rest671)
+end
+|  ( 109, ( ( _, ( MlyValue.literal literal, literal1left, 
+literal1right)) :: rest671)) => let val  result = MlyValue.disjunction
+ (( literal ))
+ in ( LrTable.NT 57, ( result, literal1left, literal1right), rest671)
+
+end
+|  ( 110, ( ( _, ( MlyValue.literal literal, _, literal1right)) :: _
+ :: ( _, ( MlyValue.disjunction disjunction, disjunction1left, _)) :: 
+rest671)) => let val  result = MlyValue.disjunction (
+( Fmla (Interpreted_Logic Or, [disjunction, literal]) ))
+ in ( LrTable.NT 57, ( result, disjunction1left, literal1right), 
+rest671)
+end
+|  ( 111, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.disjunction
+ disjunction, _, _)) :: ( _, ( _, LPAREN1left, _)) :: rest671)) => let
+ val  result = MlyValue.cnf_formula (( disjunction ))
+ in ( LrTable.NT 58, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 112, ( ( _, ( MlyValue.disjunction disjunction, disjunction1left,
+ disjunction1right)) :: rest671)) => let val  result = 
+MlyValue.cnf_formula (( disjunction ))
+ in ( LrTable.NT 58, ( result, disjunction1left, disjunction1right), 
+rest671)
+end
+|  ( 113, ( ( _, ( MlyValue.fof_logic_formula fof_logic_formula, 
+fof_logic_formula1left, fof_logic_formula1right)) :: rest671)) => let
+ val  result = MlyValue.fof_tuple_list (( [fof_logic_formula] ))
+ in ( LrTable.NT 59, ( result, fof_logic_formula1left, 
+fof_logic_formula1right), rest671)
+end
+|  ( 114, ( ( _, ( MlyValue.fof_tuple_list fof_tuple_list, _, 
+fof_tuple_list1right)) :: _ :: ( _, ( MlyValue.fof_logic_formula 
+fof_logic_formula, fof_logic_formula1left, _)) :: rest671)) => let
+ val  result = MlyValue.fof_tuple_list (
+( fof_logic_formula :: fof_tuple_list ))
+ in ( LrTable.NT 59, ( result, fof_logic_formula1left, 
+fof_tuple_list1right), rest671)
+end
+|  ( 115, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( _, LBRKT1left, _)) ::
+ rest671)) => let val  result = MlyValue.fof_tuple (( [] ))
+ in ( LrTable.NT 60, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 116, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( 
+MlyValue.fof_tuple_list fof_tuple_list, _, _)) :: ( _, ( _, LBRKT1left
+, _)) :: rest671)) => let val  result = MlyValue.fof_tuple (
+( fof_tuple_list ))
+ in ( LrTable.NT 60, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 117, ( ( _, ( MlyValue.fof_tuple fof_tuple2, _, fof_tuple2right))
+ :: _ :: ( _, ( MlyValue.fof_tuple fof_tuple1, fof_tuple1left, _)) :: 
+rest671)) => let val  result = MlyValue.fof_sequent (
+( Sequent (fof_tuple1, fof_tuple2) ))
+ in ( LrTable.NT 61, ( result, fof_tuple1left, fof_tuple2right), 
+rest671)
+end
+|  ( 118, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.fof_sequent
+ fof_sequent, _, _)) :: ( _, ( _, LPAREN1left, _)) :: rest671)) => let
+ val  result = MlyValue.fof_sequent (( fof_sequent ))
+ in ( LrTable.NT 61, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 119, ( ( _, ( _, TILDE1left, TILDE1right)) :: rest671)) => let
+ val  result = MlyValue.unary_connective (( Interpreted_Logic Not ))
+ in ( LrTable.NT 45, ( result, TILDE1left, TILDE1right), rest671)
+end
+|  ( 120, ( ( _, ( MlyValue.fof_unitary_formula fof_unitary_formula, _
+, fof_unitary_formula1right)) :: ( _, ( MlyValue.unary_connective 
+unary_connective, unary_connective1left, _)) :: rest671)) => let val  
+result = MlyValue.fof_unary_formula (
+( Fmla (unary_connective, [fof_unitary_formula]) ))
+ in ( LrTable.NT 62, ( result, unary_connective1left, 
+fof_unitary_formula1right), rest671)
+end
+|  ( 121, ( ( _, ( MlyValue.fol_infix_unary fol_infix_unary, 
+fol_infix_unary1left, fol_infix_unary1right)) :: rest671)) => let val 
+ result = MlyValue.fof_unary_formula (( fol_infix_unary ))
+ in ( LrTable.NT 62, ( result, fol_infix_unary1left, 
+fol_infix_unary1right), rest671)
+end
+|  ( 122, ( ( _, ( MlyValue.variable_ variable_, variable_1left, 
+variable_1right)) :: rest671)) => let val  result = 
+MlyValue.fof_variable_list (( [variable_] ))
+ in ( LrTable.NT 63, ( result, variable_1left, variable_1right), 
+rest671)
+end
+|  ( 123, ( ( _, ( MlyValue.fof_variable_list fof_variable_list, _, 
+fof_variable_list1right)) :: _ :: ( _, ( MlyValue.variable_ variable_,
+ variable_1left, _)) :: rest671)) => let val  result = 
+MlyValue.fof_variable_list (( variable_ :: fof_variable_list ))
+ in ( LrTable.NT 63, ( result, variable_1left, fof_variable_list1right
+), rest671)
+end
+|  ( 124, ( ( _, ( MlyValue.fof_unitary_formula fof_unitary_formula, _
+, fof_unitary_formula1right)) :: _ :: _ :: ( _, ( 
+MlyValue.fof_variable_list fof_variable_list, _, _)) :: _ :: ( _, ( 
+MlyValue.fol_quantifier fol_quantifier, fol_quantifier1left, _)) :: 
+rest671)) => let val  result = MlyValue.fof_quantified_formula (
+(
+  Quant (fol_quantifier, map (fn v => (v, NONE)) fof_variable_list, fof_unitary_formula)
+)
+)
+ in ( LrTable.NT 64, ( result, fol_quantifier1left, 
+fof_unitary_formula1right), rest671)
+end
+|  ( 125, ( ( _, ( MlyValue.fof_quantified_formula 
+fof_quantified_formula, fof_quantified_formula1left, 
+fof_quantified_formula1right)) :: rest671)) => let val  result = 
+MlyValue.fof_unitary_formula (( fof_quantified_formula ))
+ in ( LrTable.NT 65, ( result, fof_quantified_formula1left, 
+fof_quantified_formula1right), rest671)
+end
+|  ( 126, ( ( _, ( MlyValue.fof_unary_formula fof_unary_formula, 
+fof_unary_formula1left, fof_unary_formula1right)) :: rest671)) => let
+ val  result = MlyValue.fof_unitary_formula (( fof_unary_formula ))
+ in ( LrTable.NT 65, ( result, fof_unary_formula1left, 
+fof_unary_formula1right), rest671)
+end
+|  ( 127, ( ( _, ( MlyValue.atomic_formula atomic_formula, 
+atomic_formula1left, atomic_formula1right)) :: rest671)) => let val  
+result = MlyValue.fof_unitary_formula (( atomic_formula ))
+ in ( LrTable.NT 65, ( result, atomic_formula1left, 
+atomic_formula1right), rest671)
+end
+|  ( 128, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.fof_logic_formula fof_logic_formula, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.fof_unitary_formula (( fof_logic_formula ))
+ in ( LrTable.NT 65, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 129, ( ( _, ( MlyValue.fof_unitary_formula fof_unitary_formula2,
+ _, fof_unitary_formula2right)) :: _ :: ( _, ( 
+MlyValue.fof_unitary_formula fof_unitary_formula1, 
+fof_unitary_formula1left, _)) :: rest671)) => let val  result = 
+MlyValue.fof_and_formula (
+( Fmla (Interpreted_Logic And, [fof_unitary_formula1, fof_unitary_formula2]) )
+)
+ in ( LrTable.NT 66, ( result, fof_unitary_formula1left, 
+fof_unitary_formula2right), rest671)
+end
+|  ( 130, ( ( _, ( MlyValue.fof_unitary_formula fof_unitary_formula, _
+, fof_unitary_formula1right)) :: _ :: ( _, ( MlyValue.fof_and_formula 
+fof_and_formula, fof_and_formula1left, _)) :: rest671)) => let val  
+result = MlyValue.fof_and_formula (
+( Fmla (Interpreted_Logic And, [fof_and_formula, fof_unitary_formula]) )
+)
+ in ( LrTable.NT 66, ( result, fof_and_formula1left, 
+fof_unitary_formula1right), rest671)
+end
+|  ( 131, ( ( _, ( MlyValue.fof_unitary_formula fof_unitary_formula2,
+ _, fof_unitary_formula2right)) :: _ :: ( _, ( 
+MlyValue.fof_unitary_formula fof_unitary_formula1, 
+fof_unitary_formula1left, _)) :: rest671)) => let val  result = 
+MlyValue.fof_or_formula (
+( Fmla (Interpreted_Logic Or, [fof_unitary_formula1, fof_unitary_formula2]) )
+)
+ in ( LrTable.NT 67, ( result, fof_unitary_formula1left, 
+fof_unitary_formula2right), rest671)
+end
+|  ( 132, ( ( _, ( MlyValue.fof_unitary_formula fof_unitary_formula, _
+, fof_unitary_formula1right)) :: _ :: ( _, ( MlyValue.fof_or_formula 
+fof_or_formula, fof_or_formula1left, _)) :: rest671)) => let val  
+result = MlyValue.fof_or_formula (
+( Fmla (Interpreted_Logic Or, [fof_or_formula, fof_unitary_formula]) )
+)
+ in ( LrTable.NT 67, ( result, fof_or_formula1left, 
+fof_unitary_formula1right), rest671)
+end
+|  ( 133, ( ( _, ( MlyValue.fof_or_formula fof_or_formula, 
+fof_or_formula1left, fof_or_formula1right)) :: rest671)) => let val  
+result = MlyValue.fof_binary_assoc (( fof_or_formula ))
+ in ( LrTable.NT 68, ( result, fof_or_formula1left, 
+fof_or_formula1right), rest671)
+end
+|  ( 134, ( ( _, ( MlyValue.fof_and_formula fof_and_formula, 
+fof_and_formula1left, fof_and_formula1right)) :: rest671)) => let val 
+ result = MlyValue.fof_binary_assoc (( fof_and_formula ))
+ in ( LrTable.NT 68, ( result, fof_and_formula1left, 
+fof_and_formula1right), rest671)
+end
+|  ( 135, ( ( _, ( MlyValue.fof_unitary_formula fof_unitary_formula2,
+ _, fof_unitary_formula2right)) :: ( _, ( MlyValue.binary_connective 
+binary_connective, _, _)) :: ( _, ( MlyValue.fof_unitary_formula 
+fof_unitary_formula1, fof_unitary_formula1left, _)) :: rest671)) =>
+ let val  result = MlyValue.fof_binary_nonassoc (
+(
+  Fmla (binary_connective, [fof_unitary_formula1, fof_unitary_formula2] )
+)
+)
+ in ( LrTable.NT 69, ( result, fof_unitary_formula1left, 
+fof_unitary_formula2right), rest671)
+end
+|  ( 136, ( ( _, ( MlyValue.fof_binary_nonassoc fof_binary_nonassoc, 
+fof_binary_nonassoc1left, fof_binary_nonassoc1right)) :: rest671)) =>
+ let val  result = MlyValue.fof_binary_formula (
+( fof_binary_nonassoc ))
+ in ( LrTable.NT 70, ( result, fof_binary_nonassoc1left, 
+fof_binary_nonassoc1right), rest671)
+end
+|  ( 137, ( ( _, ( MlyValue.fof_binary_assoc fof_binary_assoc, 
+fof_binary_assoc1left, fof_binary_assoc1right)) :: rest671)) => let
+ val  result = MlyValue.fof_binary_formula (( fof_binary_assoc ))
+ in ( LrTable.NT 70, ( result, fof_binary_assoc1left, 
+fof_binary_assoc1right), rest671)
+end
+|  ( 138, ( ( _, ( MlyValue.fof_binary_formula fof_binary_formula, 
+fof_binary_formula1left, fof_binary_formula1right)) :: rest671)) =>
+ let val  result = MlyValue.fof_logic_formula (( fof_binary_formula ))
+ in ( LrTable.NT 71, ( result, fof_binary_formula1left, 
+fof_binary_formula1right), rest671)
+end
+|  ( 139, ( ( _, ( MlyValue.fof_unitary_formula fof_unitary_formula, 
+fof_unitary_formula1left, fof_unitary_formula1right)) :: rest671)) =>
+ let val  result = MlyValue.fof_logic_formula (( fof_unitary_formula )
+)
+ in ( LrTable.NT 71, ( result, fof_unitary_formula1left, 
+fof_unitary_formula1right), rest671)
+end
+|  ( 140, ( ( _, ( MlyValue.fof_logic_formula fof_logic_formula, 
+fof_logic_formula1left, fof_logic_formula1right)) :: rest671)) => let
+ val  result = MlyValue.fof_formula (( fof_logic_formula ))
+ in ( LrTable.NT 72, ( result, fof_logic_formula1left, 
+fof_logic_formula1right), rest671)
+end
+|  ( 141, ( ( _, ( MlyValue.fof_sequent fof_sequent, fof_sequent1left,
+ fof_sequent1right)) :: rest671)) => let val  result = 
+MlyValue.fof_formula (( fof_sequent ))
+ in ( LrTable.NT 72, ( result, fof_sequent1left, fof_sequent1right), 
+rest671)
+end
+|  ( 142, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( _, LBRKT1left, _)) ::
+ rest671)) => let val  result = MlyValue.tff_tuple (( [] ))
+ in ( LrTable.NT 73, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 143, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( 
+MlyValue.tff_tuple_list tff_tuple_list, _, _)) :: ( _, ( _, LBRKT1left
+, _)) :: rest671)) => let val  result = MlyValue.tff_tuple (
+( tff_tuple_list ))
+ in ( LrTable.NT 73, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 144, ( ( _, ( MlyValue.tff_tuple_list tff_tuple_list, _, 
+tff_tuple_list1right)) :: _ :: ( _, ( MlyValue.tff_logic_formula 
+tff_logic_formula, tff_logic_formula1left, _)) :: rest671)) => let
+ val  result = MlyValue.tff_tuple_list (
+( tff_logic_formula :: tff_tuple_list ))
+ in ( LrTable.NT 74, ( result, tff_logic_formula1left, 
+tff_tuple_list1right), rest671)
+end
+|  ( 145, ( ( _, ( MlyValue.tff_logic_formula tff_logic_formula, 
+tff_logic_formula1left, tff_logic_formula1right)) :: rest671)) => let
+ val  result = MlyValue.tff_tuple_list (( [tff_logic_formula] ))
+ in ( LrTable.NT 74, ( result, tff_logic_formula1left, 
+tff_logic_formula1right), rest671)
+end
+|  ( 146, ( ( _, ( MlyValue.tff_tuple tff_tuple2, _, tff_tuple2right))
+ :: _ :: ( _, ( MlyValue.tff_tuple tff_tuple1, tff_tuple1left, _)) :: 
+rest671)) => let val  result = MlyValue.tff_sequent (
+( Sequent (tff_tuple1, tff_tuple2) ))
+ in ( LrTable.NT 75, ( result, tff_tuple1left, tff_tuple2right), 
+rest671)
+end
+|  ( 147, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.tff_sequent
+ tff_sequent, _, _)) :: ( _, ( _, LPAREN1left, _)) :: rest671)) => let
+ val  result = MlyValue.tff_sequent (( tff_sequent ))
+ in ( LrTable.NT 75, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 148, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_logic_formula tff_logic_formula3, _, _)) :: _ :: ( _, ( 
+MlyValue.tff_logic_formula tff_logic_formula2, _, _)) :: _ :: ( _, ( 
+MlyValue.tff_logic_formula tff_logic_formula1, _, _)) :: _ :: ( _, ( _
+, ITE_F1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_conditional (
+(
+  Conditional (tff_logic_formula1, tff_logic_formula2, tff_logic_formula3)
+)
+)
+ in ( LrTable.NT 76, ( result, ITE_F1left, RPAREN1right), rest671)
+end
+|  ( 149, ( ( _, ( MlyValue.tff_logic_formula tff_logic_formula, _, 
+tff_logic_formula1right)) :: _ :: ( _, ( MlyValue.variable_ variable_,
+ variable_1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_defined_var (
+( Let_fmla ((variable_, NONE), tff_logic_formula) ))
+ in ( LrTable.NT 77, ( result, variable_1left, tff_logic_formula1right
+), rest671)
+end
+|  ( 150, ( ( _, ( MlyValue.term term, _, term1right)) :: _ :: ( _, ( 
+MlyValue.variable_ variable_, variable_1left, _)) :: rest671)) => let
+ val  result = MlyValue.tff_defined_var (
+( Let_term ((variable_, NONE), term) ))
+ in ( LrTable.NT 77, ( result, variable_1left, term1right), rest671)
+
+end
+|  ( 151, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_defined_var tff_defined_var, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_defined_var (( tff_defined_var ))
+ in ( LrTable.NT 77, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 152, ( ( _, ( MlyValue.tff_defined_var tff_defined_var, 
+tff_defined_var1left, tff_defined_var1right)) :: rest671)) => let val 
+ result = MlyValue.tff_let_list (( [tff_defined_var] ))
+ in ( LrTable.NT 78, ( result, tff_defined_var1left, 
+tff_defined_var1right), rest671)
+end
+|  ( 153, ( ( _, ( MlyValue.tff_let_list tff_let_list, _, 
+tff_let_list1right)) :: _ :: ( _, ( MlyValue.tff_defined_var 
+tff_defined_var, tff_defined_var1left, _)) :: rest671)) => let val  
+result = MlyValue.tff_let_list (( tff_defined_var :: tff_let_list ))
+ in ( LrTable.NT 78, ( result, tff_defined_var1left, 
+tff_let_list1right), rest671)
+end
+|  ( 154, ( ( _, ( MlyValue.tff_unitary_formula tff_unitary_formula, _
+, tff_unitary_formula1right)) :: _ :: _ :: ( _, ( 
+MlyValue.tff_let_list tff_let_list, _, _)) :: _ :: ( _, ( _, LET1left,
+ _)) :: rest671)) => let val  result = MlyValue.tptp_let (
+(
+  Let (tff_let_list, tff_unitary_formula)
+))
+ in ( LrTable.NT 79, ( result, LET1left, tff_unitary_formula1right), 
+rest671)
+end
+|  ( 155, ( ( _, ( MlyValue.tff_atomic_type tff_atomic_type2, _, 
+tff_atomic_type2right)) :: _ :: ( _, ( MlyValue.tff_atomic_type 
+tff_atomic_type1, tff_atomic_type1left, _)) :: rest671)) => let val  
+result = MlyValue.tff_xprod_type (
+( Prod_type(tff_atomic_type1, tff_atomic_type2) ))
+ in ( LrTable.NT 80, ( result, tff_atomic_type1left, 
+tff_atomic_type2right), rest671)
+end
+|  ( 156, ( ( _, ( MlyValue.tff_atomic_type tff_atomic_type, _, 
+tff_atomic_type1right)) :: _ :: ( _, ( MlyValue.tff_xprod_type 
+tff_xprod_type, tff_xprod_type1left, _)) :: rest671)) => let val  
+result = MlyValue.tff_xprod_type (
+( Prod_type(tff_xprod_type, tff_atomic_type) ))
+ in ( LrTable.NT 80, ( result, tff_xprod_type1left, 
+tff_atomic_type1right), rest671)
+end
+|  ( 157, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_xprod_type tff_xprod_type, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_xprod_type (( tff_xprod_type ))
+ in ( LrTable.NT 80, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 158, ( ( _, ( MlyValue.tff_atomic_type tff_atomic_type, _, 
+tff_atomic_type1right)) :: _ :: ( _, ( MlyValue.tff_unitary_type 
+tff_unitary_type, tff_unitary_type1left, _)) :: rest671)) => let val  
+result = MlyValue.tff_mapping_type (
+( Fn_type(tff_unitary_type, tff_atomic_type) ))
+ in ( LrTable.NT 81, ( result, tff_unitary_type1left, 
+tff_atomic_type1right), rest671)
+end
+|  ( 159, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_mapping_type tff_mapping_type, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_mapping_type (( tff_mapping_type ))
+ in ( LrTable.NT 81, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 160, ( ( _, ( MlyValue.atomic_word atomic_word, atomic_word1left,
+ atomic_word1right)) :: rest671)) => let val  result = 
+MlyValue.tff_atomic_type (( Atom_type atomic_word ))
+ in ( LrTable.NT 82, ( result, atomic_word1left, atomic_word1right), 
+rest671)
+end
+|  ( 161, ( ( _, ( MlyValue.defined_type defined_type, 
+defined_type1left, defined_type1right)) :: rest671)) => let val  
+result = MlyValue.tff_atomic_type (( Defined_type defined_type ))
+ in ( LrTable.NT 82, ( result, defined_type1left, defined_type1right),
+ rest671)
+end
+|  ( 162, ( ( _, ( MlyValue.tff_atomic_type tff_atomic_type, 
+tff_atomic_type1left, tff_atomic_type1right)) :: rest671)) => let val 
+ result = MlyValue.tff_unitary_type (( tff_atomic_type ))
+ in ( LrTable.NT 83, ( result, tff_atomic_type1left, 
+tff_atomic_type1right), rest671)
+end
+|  ( 163, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_xprod_type tff_xprod_type, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_unitary_type (( tff_xprod_type ))
+ in ( LrTable.NT 83, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 164, ( ( _, ( MlyValue.tff_atomic_type tff_atomic_type, 
+tff_atomic_type1left, tff_atomic_type1right)) :: rest671)) => let val 
+ result = MlyValue.tff_top_level_type (( tff_atomic_type ))
+ in ( LrTable.NT 84, ( result, tff_atomic_type1left, 
+tff_atomic_type1right), rest671)
+end
+|  ( 165, ( ( _, ( MlyValue.tff_mapping_type tff_mapping_type, 
+tff_mapping_type1left, tff_mapping_type1right)) :: rest671)) => let
+ val  result = MlyValue.tff_top_level_type (( tff_mapping_type ))
+ in ( LrTable.NT 84, ( result, tff_mapping_type1left, 
+tff_mapping_type1right), rest671)
+end
+|  ( 166, ( ( _, ( MlyValue.functor_ functor_, functor_1left, 
+functor_1right)) :: rest671)) => let val  result = 
+MlyValue.tff_untyped_atom (( (functor_, NONE) ))
+ in ( LrTable.NT 85, ( result, functor_1left, functor_1right), rest671
+)
+end
+|  ( 167, ( ( _, ( MlyValue.system_functor system_functor, 
+system_functor1left, system_functor1right)) :: rest671)) => let val  
+result = MlyValue.tff_untyped_atom (( (system_functor, NONE) ))
+ in ( LrTable.NT 85, ( result, system_functor1left, 
+system_functor1right), rest671)
+end
+|  ( 168, ( ( _, ( MlyValue.tff_top_level_type tff_top_level_type, _, 
+tff_top_level_type1right)) :: _ :: ( _, ( MlyValue.tff_untyped_atom 
+tff_untyped_atom, tff_untyped_atom1left, _)) :: rest671)) => let val  
+result = MlyValue.tff_typed_atom (
+( (fst tff_untyped_atom, SOME tff_top_level_type) ))
+ in ( LrTable.NT 86, ( result, tff_untyped_atom1left, 
+tff_top_level_type1right), rest671)
+end
+|  ( 169, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_typed_atom tff_typed_atom, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_typed_atom (( tff_typed_atom ))
+ in ( LrTable.NT 86, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 170, ( ( _, ( MlyValue.tff_unitary_formula tff_unitary_formula, _
+, tff_unitary_formula1right)) :: ( _, ( MlyValue.unary_connective 
+unary_connective, unary_connective1left, _)) :: rest671)) => let val  
+result = MlyValue.tff_unary_formula (
+( Fmla (unary_connective, [tff_unitary_formula]) ))
+ in ( LrTable.NT 87, ( result, unary_connective1left, 
+tff_unitary_formula1right), rest671)
+end
+|  ( 171, ( ( _, ( MlyValue.fol_infix_unary fol_infix_unary, 
+fol_infix_unary1left, fol_infix_unary1right)) :: rest671)) => let val 
+ result = MlyValue.tff_unary_formula (( fol_infix_unary ))
+ in ( LrTable.NT 87, ( result, fol_infix_unary1left, 
+fol_infix_unary1right), rest671)
+end
+|  ( 172, ( ( _, ( MlyValue.tff_atomic_type tff_atomic_type, _, 
+tff_atomic_type1right)) :: _ :: ( _, ( MlyValue.variable_ variable_, 
+variable_1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_typed_variable (( (variable_, SOME tff_atomic_type) ))
+ in ( LrTable.NT 88, ( result, variable_1left, tff_atomic_type1right),
+ rest671)
+end
+|  ( 173, ( ( _, ( MlyValue.tff_typed_variable tff_typed_variable, 
+tff_typed_variable1left, tff_typed_variable1right)) :: rest671)) =>
+ let val  result = MlyValue.tff_variable (( tff_typed_variable ))
+ in ( LrTable.NT 89, ( result, tff_typed_variable1left, 
+tff_typed_variable1right), rest671)
+end
+|  ( 174, ( ( _, ( MlyValue.variable_ variable_, variable_1left, 
+variable_1right)) :: rest671)) => let val  result = 
+MlyValue.tff_variable (( (variable_, NONE) ))
+ in ( LrTable.NT 89, ( result, variable_1left, variable_1right), 
+rest671)
+end
+|  ( 175, ( ( _, ( MlyValue.tff_variable tff_variable, 
+tff_variable1left, tff_variable1right)) :: rest671)) => let val  
+result = MlyValue.tff_variable_list (( [tff_variable] ))
+ in ( LrTable.NT 90, ( result, tff_variable1left, tff_variable1right),
+ rest671)
+end
+|  ( 176, ( ( _, ( MlyValue.tff_variable_list tff_variable_list, _, 
+tff_variable_list1right)) :: _ :: ( _, ( MlyValue.tff_variable 
+tff_variable, tff_variable1left, _)) :: rest671)) => let val  result =
+ MlyValue.tff_variable_list (( tff_variable :: tff_variable_list ))
+ in ( LrTable.NT 90, ( result, tff_variable1left, 
+tff_variable_list1right), rest671)
+end
+|  ( 177, ( ( _, ( MlyValue.tff_unitary_formula tff_unitary_formula, _
+, tff_unitary_formula1right)) :: _ :: _ :: ( _, ( 
+MlyValue.tff_variable_list tff_variable_list, _, _)) :: _ :: ( _, ( 
+MlyValue.fol_quantifier fol_quantifier, fol_quantifier1left, _)) :: 
+rest671)) => let val  result = MlyValue.tff_quantified_formula (
+(
+  Quant (fol_quantifier, tff_variable_list, tff_unitary_formula)
+))
+ in ( LrTable.NT 91, ( result, fol_quantifier1left, 
+tff_unitary_formula1right), rest671)
+end
+|  ( 178, ( ( _, ( MlyValue.tff_quantified_formula 
+tff_quantified_formula, tff_quantified_formula1left, 
+tff_quantified_formula1right)) :: rest671)) => let val  result = 
+MlyValue.tff_unitary_formula (( tff_quantified_formula ))
+ in ( LrTable.NT 92, ( result, tff_quantified_formula1left, 
+tff_quantified_formula1right), rest671)
+end
+|  ( 179, ( ( _, ( MlyValue.tff_unary_formula tff_unary_formula, 
+tff_unary_formula1left, tff_unary_formula1right)) :: rest671)) => let
+ val  result = MlyValue.tff_unitary_formula (( tff_unary_formula ))
+ in ( LrTable.NT 92, ( result, tff_unary_formula1left, 
+tff_unary_formula1right), rest671)
+end
+|  ( 180, ( ( _, ( MlyValue.atomic_formula atomic_formula, 
+atomic_formula1left, atomic_formula1right)) :: rest671)) => let val  
+result = MlyValue.tff_unitary_formula (( atomic_formula ))
+ in ( LrTable.NT 92, ( result, atomic_formula1left, 
+atomic_formula1right), rest671)
+end
+|  ( 181, ( ( _, ( MlyValue.tptp_let tptp_let, tptp_let1left, 
+tptp_let1right)) :: rest671)) => let val  result = 
+MlyValue.tff_unitary_formula (( tptp_let ))
+ in ( LrTable.NT 92, ( result, tptp_let1left, tptp_let1right), rest671
+)
+end
+|  ( 182, ( ( _, ( MlyValue.variable_ variable_, variable_1left, 
+variable_1right)) :: rest671)) => let val  result = 
+MlyValue.tff_unitary_formula (( Pred (Uninterpreted variable_, []) ))
+ in ( LrTable.NT 92, ( result, variable_1left, variable_1right), 
+rest671)
+end
+|  ( 183, ( ( _, ( MlyValue.tff_conditional tff_conditional, 
+tff_conditional1left, tff_conditional1right)) :: rest671)) => let val 
+ result = MlyValue.tff_unitary_formula (( tff_conditional ))
+ in ( LrTable.NT 92, ( result, tff_conditional1left, 
+tff_conditional1right), rest671)
+end
+|  ( 184, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_logic_formula tff_logic_formula, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_unitary_formula (( tff_logic_formula ))
+ in ( LrTable.NT 92, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 185, ( ( _, ( MlyValue.tff_unitary_formula tff_unitary_formula2,
+ _, tff_unitary_formula2right)) :: _ :: ( _, ( 
+MlyValue.tff_unitary_formula tff_unitary_formula1, 
+tff_unitary_formula1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_and_formula (
+( Fmla (Interpreted_Logic And, [tff_unitary_formula1, tff_unitary_formula2]) )
+)
+ in ( LrTable.NT 93, ( result, tff_unitary_formula1left, 
+tff_unitary_formula2right), rest671)
+end
+|  ( 186, ( ( _, ( MlyValue.tff_unitary_formula tff_unitary_formula, _
+, tff_unitary_formula1right)) :: _ :: ( _, ( MlyValue.tff_and_formula 
+tff_and_formula, tff_and_formula1left, _)) :: rest671)) => let val  
+result = MlyValue.tff_and_formula (
+( Fmla (Interpreted_Logic And, [tff_and_formula, tff_unitary_formula]) )
+)
+ in ( LrTable.NT 93, ( result, tff_and_formula1left, 
+tff_unitary_formula1right), rest671)
+end
+|  ( 187, ( ( _, ( MlyValue.tff_unitary_formula tff_unitary_formula2,
+ _, tff_unitary_formula2right)) :: _ :: ( _, ( 
+MlyValue.tff_unitary_formula tff_unitary_formula1, 
+tff_unitary_formula1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_or_formula (
+( Fmla (Interpreted_Logic Or, [tff_unitary_formula1, tff_unitary_formula2]) )
+)
+ in ( LrTable.NT 94, ( result, tff_unitary_formula1left, 
+tff_unitary_formula2right), rest671)
+end
+|  ( 188, ( ( _, ( MlyValue.tff_unitary_formula tff_unitary_formula, _
+, tff_unitary_formula1right)) :: _ :: ( _, ( MlyValue.tff_or_formula 
+tff_or_formula, tff_or_formula1left, _)) :: rest671)) => let val  
+result = MlyValue.tff_or_formula (
+( Fmla (Interpreted_Logic Or, [tff_or_formula, tff_unitary_formula]) )
+)
+ in ( LrTable.NT 94, ( result, tff_or_formula1left, 
+tff_unitary_formula1right), rest671)
+end
+|  ( 189, ( ( _, ( MlyValue.tff_or_formula tff_or_formula, 
+tff_or_formula1left, tff_or_formula1right)) :: rest671)) => let val  
+result = MlyValue.tff_binary_assoc (( tff_or_formula ))
+ in ( LrTable.NT 95, ( result, tff_or_formula1left, 
+tff_or_formula1right), rest671)
+end
+|  ( 190, ( ( _, ( MlyValue.tff_and_formula tff_and_formula, 
+tff_and_formula1left, tff_and_formula1right)) :: rest671)) => let val 
+ result = MlyValue.tff_binary_assoc (( tff_and_formula ))
+ in ( LrTable.NT 95, ( result, tff_and_formula1left, 
+tff_and_formula1right), rest671)
+end
+|  ( 191, ( ( _, ( MlyValue.tff_unitary_formula tff_unitary_formula2,
+ _, tff_unitary_formula2right)) :: ( _, ( MlyValue.binary_connective 
+binary_connective, _, _)) :: ( _, ( MlyValue.tff_unitary_formula 
+tff_unitary_formula1, tff_unitary_formula1left, _)) :: rest671)) =>
+ let val  result = MlyValue.tff_binary_nonassoc (
+( Fmla (binary_connective, [tff_unitary_formula1, tff_unitary_formula2]) )
+)
+ in ( LrTable.NT 96, ( result, tff_unitary_formula1left, 
+tff_unitary_formula2right), rest671)
+end
+|  ( 192, ( ( _, ( MlyValue.tff_binary_nonassoc tff_binary_nonassoc, 
+tff_binary_nonassoc1left, tff_binary_nonassoc1right)) :: rest671)) =>
+ let val  result = MlyValue.tff_binary_formula (
+( tff_binary_nonassoc ))
+ in ( LrTable.NT 97, ( result, tff_binary_nonassoc1left, 
+tff_binary_nonassoc1right), rest671)
+end
+|  ( 193, ( ( _, ( MlyValue.tff_binary_assoc tff_binary_assoc, 
+tff_binary_assoc1left, tff_binary_assoc1right)) :: rest671)) => let
+ val  result = MlyValue.tff_binary_formula (( tff_binary_assoc ))
+ in ( LrTable.NT 97, ( result, tff_binary_assoc1left, 
+tff_binary_assoc1right), rest671)
+end
+|  ( 194, ( ( _, ( MlyValue.tff_binary_formula tff_binary_formula, 
+tff_binary_formula1left, tff_binary_formula1right)) :: rest671)) =>
+ let val  result = MlyValue.tff_logic_formula (( tff_binary_formula ))
+ in ( LrTable.NT 98, ( result, tff_binary_formula1left, 
+tff_binary_formula1right), rest671)
+end
+|  ( 195, ( ( _, ( MlyValue.tff_unitary_formula tff_unitary_formula, 
+tff_unitary_formula1left, tff_unitary_formula1right)) :: rest671)) =>
+ let val  result = MlyValue.tff_logic_formula (( tff_unitary_formula )
+)
+ in ( LrTable.NT 98, ( result, tff_unitary_formula1left, 
+tff_unitary_formula1right), rest671)
+end
+|  ( 196, ( ( _, ( MlyValue.tff_logic_formula tff_logic_formula, 
+tff_logic_formula1left, tff_logic_formula1right)) :: rest671)) => let
+ val  result = MlyValue.tff_formula (( tff_logic_formula ))
+ in ( LrTable.NT 99, ( result, tff_logic_formula1left, 
+tff_logic_formula1right), rest671)
+end
+|  ( 197, ( ( _, ( MlyValue.tff_typed_atom tff_typed_atom, 
+tff_typed_atom1left, tff_typed_atom1right)) :: rest671)) => let val  
+result = MlyValue.tff_formula (
+( Atom (TFF_Typed_Atom tff_typed_atom) ))
+ in ( LrTable.NT 99, ( result, tff_typed_atom1left, 
+tff_typed_atom1right), rest671)
+end
+|  ( 198, ( ( _, ( MlyValue.tff_sequent tff_sequent, tff_sequent1left,
+ tff_sequent1right)) :: rest671)) => let val  result = 
+MlyValue.tff_formula (( tff_sequent ))
+ in ( LrTable.NT 99, ( result, tff_sequent1left, tff_sequent1right), 
+rest671)
+end
+|  ( 199, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( _, LBRKT1left, _)) ::
+ rest671)) => let val  result = MlyValue.thf_tuple (( [] ))
+ in ( LrTable.NT 100, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 200, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( 
+MlyValue.thf_tuple_list thf_tuple_list, _, _)) :: ( _, ( _, LBRKT1left
+, _)) :: rest671)) => let val  result = MlyValue.thf_tuple (
+( thf_tuple_list ))
+ in ( LrTable.NT 100, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 201, ( ( _, ( MlyValue.thf_logic_formula thf_logic_formula, 
+thf_logic_formula1left, thf_logic_formula1right)) :: rest671)) => let
+ val  result = MlyValue.thf_tuple_list (( [thf_logic_formula] ))
+ in ( LrTable.NT 101, ( result, thf_logic_formula1left, 
+thf_logic_formula1right), rest671)
+end
+|  ( 202, ( ( _, ( MlyValue.thf_tuple_list thf_tuple_list, _, 
+thf_tuple_list1right)) :: _ :: ( _, ( MlyValue.thf_logic_formula 
+thf_logic_formula, thf_logic_formula1left, _)) :: rest671)) => let
+ val  result = MlyValue.thf_tuple_list (
+( thf_logic_formula :: thf_tuple_list ))
+ in ( LrTable.NT 101, ( result, thf_logic_formula1left, 
+thf_tuple_list1right), rest671)
+end
+|  ( 203, ( ( _, ( MlyValue.thf_tuple thf_tuple2, _, thf_tuple2right))
+ :: _ :: ( _, ( MlyValue.thf_tuple thf_tuple1, thf_tuple1left, _)) :: 
+rest671)) => let val  result = MlyValue.thf_sequent (
+( Sequent(thf_tuple1, thf_tuple2) ))
+ in ( LrTable.NT 102, ( result, thf_tuple1left, thf_tuple2right), 
+rest671)
+end
+|  ( 204, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.thf_sequent
+ thf_sequent, _, _)) :: ( _, ( _, LPAREN1left, _)) :: rest671)) => let
+ val  result = MlyValue.thf_sequent (( thf_sequent ))
+ in ( LrTable.NT 102, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 205, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.thf_logic_formula thf_logic_formula3, _, _)) :: _ :: ( _, ( 
+MlyValue.thf_logic_formula thf_logic_formula2, _, _)) :: _ :: ( _, ( 
+MlyValue.thf_logic_formula thf_logic_formula1, _, _)) :: _ :: ( _, ( _
+, ITE_F1left, _)) :: rest671)) => let val  result = 
+MlyValue.thf_conditional (
+(
+  Conditional (thf_logic_formula1, thf_logic_formula2, thf_logic_formula3)
+)
+)
+ in ( LrTable.NT 103, ( result, ITE_F1left, RPAREN1right), rest671)
+
+end
+|  ( 206, ( ( _, ( MlyValue.thf_logic_formula thf_logic_formula, _, 
+thf_logic_formula1right)) :: _ :: ( _, ( MlyValue.thf_variable 
+thf_variable, thf_variable1left, _)) :: rest671)) => let val  result =
+ MlyValue.thf_defined_var (
+( Let_fmla (thf_variable, thf_logic_formula) ))
+ in ( LrTable.NT 104, ( result, thf_variable1left, 
+thf_logic_formula1right), rest671)
+end
+|  ( 207, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.thf_defined_var thf_defined_var, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.thf_defined_var (( thf_defined_var ))
+ in ( LrTable.NT 104, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 208, ( ( _, ( MlyValue.thf_defined_var thf_defined_var, 
+thf_defined_var1left, thf_defined_var1right)) :: rest671)) => let val 
+ result = MlyValue.thf_let_list (( [thf_defined_var] ))
+ in ( LrTable.NT 105, ( result, thf_defined_var1left, 
+thf_defined_var1right), rest671)
+end
+|  ( 209, ( ( _, ( MlyValue.thf_let_list thf_let_list, _, 
+thf_let_list1right)) :: _ :: ( _, ( MlyValue.thf_defined_var 
+thf_defined_var, thf_defined_var1left, _)) :: rest671)) => let val  
+result = MlyValue.thf_let_list (( thf_defined_var :: thf_let_list ))
+ in ( LrTable.NT 105, ( result, thf_defined_var1left, 
+thf_let_list1right), rest671)
+end
+|  ( 210, ( ( _, ( MlyValue.thf_unitary_formula thf_unitary_formula, _
+, thf_unitary_formula1right)) :: _ :: _ :: ( _, ( 
+MlyValue.thf_let_list thf_let_list, _, _)) :: _ :: ( _, ( _, LET1left,
+ _)) :: rest671)) => let val  result = MlyValue.thf_let (
+(
+  Let (thf_let_list, thf_unitary_formula)
+))
+ in ( LrTable.NT 106, ( result, LET1left, thf_unitary_formula1right), 
+rest671)
+end
+|  ( 211, ( ( _, ( MlyValue.term term, term1left, term1right)) :: 
+rest671)) => let val  result = MlyValue.thf_atom (
+( Atom (THF_Atom_term term) ))
+ in ( LrTable.NT 107, ( result, term1left, term1right), rest671)
+end
+|  ( 212, ( ( _, ( MlyValue.thf_conn_term thf_conn_term, 
+thf_conn_term1left, thf_conn_term1right)) :: rest671)) => let val  
+result = MlyValue.thf_atom (
+( Atom (THF_Atom_conn_term thf_conn_term) ))
+ in ( LrTable.NT 107, ( result, thf_conn_term1left, 
+thf_conn_term1right), rest671)
+end
+|  ( 213, ( ( _, ( MlyValue.thf_unitary_type thf_unitary_type2, _, 
+thf_unitary_type2right)) :: _ :: ( _, ( MlyValue.thf_unitary_type 
+thf_unitary_type1, thf_unitary_type1left, _)) :: rest671)) => let val 
+ result = MlyValue.thf_union_type (
+( Sum_type(thf_unitary_type1, thf_unitary_type2) ))
+ in ( LrTable.NT 108, ( result, thf_unitary_type1left, 
+thf_unitary_type2right), rest671)
+end
+|  ( 214, ( ( _, ( MlyValue.thf_unitary_type thf_unitary_type, _, 
+thf_unitary_type1right)) :: _ :: ( _, ( MlyValue.thf_union_type 
+thf_union_type, thf_union_type1left, _)) :: rest671)) => let val  
+result = MlyValue.thf_union_type (
+( Sum_type(thf_union_type, thf_unitary_type) ))
+ in ( LrTable.NT 108, ( result, thf_union_type1left, 
+thf_unitary_type1right), rest671)
+end
+|  ( 215, ( ( _, ( MlyValue.thf_unitary_type thf_unitary_type2, _, 
+thf_unitary_type2right)) :: _ :: ( _, ( MlyValue.thf_unitary_type 
+thf_unitary_type1, thf_unitary_type1left, _)) :: rest671)) => let val 
+ result = MlyValue.thf_xprod_type (
+( Prod_type(thf_unitary_type1, thf_unitary_type2) ))
+ in ( LrTable.NT 109, ( result, thf_unitary_type1left, 
+thf_unitary_type2right), rest671)
+end
+|  ( 216, ( ( _, ( MlyValue.thf_unitary_type thf_unitary_type, _, 
+thf_unitary_type1right)) :: _ :: ( _, ( MlyValue.thf_xprod_type 
+thf_xprod_type, thf_xprod_type1left, _)) :: rest671)) => let val  
+result = MlyValue.thf_xprod_type (
+( Prod_type(thf_xprod_type, thf_unitary_type) ))
+ in ( LrTable.NT 109, ( result, thf_xprod_type1left, 
+thf_unitary_type1right), rest671)
+end
+|  ( 217, ( ( _, ( MlyValue.thf_unitary_type thf_unitary_type2, _, 
+thf_unitary_type2right)) :: _ :: ( _, ( MlyValue.thf_unitary_type 
+thf_unitary_type1, thf_unitary_type1left, _)) :: rest671)) => let val 
+ result = MlyValue.thf_mapping_type (
+( Fn_type(thf_unitary_type1, thf_unitary_type2) ))
+ in ( LrTable.NT 110, ( result, thf_unitary_type1left, 
+thf_unitary_type2right), rest671)
+end
+|  ( 218, ( ( _, ( MlyValue.thf_mapping_type thf_mapping_type, _, 
+thf_mapping_type1right)) :: _ :: ( _, ( MlyValue.thf_unitary_type 
+thf_unitary_type, thf_unitary_type1left, _)) :: rest671)) => let val  
+result = MlyValue.thf_mapping_type (
+( Fn_type(thf_unitary_type, thf_mapping_type) ))
+ in ( LrTable.NT 110, ( result, thf_unitary_type1left, 
+thf_mapping_type1right), rest671)
+end
+|  ( 219, ( ( _, ( MlyValue.thf_mapping_type thf_mapping_type, 
+thf_mapping_type1left, thf_mapping_type1right)) :: rest671)) => let
+ val  result = MlyValue.thf_binary_type (( thf_mapping_type ))
+ in ( LrTable.NT 111, ( result, thf_mapping_type1left, 
+thf_mapping_type1right), rest671)
+end
+|  ( 220, ( ( _, ( MlyValue.thf_xprod_type thf_xprod_type, 
+thf_xprod_type1left, thf_xprod_type1right)) :: rest671)) => let val  
+result = MlyValue.thf_binary_type (( thf_xprod_type ))
+ in ( LrTable.NT 111, ( result, thf_xprod_type1left, 
+thf_xprod_type1right), rest671)
+end
+|  ( 221, ( ( _, ( MlyValue.thf_union_type thf_union_type, 
+thf_union_type1left, thf_union_type1right)) :: rest671)) => let val  
+result = MlyValue.thf_binary_type (( thf_union_type ))
+ in ( LrTable.NT 111, ( result, thf_union_type1left, 
+thf_union_type1right), rest671)
+end
+|  ( 222, ( ( _, ( MlyValue.thf_unitary_formula thf_unitary_formula, 
+thf_unitary_formula1left, thf_unitary_formula1right)) :: rest671)) =>
+ let val  result = MlyValue.thf_unitary_type (
+( Fmla_type thf_unitary_formula ))
+ in ( LrTable.NT 112, ( result, thf_unitary_formula1left, 
+thf_unitary_formula1right), rest671)
+end
+|  ( 223, ( ( _, ( MlyValue.thf_logic_formula thf_logic_formula, 
+thf_logic_formula1left, thf_logic_formula1right)) :: rest671)) => let
+ val  result = MlyValue.thf_top_level_type (
+( Fmla_type thf_logic_formula ))
+ in ( LrTable.NT 113, ( result, thf_logic_formula1left, 
+thf_logic_formula1right), rest671)
+end
+|  ( 224, ( ( _, ( MlyValue.constant constant2, _, constant2right)) ::
+ _ :: ( _, ( MlyValue.constant constant1, constant1left, _)) :: 
+rest671)) => let val  result = MlyValue.thf_subtype (
+( Subtype(constant1, constant2) ))
+ in ( LrTable.NT 114, ( result, constant1left, constant2right), 
+rest671)
+end
+|  ( 225, ( ( _, ( MlyValue.thf_atom thf_atom, thf_atom1left, 
+thf_atom1right)) :: rest671)) => let val  result = 
+MlyValue.thf_typeable_formula (( thf_atom ))
+ in ( LrTable.NT 115, ( result, thf_atom1left, thf_atom1right), 
+rest671)
+end
+|  ( 226, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.thf_logic_formula thf_logic_formula, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.thf_typeable_formula (( thf_logic_formula ))
+ in ( LrTable.NT 115, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 227, ( ( _, ( MlyValue.thf_top_level_type thf_top_level_type, _, 
+thf_top_level_type1right)) :: _ :: ( _, ( 
+MlyValue.thf_typeable_formula thf_typeable_formula, 
+thf_typeable_formula1left, _)) :: rest671)) => let val  result = 
+MlyValue.thf_type_formula (
+( (thf_typeable_formula, thf_top_level_type) ))
+ in ( LrTable.NT 116, ( result, thf_typeable_formula1left, 
+thf_top_level_type1right), rest671)
+end
+|  ( 228, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.thf_logic_formula thf_logic_formula, _, _)) :: _ :: ( _, ( 
+MlyValue.thf_unary_connective thf_unary_connective, 
+thf_unary_connective1left, _)) :: rest671)) => let val  result = 
+MlyValue.thf_unary_formula (
+(
+  Fmla (thf_unary_connective, [thf_logic_formula])
+))
+ in ( LrTable.NT 117, ( result, thf_unary_connective1left, 
+RPAREN1right), rest671)
+end
+|  ( 229, ( ( _, ( MlyValue.thf_top_level_type thf_top_level_type, _, 
+thf_top_level_type1right)) :: _ :: ( _, ( MlyValue.variable_ variable_
+, variable_1left, _)) :: rest671)) => let val  result = 
+MlyValue.thf_typed_variable (( (variable_, SOME thf_top_level_type) ))
+ in ( LrTable.NT 118, ( result, variable_1left, 
+thf_top_level_type1right), rest671)
+end
+|  ( 230, ( ( _, ( MlyValue.thf_typed_variable thf_typed_variable, 
+thf_typed_variable1left, thf_typed_variable1right)) :: rest671)) =>
+ let val  result = MlyValue.thf_variable (( thf_typed_variable ))
+ in ( LrTable.NT 119, ( result, thf_typed_variable1left, 
+thf_typed_variable1right), rest671)
+end
+|  ( 231, ( ( _, ( MlyValue.variable_ variable_, variable_1left, 
+variable_1right)) :: rest671)) => let val  result = 
+MlyValue.thf_variable (( (variable_, NONE) ))
+ in ( LrTable.NT 119, ( result, variable_1left, variable_1right), 
+rest671)
+end
+|  ( 232, ( ( _, ( MlyValue.thf_variable thf_variable, 
+thf_variable1left, thf_variable1right)) :: rest671)) => let val  
+result = MlyValue.thf_variable_list (( [thf_variable] ))
+ in ( LrTable.NT 120, ( result, thf_variable1left, thf_variable1right)
+, rest671)
+end
+|  ( 233, ( ( _, ( MlyValue.thf_variable_list thf_variable_list, _, 
+thf_variable_list1right)) :: _ :: ( _, ( MlyValue.thf_variable 
+thf_variable, thf_variable1left, _)) :: rest671)) => let val  result =
+ MlyValue.thf_variable_list (( thf_variable :: thf_variable_list ))
+ in ( LrTable.NT 120, ( result, thf_variable1left, 
+thf_variable_list1right), rest671)
+end
+|  ( 234, ( ( _, ( MlyValue.thf_unitary_formula thf_unitary_formula, _
+, thf_unitary_formula1right)) :: _ :: _ :: ( _, ( 
+MlyValue.thf_variable_list thf_variable_list, _, _)) :: _ :: ( _, ( 
+MlyValue.thf_quantifier thf_quantifier, thf_quantifier1left, _)) :: 
+rest671)) => let val  result = MlyValue.thf_quantified_formula (
+(
+  Quant (thf_quantifier, thf_variable_list, thf_unitary_formula)
+))
+ in ( LrTable.NT 121, ( result, thf_quantifier1left, 
+thf_unitary_formula1right), rest671)
+end
+|  ( 235, ( ( _, ( MlyValue.thf_quantified_formula 
+thf_quantified_formula, thf_quantified_formula1left, 
+thf_quantified_formula1right)) :: rest671)) => let val  result = 
+MlyValue.thf_unitary_formula (( thf_quantified_formula ))
+ in ( LrTable.NT 122, ( result, thf_quantified_formula1left, 
+thf_quantified_formula1right), rest671)
+end
+|  ( 236, ( ( _, ( MlyValue.thf_unary_formula thf_unary_formula, 
+thf_unary_formula1left, thf_unary_formula1right)) :: rest671)) => let
+ val  result = MlyValue.thf_unitary_formula (( thf_unary_formula ))
+ in ( LrTable.NT 122, ( result, thf_unary_formula1left, 
+thf_unary_formula1right), rest671)
+end
+|  ( 237, ( ( _, ( MlyValue.thf_atom thf_atom, thf_atom1left, 
+thf_atom1right)) :: rest671)) => let val  result = 
+MlyValue.thf_unitary_formula (( thf_atom ))
+ in ( LrTable.NT 122, ( result, thf_atom1left, thf_atom1right), 
+rest671)
+end
+|  ( 238, ( ( _, ( MlyValue.thf_let thf_let, thf_let1left, 
+thf_let1right)) :: rest671)) => let val  result = 
+MlyValue.thf_unitary_formula (( thf_let ))
+ in ( LrTable.NT 122, ( result, thf_let1left, thf_let1right), rest671)
+
+end
+|  ( 239, ( ( _, ( MlyValue.thf_conditional thf_conditional, 
+thf_conditional1left, thf_conditional1right)) :: rest671)) => let val 
+ result = MlyValue.thf_unitary_formula (( thf_conditional ))
+ in ( LrTable.NT 122, ( result, thf_conditional1left, 
+thf_conditional1right), rest671)
+end
+|  ( 240, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.thf_logic_formula thf_logic_formula, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.thf_unitary_formula (( thf_logic_formula ))
+ in ( LrTable.NT 122, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 241, ( ( _, ( MlyValue.thf_unitary_formula thf_unitary_formula2,
+ _, thf_unitary_formula2right)) :: _ :: ( _, ( 
+MlyValue.thf_unitary_formula thf_unitary_formula1, 
+thf_unitary_formula1left, _)) :: rest671)) => let val  result = 
+MlyValue.thf_apply_formula (
+( Fmla (Interpreted_ExtraLogic Apply, [thf_unitary_formula1, thf_unitary_formula2]) )
+)
+ in ( LrTable.NT 123, ( result, thf_unitary_formula1left, 
+thf_unitary_formula2right), rest671)
+end
+|  ( 242, ( ( _, ( MlyValue.thf_unitary_formula thf_unitary_formula, _
+, thf_unitary_formula1right)) :: _ :: ( _, ( 
+MlyValue.thf_apply_formula thf_apply_formula, thf_apply_formula1left,
+ _)) :: rest671)) => let val  result = MlyValue.thf_apply_formula (
+( Fmla (Interpreted_ExtraLogic Apply, [thf_apply_formula, thf_unitary_formula]) )
+)
+ in ( LrTable.NT 123, ( result, thf_apply_formula1left, 
+thf_unitary_formula1right), rest671)
+end
+|  ( 243, ( ( _, ( MlyValue.thf_unitary_formula thf_unitary_formula2,
+ _, thf_unitary_formula2right)) :: _ :: ( _, ( 
+MlyValue.thf_unitary_formula thf_unitary_formula1, 
+thf_unitary_formula1left, _)) :: rest671)) => let val  result = 
+MlyValue.thf_and_formula (
+( Fmla (Interpreted_Logic And, [thf_unitary_formula1, thf_unitary_formula2]) )
+)
+ in ( LrTable.NT 124, ( result, thf_unitary_formula1left, 
+thf_unitary_formula2right), rest671)
+end
+|  ( 244, ( ( _, ( MlyValue.thf_unitary_formula thf_unitary_formula, _
+, thf_unitary_formula1right)) :: _ :: ( _, ( MlyValue.thf_and_formula 
+thf_and_formula, thf_and_formula1left, _)) :: rest671)) => let val  
+result = MlyValue.thf_and_formula (
+( Fmla (Interpreted_Logic And, [thf_and_formula, thf_unitary_formula]) )
+)
+ in ( LrTable.NT 124, ( result, thf_and_formula1left, 
+thf_unitary_formula1right), rest671)
+end
+|  ( 245, ( ( _, ( MlyValue.thf_unitary_formula thf_unitary_formula2,
+ _, thf_unitary_formula2right)) :: _ :: ( _, ( 
+MlyValue.thf_unitary_formula thf_unitary_formula1, 
+thf_unitary_formula1left, _)) :: rest671)) => let val  result = 
+MlyValue.thf_or_formula (
+( Fmla (Interpreted_Logic Or, [thf_unitary_formula1, thf_unitary_formula2]) )
+)
+ in ( LrTable.NT 125, ( result, thf_unitary_formula1left, 
+thf_unitary_formula2right), rest671)
+end
+|  ( 246, ( ( _, ( MlyValue.thf_unitary_formula thf_unitary_formula, _
+, thf_unitary_formula1right)) :: _ :: ( _, ( MlyValue.thf_or_formula 
+thf_or_formula, thf_or_formula1left, _)) :: rest671)) => let val  
+result = MlyValue.thf_or_formula (
+( Fmla (Interpreted_Logic Or, [thf_or_formula, thf_unitary_formula]) )
+)
+ in ( LrTable.NT 125, ( result, thf_or_formula1left, 
+thf_unitary_formula1right), rest671)
+end
+|  ( 247, ( ( _, ( MlyValue.thf_or_formula thf_or_formula, 
+thf_or_formula1left, thf_or_formula1right)) :: rest671)) => let val  
+result = MlyValue.thf_binary_tuple (( thf_or_formula ))
+ in ( LrTable.NT 126, ( result, thf_or_formula1left, 
+thf_or_formula1right), rest671)
+end
+|  ( 248, ( ( _, ( MlyValue.thf_and_formula thf_and_formula, 
+thf_and_formula1left, thf_and_formula1right)) :: rest671)) => let val 
+ result = MlyValue.thf_binary_tuple (( thf_and_formula ))
+ in ( LrTable.NT 126, ( result, thf_and_formula1left, 
+thf_and_formula1right), rest671)
+end
+|  ( 249, ( ( _, ( MlyValue.thf_apply_formula thf_apply_formula, 
+thf_apply_formula1left, thf_apply_formula1right)) :: rest671)) => let
+ val  result = MlyValue.thf_binary_tuple (( thf_apply_formula ))
+ in ( LrTable.NT 126, ( result, thf_apply_formula1left, 
+thf_apply_formula1right), rest671)
+end
+|  ( 250, ( ( _, ( MlyValue.thf_unitary_formula thf_unitary_formula2,
+ _, thf_unitary_formula2right)) :: ( _, ( MlyValue.thf_pair_connective
+ thf_pair_connective, _, _)) :: ( _, ( MlyValue.thf_unitary_formula 
+thf_unitary_formula1, thf_unitary_formula1left, _)) :: rest671)) =>
+ let val  result = MlyValue.thf_binary_pair (
+(
+  Fmla (thf_pair_connective, [thf_unitary_formula1, thf_unitary_formula2])
+)
+)
+ in ( LrTable.NT 127, ( result, thf_unitary_formula1left, 
+thf_unitary_formula2right), rest671)
+end
+|  ( 251, ( ( _, ( MlyValue.thf_binary_pair thf_binary_pair, 
+thf_binary_pair1left, thf_binary_pair1right)) :: rest671)) => let val 
+ result = MlyValue.thf_binary_formula (( thf_binary_pair ))
+ in ( LrTable.NT 128, ( result, thf_binary_pair1left, 
+thf_binary_pair1right), rest671)
+end
+|  ( 252, ( ( _, ( MlyValue.thf_binary_tuple thf_binary_tuple, 
+thf_binary_tuple1left, thf_binary_tuple1right)) :: rest671)) => let
+ val  result = MlyValue.thf_binary_formula (( thf_binary_tuple ))
+ in ( LrTable.NT 128, ( result, thf_binary_tuple1left, 
+thf_binary_tuple1right), rest671)
+end
+|  ( 253, ( ( _, ( MlyValue.thf_binary_type thf_binary_type, 
+thf_binary_type1left, thf_binary_type1right)) :: rest671)) => let val 
+ result = MlyValue.thf_binary_formula (( THF_type thf_binary_type ))
+ in ( LrTable.NT 128, ( result, thf_binary_type1left, 
+thf_binary_type1right), rest671)
+end
+|  ( 254, ( ( _, ( MlyValue.thf_binary_formula thf_binary_formula, 
+thf_binary_formula1left, thf_binary_formula1right)) :: rest671)) =>
+ let val  result = MlyValue.thf_logic_formula (( thf_binary_formula ))
+ in ( LrTable.NT 129, ( result, thf_binary_formula1left, 
+thf_binary_formula1right), rest671)
+end
+|  ( 255, ( ( _, ( MlyValue.thf_unitary_formula thf_unitary_formula, 
+thf_unitary_formula1left, thf_unitary_formula1right)) :: rest671)) =>
+ let val  result = MlyValue.thf_logic_formula (( thf_unitary_formula )
+)
+ in ( LrTable.NT 129, ( result, thf_unitary_formula1left, 
+thf_unitary_formula1right), rest671)
+end
+|  ( 256, ( ( _, ( MlyValue.thf_type_formula thf_type_formula, 
+thf_type_formula1left, thf_type_formula1right)) :: rest671)) => let
+ val  result = MlyValue.thf_logic_formula (
+( THF_typing thf_type_formula ))
+ in ( LrTable.NT 129, ( result, thf_type_formula1left, 
+thf_type_formula1right), rest671)
+end
+|  ( 257, ( ( _, ( MlyValue.thf_subtype thf_subtype, thf_subtype1left,
+ thf_subtype1right)) :: rest671)) => let val  result = 
+MlyValue.thf_logic_formula (( THF_type thf_subtype ))
+ in ( LrTable.NT 129, ( result, thf_subtype1left, thf_subtype1right), 
+rest671)
+end
+|  ( 258, ( ( _, ( MlyValue.thf_logic_formula thf_logic_formula, 
+thf_logic_formula1left, thf_logic_formula1right)) :: rest671)) => let
+ val  result = MlyValue.thf_formula (( thf_logic_formula ))
+ in ( LrTable.NT 130, ( result, thf_logic_formula1left, 
+thf_logic_formula1right), rest671)
+end
+|  ( 259, ( ( _, ( MlyValue.thf_sequent thf_sequent, thf_sequent1left,
+ thf_sequent1right)) :: rest671)) => let val  result = 
+MlyValue.thf_formula (( thf_sequent ))
+ in ( LrTable.NT 130, ( result, thf_sequent1left, thf_sequent1right), 
+rest671)
+end
+|  ( 260, ( ( _, ( MlyValue.LOWER_WORD LOWER_WORD, LOWER_WORD1left, 
+LOWER_WORD1right)) :: rest671)) => let val  result = 
+MlyValue.formula_role (( classify_role LOWER_WORD ))
+ in ( LrTable.NT 131, ( result, LOWER_WORD1left, LOWER_WORD1right), 
+rest671)
+end
+|  ( 261, ( ( _, ( _, _, PERIOD1right)) :: _ :: ( _, ( 
+MlyValue.annotations annotations, _, _)) :: ( _, ( 
+MlyValue.thf_formula thf_formula, _, _)) :: _ :: ( _, ( 
+MlyValue.formula_role formula_role, _, _)) :: _ :: ( _, ( 
+MlyValue.name name, _, _)) :: _ :: ( _, ( _, (THFleft as THF1left), 
+THFright)) :: rest671)) => let val  result = MlyValue.thf_annotated (
+(
+  Annotated_Formula ((file_name, THFleft + 1, THFright + 1),
+   THF, name, formula_role, thf_formula, annotations)
+)
+)
+ in ( LrTable.NT 135, ( result, THF1left, PERIOD1right), rest671)
+end
+|  ( 262, ( ( _, ( _, _, PERIOD1right)) :: _ :: ( _, ( 
+MlyValue.annotations annotations, _, _)) :: ( _, ( 
+MlyValue.tff_formula tff_formula, _, _)) :: _ :: ( _, ( 
+MlyValue.formula_role formula_role, _, _)) :: _ :: ( _, ( 
+MlyValue.name name, _, _)) :: _ :: ( _, ( _, (TFFleft as TFF1left), 
+TFFright)) :: rest671)) => let val  result = MlyValue.tff_annotated (
+(
+  Annotated_Formula ((file_name, TFFleft + 1, TFFright + 1),
+   TFF, name, formula_role, tff_formula, annotations)
+)
+)
+ in ( LrTable.NT 134, ( result, TFF1left, PERIOD1right), rest671)
+end
+|  ( 263, ( ( _, ( _, _, PERIOD1right)) :: _ :: ( _, ( 
+MlyValue.annotations annotations, _, _)) :: ( _, ( 
+MlyValue.fof_formula fof_formula, _, _)) :: _ :: ( _, ( 
+MlyValue.formula_role formula_role, _, _)) :: _ :: ( _, ( 
+MlyValue.name name, _, _)) :: _ :: ( _, ( _, (FOFleft as FOF1left), 
+FOFright)) :: rest671)) => let val  result = MlyValue.fof_annotated (
+(
+  Annotated_Formula ((file_name, FOFleft + 1, FOFright + 1),
+   FOF, name, formula_role, fof_formula, annotations)
+)
+)
+ in ( LrTable.NT 133, ( result, FOF1left, PERIOD1right), rest671)
+end
+|  ( 264, ( ( _, ( _, _, PERIOD1right)) :: _ :: ( _, ( 
+MlyValue.annotations annotations, _, _)) :: ( _, ( 
+MlyValue.cnf_formula cnf_formula, _, _)) :: _ :: ( _, ( 
+MlyValue.formula_role formula_role, _, _)) :: _ :: ( _, ( 
+MlyValue.name name, _, _)) :: _ :: ( _, ( _, (CNFleft as CNF1left), 
+CNFright)) :: rest671)) => let val  result = MlyValue.cnf_annotated (
+(
+  Annotated_Formula ((file_name, CNFleft + 1, CNFright + 1),
+   CNF, name, formula_role, cnf_formula, annotations)
+)
+)
+ in ( LrTable.NT 132, ( result, CNF1left, PERIOD1right), rest671)
+end
+|  ( 265, ( ( _, ( MlyValue.cnf_annotated cnf_annotated, 
+cnf_annotated1left, cnf_annotated1right)) :: rest671)) => let val  
+result = MlyValue.annotated_formula (( cnf_annotated ))
+ in ( LrTable.NT 136, ( result, cnf_annotated1left, 
+cnf_annotated1right), rest671)
+end
+|  ( 266, ( ( _, ( MlyValue.fof_annotated fof_annotated, 
+fof_annotated1left, fof_annotated1right)) :: rest671)) => let val  
+result = MlyValue.annotated_formula (( fof_annotated ))
+ in ( LrTable.NT 136, ( result, fof_annotated1left, 
+fof_annotated1right), rest671)
+end
+|  ( 267, ( ( _, ( MlyValue.tff_annotated tff_annotated, 
+tff_annotated1left, tff_annotated1right)) :: rest671)) => let val  
+result = MlyValue.annotated_formula (( tff_annotated ))
+ in ( LrTable.NT 136, ( result, tff_annotated1left, 
+tff_annotated1right), rest671)
+end
+|  ( 268, ( ( _, ( MlyValue.thf_annotated thf_annotated, 
+thf_annotated1left, thf_annotated1right)) :: rest671)) => let val  
+result = MlyValue.annotated_formula (( thf_annotated ))
+ in ( LrTable.NT 136, ( result, thf_annotated1left, 
+thf_annotated1right), rest671)
+end
+|  ( 269, ( ( _, ( _, _, PERIOD1right)) :: _ :: ( _, ( 
+MlyValue.formula_selection formula_selection, _, _)) :: ( _, ( 
+MlyValue.file_name file_name, _, _)) :: _ :: ( _, ( _, INCLUDE1left, _
+)) :: rest671)) => let val  result = MlyValue.include_ (
+(
+  Include (file_name, formula_selection)
+))
+ in ( LrTable.NT 137, ( result, INCLUDE1left, PERIOD1right), rest671)
+
+end
+|  ( 270, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( MlyValue.name_list 
+name_list, _, _)) :: _ :: ( _, ( _, COMMA1left, _)) :: rest671)) =>
+ let val  result = MlyValue.formula_selection (( name_list  ))
+ in ( LrTable.NT 3, ( result, COMMA1left, RBRKT1right), rest671)
+end
+|  ( 271, ( rest671)) => let val  result = MlyValue.formula_selection
+ (( [] ))
+ in ( LrTable.NT 3, ( result, defaultPos, defaultPos), rest671)
+end
+|  ( 272, ( ( _, ( MlyValue.name_list name_list, _, name_list1right))
+ :: _ :: ( _, ( MlyValue.name name, name1left, _)) :: rest671)) => let
+ val  result = MlyValue.name_list (( name :: name_list ))
+ in ( LrTable.NT 2, ( result, name1left, name_list1right), rest671)
+
+end
+|  ( 273, ( ( _, ( MlyValue.name name, name1left, name1right)) :: 
+rest671)) => let val  result = MlyValue.name_list (( [name] ))
+ in ( LrTable.NT 2, ( result, name1left, name1right), rest671)
+end
+|  ( 274, ( ( _, ( MlyValue.atomic_word atomic_word, atomic_word1left,
+ atomic_word1right)) :: rest671)) => let val  result = MlyValue.name (
+( atomic_word ))
+ in ( LrTable.NT 1, ( result, atomic_word1left, atomic_word1right), 
+rest671)
+end
+|  ( 275, ( ( _, ( MlyValue.integer integer, integer1left, 
+integer1right)) :: rest671)) => let val  result = MlyValue.name (
+( integer ))
+ in ( LrTable.NT 1, ( result, integer1left, integer1right), rest671)
+
+end
+|  ( 276, ( ( _, ( MlyValue.annotated_formula annotated_formula, 
+annotated_formula1left, annotated_formula1right)) :: rest671)) => let
+ val  result = MlyValue.tptp_input (( annotated_formula ))
+ in ( LrTable.NT 138, ( result, annotated_formula1left, 
+annotated_formula1right), rest671)
+end
+|  ( 277, ( ( _, ( MlyValue.include_ include_, include_1left, 
+include_1right)) :: rest671)) => let val  result = MlyValue.tptp_input
+ (( include_ ))
+ in ( LrTable.NT 138, ( result, include_1left, include_1right), 
+rest671)
+end
+|  ( 278, ( ( _, ( MlyValue.tptp_file tptp_file, _, tptp_file1right))
+ :: ( _, ( MlyValue.tptp_input tptp_input, tptp_input1left, _)) :: 
+rest671)) => let val  result = MlyValue.tptp_file (
+( tptp_input :: tptp_file ))
+ in ( LrTable.NT 139, ( result, tptp_input1left, tptp_file1right), 
+rest671)
+end
+|  ( 279, ( ( _, ( MlyValue.tptp_file tptp_file, _, tptp_file1right))
+ :: ( _, ( _, COMMENT1left, _)) :: rest671)) => let val  result = 
+MlyValue.tptp_file (( tptp_file ))
+ in ( LrTable.NT 139, ( result, COMMENT1left, tptp_file1right), 
+rest671)
+end
+|  ( 280, ( rest671)) => let val  result = MlyValue.tptp_file (( [] ))
+ in ( LrTable.NT 139, ( result, defaultPos, defaultPos), rest671)
+end
+|  ( 281, ( ( _, ( MlyValue.tptp_file tptp_file, tptp_file1left, 
+tptp_file1right)) :: rest671)) => let val  result = MlyValue.tptp (
+( tptp_file ))
+ in ( LrTable.NT 140, ( result, tptp_file1left, tptp_file1right), 
+rest671)
+end
+| _ => raise (mlyAction i392)
+end
+val void = MlyValue.VOID
+val extract = fn a => (fn MlyValue.tptp x => x
+| _ => let exception ParseInternal
+	in raise ParseInternal end) a 
+end
+end
+structure Tokens : TPTP_TOKENS =
+struct
+type svalue = ParserData.svalue
+type ('a,'b) token = ('a,'b) Token.token
+fun AMPERSAND (p1,p2) = Token.TOKEN (ParserData.LrTable.T 0,(
+ParserData.MlyValue.VOID,p1,p2))
+fun AT_SIGN (p1,p2) = Token.TOKEN (ParserData.LrTable.T 1,(
+ParserData.MlyValue.VOID,p1,p2))
+fun CARET (p1,p2) = Token.TOKEN (ParserData.LrTable.T 2,(
+ParserData.MlyValue.VOID,p1,p2))
+fun COLON (p1,p2) = Token.TOKEN (ParserData.LrTable.T 3,(
+ParserData.MlyValue.VOID,p1,p2))
+fun COMMA (p1,p2) = Token.TOKEN (ParserData.LrTable.T 4,(
+ParserData.MlyValue.VOID,p1,p2))
+fun EQUALS (p1,p2) = Token.TOKEN (ParserData.LrTable.T 5,(
+ParserData.MlyValue.VOID,p1,p2))
+fun EXCLAMATION (p1,p2) = Token.TOKEN (ParserData.LrTable.T 6,(
+ParserData.MlyValue.VOID,p1,p2))
+fun LET (p1,p2) = Token.TOKEN (ParserData.LrTable.T 7,(
+ParserData.MlyValue.VOID,p1,p2))
+fun ARROW (p1,p2) = Token.TOKEN (ParserData.LrTable.T 8,(
+ParserData.MlyValue.VOID,p1,p2))
+fun IF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 9,(
+ParserData.MlyValue.VOID,p1,p2))
+fun IFF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 10,(
+ParserData.MlyValue.VOID,p1,p2))
+fun IMPLIES (p1,p2) = Token.TOKEN (ParserData.LrTable.T 11,(
+ParserData.MlyValue.VOID,p1,p2))
+fun INCLUDE (p1,p2) = Token.TOKEN (ParserData.LrTable.T 12,(
+ParserData.MlyValue.VOID,p1,p2))
+fun LAMBDA (p1,p2) = Token.TOKEN (ParserData.LrTable.T 13,(
+ParserData.MlyValue.VOID,p1,p2))
+fun LBRKT (p1,p2) = Token.TOKEN (ParserData.LrTable.T 14,(
+ParserData.MlyValue.VOID,p1,p2))
+fun LPAREN (p1,p2) = Token.TOKEN (ParserData.LrTable.T 15,(
+ParserData.MlyValue.VOID,p1,p2))
+fun MAP_TO (p1,p2) = Token.TOKEN (ParserData.LrTable.T 16,(
+ParserData.MlyValue.VOID,p1,p2))
+fun MMINUS (p1,p2) = Token.TOKEN (ParserData.LrTable.T 17,(
+ParserData.MlyValue.VOID,p1,p2))
+fun NAND (p1,p2) = Token.TOKEN (ParserData.LrTable.T 18,(
+ParserData.MlyValue.VOID,p1,p2))
+fun NEQUALS (p1,p2) = Token.TOKEN (ParserData.LrTable.T 19,(
+ParserData.MlyValue.VOID,p1,p2))
+fun XOR (p1,p2) = Token.TOKEN (ParserData.LrTable.T 20,(
+ParserData.MlyValue.VOID,p1,p2))
+fun NOR (p1,p2) = Token.TOKEN (ParserData.LrTable.T 21,(
+ParserData.MlyValue.VOID,p1,p2))
+fun PERIOD (p1,p2) = Token.TOKEN (ParserData.LrTable.T 22,(
+ParserData.MlyValue.VOID,p1,p2))
+fun PPLUS (p1,p2) = Token.TOKEN (ParserData.LrTable.T 23,(
+ParserData.MlyValue.VOID,p1,p2))
+fun QUESTION (p1,p2) = Token.TOKEN (ParserData.LrTable.T 24,(
+ParserData.MlyValue.VOID,p1,p2))
+fun RBRKT (p1,p2) = Token.TOKEN (ParserData.LrTable.T 25,(
+ParserData.MlyValue.VOID,p1,p2))
+fun RPAREN (p1,p2) = Token.TOKEN (ParserData.LrTable.T 26,(
+ParserData.MlyValue.VOID,p1,p2))
+fun TILDE (p1,p2) = Token.TOKEN (ParserData.LrTable.T 27,(
+ParserData.MlyValue.VOID,p1,p2))
+fun TOK_FALSE (p1,p2) = Token.TOKEN (ParserData.LrTable.T 28,(
+ParserData.MlyValue.VOID,p1,p2))
+fun TOK_I (p1,p2) = Token.TOKEN (ParserData.LrTable.T 29,(
+ParserData.MlyValue.VOID,p1,p2))
+fun TOK_O (p1,p2) = Token.TOKEN (ParserData.LrTable.T 30,(
+ParserData.MlyValue.VOID,p1,p2))
+fun TOK_INT (p1,p2) = Token.TOKEN (ParserData.LrTable.T 31,(
+ParserData.MlyValue.VOID,p1,p2))
+fun TOK_REAL (p1,p2) = Token.TOKEN (ParserData.LrTable.T 32,(
+ParserData.MlyValue.VOID,p1,p2))
+fun TOK_RAT (p1,p2) = Token.TOKEN (ParserData.LrTable.T 33,(
+ParserData.MlyValue.VOID,p1,p2))
+fun TOK_TRUE (p1,p2) = Token.TOKEN (ParserData.LrTable.T 34,(
+ParserData.MlyValue.VOID,p1,p2))
+fun TOK_TYPE (p1,p2) = Token.TOKEN (ParserData.LrTable.T 35,(
+ParserData.MlyValue.VOID,p1,p2))
+fun VLINE (p1,p2) = Token.TOKEN (ParserData.LrTable.T 36,(
+ParserData.MlyValue.VOID,p1,p2))
+fun EOF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 37,(
+ParserData.MlyValue.VOID,p1,p2))
+fun DTHF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 38,(
+ParserData.MlyValue.VOID,p1,p2))
+fun DFOF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 39,(
+ParserData.MlyValue.VOID,p1,p2))
+fun DCNF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 40,(
+ParserData.MlyValue.VOID,p1,p2))
+fun DFOT (p1,p2) = Token.TOKEN (ParserData.LrTable.T 41,(
+ParserData.MlyValue.VOID,p1,p2))
+fun DTFF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 42,(
+ParserData.MlyValue.VOID,p1,p2))
+fun REAL (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 43,(
+ParserData.MlyValue.REAL i,p1,p2))
+fun RATIONAL (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 44,(
+ParserData.MlyValue.RATIONAL i,p1,p2))
+fun SIGNED_INTEGER (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 45,(
+ParserData.MlyValue.SIGNED_INTEGER i,p1,p2))
+fun UNSIGNED_INTEGER (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 46
+,(ParserData.MlyValue.UNSIGNED_INTEGER i,p1,p2))
+fun DOT_DECIMAL (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 47,(
+ParserData.MlyValue.DOT_DECIMAL i,p1,p2))
+fun SINGLE_QUOTED (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 48,(
+ParserData.MlyValue.SINGLE_QUOTED i,p1,p2))
+fun UPPER_WORD (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 49,(
+ParserData.MlyValue.UPPER_WORD i,p1,p2))
+fun LOWER_WORD (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 50,(
+ParserData.MlyValue.LOWER_WORD i,p1,p2))
+fun COMMENT (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 51,(
+ParserData.MlyValue.COMMENT i,p1,p2))
+fun DISTINCT_OBJECT (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 52,(
+ParserData.MlyValue.DISTINCT_OBJECT i,p1,p2))
+fun DUD (p1,p2) = Token.TOKEN (ParserData.LrTable.T 53,(
+ParserData.MlyValue.VOID,p1,p2))
+fun INDEF_CHOICE (p1,p2) = Token.TOKEN (ParserData.LrTable.T 54,(
+ParserData.MlyValue.VOID,p1,p2))
+fun DEFIN_CHOICE (p1,p2) = Token.TOKEN (ParserData.LrTable.T 55,(
+ParserData.MlyValue.VOID,p1,p2))
+fun OPERATOR_FORALL (p1,p2) = Token.TOKEN (ParserData.LrTable.T 56,(
+ParserData.MlyValue.VOID,p1,p2))
+fun OPERATOR_EXISTS (p1,p2) = Token.TOKEN (ParserData.LrTable.T 57,(
+ParserData.MlyValue.VOID,p1,p2))
+fun PLUS (p1,p2) = Token.TOKEN (ParserData.LrTable.T 58,(
+ParserData.MlyValue.VOID,p1,p2))
+fun TIMES (p1,p2) = Token.TOKEN (ParserData.LrTable.T 59,(
+ParserData.MlyValue.VOID,p1,p2))
+fun GENTZEN_ARROW (p1,p2) = Token.TOKEN (ParserData.LrTable.T 60,(
+ParserData.MlyValue.VOID,p1,p2))
+fun DEP_SUM (p1,p2) = Token.TOKEN (ParserData.LrTable.T 61,(
+ParserData.MlyValue.VOID,p1,p2))
+fun DEP_PROD (p1,p2) = Token.TOKEN (ParserData.LrTable.T 62,(
+ParserData.MlyValue.VOID,p1,p2))
+fun ATOMIC_DEFINED_WORD (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 
+63,(ParserData.MlyValue.ATOMIC_DEFINED_WORD i,p1,p2))
+fun ATOMIC_SYSTEM_WORD (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 
+64,(ParserData.MlyValue.ATOMIC_SYSTEM_WORD i,p1,p2))
+fun SUBTYPE (p1,p2) = Token.TOKEN (ParserData.LrTable.T 65,(
+ParserData.MlyValue.VOID,p1,p2))
+fun LET_TERM (p1,p2) = Token.TOKEN (ParserData.LrTable.T 66,(
+ParserData.MlyValue.VOID,p1,p2))
+fun THF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 67,(
+ParserData.MlyValue.VOID,p1,p2))
+fun TFF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 68,(
+ParserData.MlyValue.VOID,p1,p2))
+fun FOF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 69,(
+ParserData.MlyValue.VOID,p1,p2))
+fun CNF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 70,(
+ParserData.MlyValue.VOID,p1,p2))
+fun ITE_F (p1,p2) = Token.TOKEN (ParserData.LrTable.T 71,(
+ParserData.MlyValue.VOID,p1,p2))
+fun ITE_T (p1,p2) = Token.TOKEN (ParserData.LrTable.T 72,(
+ParserData.MlyValue.VOID,p1,p2))
+end
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_parser.ML	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,73 @@
+(*  Title:      HOL/TPTP/TPTP_Parser/tptp_parser.ML
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+An interface for a parser, generated using ML-Yacc, to parse TPTP languages.
+*)
+
+
+(* Build the parser structure *)
+
+structure TPTPLrVals = TPTPLrValsFun(structure Token = LrParser.Token)
+structure TPTPLex = TPTPLexFun(structure Tokens = TPTPLrVals.Tokens)
+structure TPTPParser
+  = JoinWithArg
+     (structure ParserData = TPTPLrVals.ParserData
+      structure Lex = TPTPLex
+      structure LrParser = LrParser)
+
+
+(* Parser interface *)
+structure TPTP_Parser :
+sig
+  val parse_file : string -> TPTP_Syntax.tptp_problem
+  val parse_expression : string -> string -> TPTP_Syntax.tptp_problem
+  exception TPTP_PARSE_ERROR
+end =
+struct
+
+exception TPTP_PARSE_ERROR
+
+val LOOKAHEAD = 0 (*usually set to 15*)
+
+local
+  fun print_error file_name (msg, line, col) =
+    error (file_name ^ "[" ^ Int.toString line ^ ":" ^ Int.toString col ^ "] " ^
+      msg ^ "\n")
+  fun parse lookahead grab file_name =
+    TPTPParser.parse
+      (lookahead,
+       TPTPParser.makeLexer grab file_name,
+       print_error file_name,
+       file_name)
+in
+  fun parse_expression file_name expression =
+    (*file_name only used in reporting error messages*)
+    let
+      val currentPos = Unsynchronized.ref 0
+      fun grab n =
+        if !currentPos = String.size expression then ""
+        else
+          let
+            fun extractStr n =
+              let
+                val s = String.extract (expression, !currentPos, SOME n)
+              in
+                currentPos := !currentPos + n;
+                s
+              end
+            val remaining = String.size expression - !currentPos
+          in if remaining < n then extractStr remaining else extractStr n
+          end
+      val (tree, _ (*remainder*)) =
+        parse LOOKAHEAD grab file_name
+    in tree end
+
+  fun parse_file' lookahead file_name =
+    parse_expression
+     file_name
+     (File.open_input TextIO.inputAll (Path.explode file_name))
+end
+
+val parse_file = parse_file' LOOKAHEAD
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_problem_name.ML	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,150 @@
+(*  Title:      HOL/TPTP/TPTP_Parser/tptp_problem_name.ML
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+Scans a TPTP problem name. Naming convention is described
+http://www.cs.miami.edu/~tptp/TPTP/TR/TPTPTR.shtml#Problem and Axiomatization Naming
+*)
+
+signature TPTP_PROBLEM_NAME =
+sig
+  datatype suffix =
+      Problem of
+        ((*version*)int *
+        (*size parameter*)int option) *
+        (*extension*)string
+    | Axiom of
+        (*specialisation*)int *
+        (*extension*)string
+
+  type tptp_problem_name =
+    {prob_domain : string,
+     prob_number : int,
+     prob_form : TPTP_Syntax.language,
+     suffix : suffix}
+
+  datatype problem_name =
+      Standard of tptp_problem_name
+    | Nonstandard of string
+
+  exception TPTP_PROBLEM_NAME of string
+
+  val parse_problem_name : string -> problem_name
+  val mangle_problem_name : problem_name -> string
+end
+
+structure TPTP_Problem_Name: TPTP_PROBLEM_NAME =
+struct
+
+(*some basic tokens*)
+val numerics = map Char.chr (48 upto 57) (*0..9*)
+val alphabetics =
+  map Char.chr (65 upto 90) @ (*A..Z*)
+  map Char.chr (97 upto 122)  (*a..z*)
+(*TPTP formula forms*)
+val forms = [#"^", #"_", #"=", #"+", #"-"]
+
+(*lift a list of characters into a scanner combinator matching any one of the
+characters in that list.*)
+fun lift l =
+  (map (Char.toString #> ($$)) l, Scan.fail)
+  |-> fold (fn x => fn y => x || y)
+
+(*combinators for parsing letters and numbers*)
+val alpha = lift alphabetics
+val numer = lift numerics
+
+datatype suffix =
+    Problem of
+      ((*version*)int *
+      (*size parameter*)int option) *
+      (*extension*)string
+  | Axiom of
+      (*specialisation*)int *
+      (*extension*)string
+
+val to_int = Int.fromString #> the
+val rm_ending = Scan.this_string "rm"
+val ax_ending =
+  ((numer >> to_int) --|
+   $$ "." -- (Scan.this_string "eq" || Scan.this_string "ax" || rm_ending))
+  >> Axiom
+val prob_ending = $$ "p" || $$ "g" || rm_ending
+val prob_suffix =
+  ((numer >> to_int) --
+   Scan.option ($$ "." |-- numer ^^ numer ^^ numer >> to_int) --| $$ "."
+   -- prob_ending)
+  >> Problem
+
+type tptp_problem_name =
+  {prob_domain : string,
+   prob_number : int,
+   prob_form : TPTP_Syntax.language,
+   suffix : suffix}
+
+datatype problem_name =
+    Standard of tptp_problem_name
+  | Nonstandard of string
+
+exception TPTP_PROBLEM_NAME of string
+
+(*FIXME add graceful handling on non-wellformed TPTP filenames*)
+fun parse_problem_name str' : problem_name =
+  let
+    val str = Symbol.explode str'
+    (*NOTE there's an ambiguity in the spec: there's no way of knowing if a
+    file ending in "rm" used to be "ax" or "p". Here we default to "p".*)
+    val ((((prob_domain, prob_number), prob_form), suffix), rest) =
+     Scan.finite Symbol.stopper
+     ((alpha ^^ alpha ^^ alpha) --
+     (numer ^^ numer ^^ numer >> to_int) --
+     lift forms -- (prob_suffix || ax_ending)) str
+
+    fun parse_form str =
+      case str of
+        "^" => TPTP_Syntax.THF
+      | "_" => TPTP_Syntax.TFF
+      | "=" => TPTP_Syntax.TFF_with_arithmetic
+      | "+" => TPTP_Syntax.FOF
+      | "-" => TPTP_Syntax.CNF
+      | _ => raise TPTP_PROBLEM_NAME ("Unknown TPTP form: " ^ str)
+  in
+    if null rest (*check if the whole name was parsed*)
+    then
+      Standard
+        {prob_domain = prob_domain,
+         prob_number = prob_number,
+         prob_form = parse_form prob_form,
+         suffix = suffix}
+      handle _ => Nonstandard str'
+    else raise TPTP_PROBLEM_NAME ("Parse error")
+  end
+
+(*Produces an ASCII encoding of a TPTP problem-file name.*)
+fun mangle_problem_name (prob : problem_name) : string =
+  case prob of
+     Standard tptp_prob =>
+       let
+         val prob_form =
+           case #prob_form tptp_prob of
+             TPTP_Syntax.THF => "_thf_"
+           | TPTP_Syntax.TFF => "_tff_"
+           | TPTP_Syntax.TFF_with_arithmetic => "_thfwa_"
+           | TPTP_Syntax.FOF => "_fof_"
+           | TPTP_Syntax.CNF => "_cnf_"
+         val suffix =
+           case #suffix tptp_prob of
+             Problem ((version, size), extension) =>
+               Int.toString version ^ "_" ^
+               (if is_some size then Int.toString (the size) ^ "_" else "") ^
+               extension
+           | Axiom (specialisation, extension) =>
+               Int.toString specialisation ^ "_" ^ extension
+       in
+         #prob_domain tptp_prob ^
+         Int.toString (#prob_number tptp_prob) ^
+         prob_form ^
+         suffix
+       end
+   | Nonstandard str => str
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_syntax.ML	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,489 @@
+(*  Title:      HOL/TPTP/TPTP_Parser/tptp_syntax.ML
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+TPTP abstract syntax and parser-related definitions.
+*)
+
+signature TPTP_SYNTAX =
+sig
+  exception TPTP_SYNTAX of string
+  val debug: ('a -> unit) -> 'a -> unit
+
+(*Note that in THF "^ [X] : ^ [Y] : f @ g" should parse
+  as "(^ [X] : (^ [Y] : f)) @ g"
+*)
+
+  datatype number_kind = Int_num | Real_num | Rat_num
+
+  datatype status_value =
+      Suc | Unp | Sap | Esa | Sat | Fsa
+    | Thm | Eqv | Tac | Wec | Eth | Tau
+    | Wtc | Wth | Cax | Sca | Tca | Wca
+    | Cup | Csp | Ecs | Csa | Cth | Ceq
+    | Unc | Wcc | Ect | Fun | Uns | Wuc
+    | Wct | Scc | Uca | Noc
+
+  type name = string
+  type atomic_word = string
+  type inference_rule = atomic_word
+  type file_info = name option
+  type single_quoted = string
+  type file_name = single_quoted
+  type creator_name = atomic_word
+  type variable = string
+  type upper_word = string
+
+  datatype language = FOF | CNF | TFF | THF | FOT | TFF_with_arithmetic
+  and role =
+     Role_Axiom | Role_Hypothesis | Role_Definition | Role_Assumption |
+     Role_Lemma | Role_Theorem | Role_Conjecture | Role_Negated_Conjecture |
+     Role_Plain | Role_Fi_Domain | Role_Fi_Functors | Role_Fi_Predicates |
+     Role_Type | Role_Unknown
+
+  and general_data = (*Bind of variable * formula_data*)
+     Atomic_Word of string
+   | Application of string * general_term list (*general_function*)
+   | V of upper_word (*variable*)
+   | Number of number_kind * string
+   | Distinct_Object of string
+   | (*formula_data*) Formula_Data of language * tptp_formula (* $thf(<thf_formula>) *)
+   | (*formula_data*) Term_Data of tptp_term
+
+  and interpreted_symbol =
+    UMinus | Sum | Difference | Product | Quotient | Quotient_E |
+    Quotient_T | Quotient_F | Remainder_E | Remainder_T | Remainder_F |
+    Floor | Ceiling | Truncate | Round | To_Int | To_Rat | To_Real |
+    (*these should be in defined_pred, but that's not being used in TPTP*)
+    Less | LessEq | Greater | GreaterEq | EvalEq | Is_Int | Is_Rat |
+    Apply (*this is just a matter of convenience*)
+
+  and logic_symbol = Equals | NEquals | Or | And | Iff | If | Fi | Xor |
+    Nor | Nand | Not | Op_Forall | Op_Exists |
+    (*these should be in defined_pred, but that's not being used in TPTP*)
+    True | False
+
+  and quantifier = (*interpreted binders*)
+    Forall | Exists | Epsilon | Iota | Lambda | Dep_Prod | Dep_Sum
+
+  and tptp_base_type =
+    Type_Ind | Type_Bool | Type_Type | Type_Int | Type_Rat | Type_Real
+
+  and symbol =
+      Uninterpreted of string
+    | Interpreted_ExtraLogic of interpreted_symbol
+    | Interpreted_Logic of logic_symbol
+    | TypeSymbol of tptp_base_type
+    | System of string
+
+  and general_term =
+      General_Data of general_data (*general_data*)
+    | General_Term of general_data * general_term (*general_data : general_term*)
+    | General_List of general_term list
+
+  and tptp_term =
+      Term_Func of symbol * tptp_term list
+    | Term_Var of string
+    | Term_Conditional of tptp_formula * tptp_term * tptp_term
+    | Term_Num of number_kind * string
+    | Term_Distinct_Object of string
+
+  and tptp_atom =
+      TFF_Typed_Atom of symbol * tptp_type option (*only TFF*)
+    | THF_Atom_term of tptp_term   (*from here on, only THF*)
+    | THF_Atom_conn_term of symbol
+
+  and tptp_formula =
+      Pred of symbol * tptp_term list
+    | Fmla of symbol * tptp_formula list
+    | Sequent of tptp_formula list * tptp_formula list
+    | Quant of quantifier * (string * tptp_type option) list * tptp_formula
+    | Conditional of tptp_formula * tptp_formula * tptp_formula
+    | Let of tptp_let list * tptp_formula
+    | Atom of tptp_atom
+    | THF_type of tptp_type
+    | THF_typing of tptp_formula * tptp_type (*only THF*)
+
+  and tptp_let =
+      Let_fmla of (string * tptp_type option) * tptp_formula
+    | Let_term of (string * tptp_type option) * tptp_term  (*only TFF*)
+
+  and tptp_type =
+      Prod_type of tptp_type * tptp_type
+    | Fn_type of tptp_type * tptp_type
+    | Atom_type of string
+    | Defined_type of tptp_base_type
+    | Sum_type of tptp_type * tptp_type (*only THF*)
+    | Fmla_type of tptp_formula (*only THF*)
+    | Subtype of symbol * symbol (*only THF*)
+
+  type general_list = general_term list
+  type parent_details = general_list
+  type useful_info = general_term list
+  type info = useful_info
+
+  type annotation = general_term * general_term list
+
+  exception DEQUOTE of string
+
+  type position = string * int * int
+
+  datatype tptp_line =
+      Annotated_Formula of position * language * string * role * tptp_formula * annotation option
+   |  Include of string * string list
+
+  type tptp_problem = tptp_line list
+
+  val dequote : single_quoted -> single_quoted
+
+  val role_to_string : role -> string
+
+  val status_to_string : status_value -> string
+
+  val nameof_tff_atom_type : tptp_type -> string
+
+  (*Returns the list of all files included in a directory and its
+  subdirectories. This is only used for testing the parser/interpreter against
+  all THF problems.*)
+  val get_file_list : Path.T -> Path.T list
+
+  val string_of_tptp_term : tptp_term -> string
+  val string_of_tptp_formula : tptp_formula -> string
+
+end
+
+
+structure TPTP_Syntax : TPTP_SYNTAX =
+struct
+
+exception TPTP_SYNTAX of string
+
+datatype number_kind = Int_num | Real_num | Rat_num
+
+datatype status_value =
+    Suc | Unp | Sap | Esa | Sat | Fsa
+  | Thm | Eqv | Tac | Wec | Eth | Tau
+  | Wtc | Wth | Cax | Sca | Tca | Wca
+  | Cup | Csp | Ecs | Csa | Cth | Ceq
+  | Unc | Wcc | Ect | Fun | Uns | Wuc
+  | Wct | Scc | Uca | Noc
+
+type name = string
+type atomic_word = string
+type inference_rule = atomic_word
+type file_info = name option
+type single_quoted = string
+type file_name = single_quoted
+type creator_name = atomic_word
+type variable = string
+type upper_word = string
+
+datatype language = FOF | CNF | TFF | THF | FOT | TFF_with_arithmetic
+and role =
+  Role_Axiom | Role_Hypothesis | Role_Definition | Role_Assumption |
+  Role_Lemma | Role_Theorem | Role_Conjecture | Role_Negated_Conjecture |
+  Role_Plain | Role_Fi_Domain | Role_Fi_Functors | Role_Fi_Predicates |
+  Role_Type | Role_Unknown
+and general_data = (*Bind of variable * formula_data*)
+    Atomic_Word of string
+  | Application of string * (general_term list)
+  | V of upper_word (*variable*)
+  | Number of number_kind * string
+  | Distinct_Object of string
+  | (*formula_data*) Formula_Data of language * tptp_formula (* $thf(<thf_formula>) *)
+  | (*formula_data*) Term_Data of tptp_term
+
+  and interpreted_symbol =
+    UMinus | Sum | Difference | Product | Quotient | Quotient_E |
+    Quotient_T | Quotient_F | Remainder_E | Remainder_T | Remainder_F |
+    Floor | Ceiling | Truncate | Round | To_Int | To_Rat | To_Real |
+    (*these should be in defined_pred, but that's not being used in TPTP*)
+    Less | LessEq | Greater | GreaterEq | EvalEq | Is_Int | Is_Rat |
+    Apply (*this is just a matter of convenience*)
+
+  and logic_symbol = Equals | NEquals | Or | And | Iff | If | Fi | Xor |
+    Nor | Nand | Not | Op_Forall | Op_Exists |
+    (*these should be in defined_pred, but that's not being used in TPTP*)
+    True | False
+
+  and quantifier = (*interpreted binders*)
+    Forall | Exists | Epsilon | Iota | Lambda | Dep_Prod | Dep_Sum
+
+  and tptp_base_type =
+    Type_Ind | Type_Bool | Type_Type | Type_Int | Type_Rat | Type_Real
+
+  and symbol =
+      Uninterpreted of string
+    | Interpreted_ExtraLogic of interpreted_symbol
+    | Interpreted_Logic of logic_symbol
+    | TypeSymbol of tptp_base_type
+    | System of string
+
+  and general_term =
+      General_Data of general_data (*general_data*)
+    | General_Term of general_data * general_term (*general_data : general_term*)
+    | General_List of general_term list
+
+  and tptp_term =
+      Term_Func of symbol * tptp_term list
+    | Term_Var of string
+    | Term_Conditional of tptp_formula * tptp_term * tptp_term
+    | Term_Num of number_kind * string
+    | Term_Distinct_Object of string
+
+  and tptp_atom =
+      TFF_Typed_Atom of symbol * tptp_type option (*only TFF*)
+    | THF_Atom_term of tptp_term   (*from here on, only THF*)
+    | THF_Atom_conn_term of symbol
+
+  and tptp_formula =
+      Pred of symbol * tptp_term list
+    | Fmla of symbol * tptp_formula list
+    | Sequent of tptp_formula list * tptp_formula list
+    | Quant of quantifier * (string * tptp_type option) list * tptp_formula
+    | Conditional of tptp_formula * tptp_formula * tptp_formula
+    | Let of tptp_let list * tptp_formula
+    | Atom of tptp_atom
+    | THF_type of tptp_type
+    | THF_typing of tptp_formula * tptp_type
+
+  and tptp_let =
+      Let_fmla of (string * tptp_type option) * tptp_formula (*both TFF and THF*)
+    | Let_term of (string * tptp_type option) * tptp_term  (*only TFF*)
+
+  and tptp_type =
+      Prod_type of tptp_type * tptp_type
+    | Fn_type of tptp_type * tptp_type
+    | Atom_type of string
+    | Defined_type of tptp_base_type
+    | Sum_type of tptp_type * tptp_type (*only THF*)
+    | Fmla_type of tptp_formula (*only THF*)
+    | Subtype of symbol * symbol (*only THF*)
+
+type general_list = general_term list
+type parent_details = general_list
+type useful_info = general_term list
+type info = useful_info
+
+(*type annotation = (source * info option)*)
+type annotation = general_term * general_term list
+
+exception DEQUOTE of string
+
+(*
+datatype defined_functor =
+  ITE_T | UMINUS | SUM | DIFFERENCE | PRODUCT | QUOTIENT | QUOTIENT_E |
+  QUOTIENT_T | QUOTIENT_F | REMAINDER_E | REMAINDER_T | REMAINDER_F |
+  FLOOR | CEILING | TRUNCATE | ROUND | TO_INT | TO_RAT | TO_REAL
+*)
+
+type position = string * int * int
+
+datatype tptp_line =
+    Annotated_Formula of position * language * string * role * tptp_formula * annotation option
+ |  Include of string * string list
+
+type tptp_problem = tptp_line list
+
+fun debug f x = if !Runtime.debug then (f x; ()) else ()
+
+fun nameof_tff_atom_type (Atom_type str) = str
+  | nameof_tff_atom_type _ = raise TPTP_SYNTAX "nameof_tff_atom_type called on non-atom type"
+
+(*Used for debugging. Returns all files contained within a directory or its
+subdirectories. Follows symbolic links, filters away directories.*)
+fun get_file_list path =
+  let
+    fun check_file_entry f rest =
+      let
+        (*NOTE needed since no File.is_link and File.read_link*)
+        val f_str = Path.implode f
+      in
+        if File.is_dir f then
+          rest @ get_file_list f
+        else if OS.FileSys.isLink f_str then
+          (*follow links -- NOTE this breaks if links are relative paths*)
+          check_file_entry (Path.explode (OS.FileSys.readLink f_str)) rest
+        else f :: rest
+      end
+  in
+    File.read_dir path
+    |> map
+        (Path.explode
+        #> Path.append path)
+    |> (fn l => fold check_file_entry l [])
+  end
+
+fun role_to_string role =
+  case role of
+      Role_Axiom => "axiom"
+    | Role_Hypothesis => "hypothesis"
+    | Role_Definition => "definition"
+    | Role_Assumption => "assumption"
+    | Role_Lemma => "lemma"
+    | Role_Theorem => "theorem"
+    | Role_Conjecture => "conjecture"
+    | Role_Negated_Conjecture => "negated_conjecture"
+    | Role_Plain => "plain"
+    | Role_Fi_Domain => "fi_domain"
+    | Role_Fi_Functors => "fi_functors"
+    | Role_Fi_Predicates => "fi_predicates"
+    | Role_Type => "type"
+    | Role_Unknown => "unknown"
+
+(*accepts a string "'abc'" and returns "abc"*)
+fun dequote str : single_quoted =
+  if str = "" then
+    raise (DEQUOTE "empty string")
+  else
+    (unprefix "'" str
+    |> unsuffix "'"
+    handle (Fail str) =>
+      if str = "unprefix" then
+        raise DEQUOTE ("string doesn't open with quote:" ^ str)
+      else if str = "unsuffix" then
+        raise DEQUOTE ("string doesn't close with quote:" ^ str)
+      else raise Fail str)
+
+
+(* Printing parsed TPTP formulas *)
+(*FIXME this is not pretty-printing, just printing*)
+
+fun status_to_string status_value =
+  case status_value of
+      Suc => "suc"	| Unp => "unp"
+    | Sap => "sap"  | Esa => "esa"
+    | Sat => "sat"  | Fsa => "fsa"
+    | Thm => "thm"  | Wuc => "wuc"
+    | Eqv => "eqv"	| Tac => "tac"
+    | Wec => "wec"	| Eth => "eth"
+    | Tau => "tau"	| Wtc => "wtc"
+    | Wth => "wth"	| Cax => "cax"
+    | Sca => "sca"	| Tca => "tca"
+    | Wca => "wca"	| Cup => "cup"
+    | Csp => "csp"	| Ecs => "ecs"
+    | Csa => "csa"	| Cth => "cth"
+    | Ceq => "ceq"	| Unc => "unc"
+    | Wcc => "wcc"	| Ect => "ect"
+    | Fun => "fun"	| Uns => "uns"
+    | Wct => "wct"	| Scc => "scc"
+    | Uca => "uca"	| Noc => "noc"
+
+fun string_of_tptp_term x =
+  case x of
+      Term_Func (symbol, tptp_term_list) =>
+        "(" ^ string_of_symbol symbol ^ " " ^
+        String.concatWith " " (map string_of_tptp_term tptp_term_list) ^ ")"
+    | Term_Var str => str
+    | Term_Conditional (tptp_formula, tptp_term1, tptp_term2) => "" (*FIXME*)
+    | Term_Num (_, str) => str
+    | Term_Distinct_Object str => str
+
+and string_of_symbol (Uninterpreted str) = str
+  | string_of_symbol (Interpreted_ExtraLogic interpreted_symbol) = string_of_interpreted_symbol interpreted_symbol
+  | string_of_symbol (Interpreted_Logic logic_symbol) = string_of_logic_symbol logic_symbol
+  | string_of_symbol (TypeSymbol tptp_base_type) = string_of_tptp_base_type tptp_base_type
+  | string_of_symbol (System str) = str
+
+and string_of_tptp_base_type Type_Ind = "$i"
+  | string_of_tptp_base_type Type_Bool = "$o"
+  | string_of_tptp_base_type Type_Type = "$tType"
+  | string_of_tptp_base_type Type_Int = "$int"
+  | string_of_tptp_base_type Type_Rat = "$rat"
+  | string_of_tptp_base_type Type_Real = "$real"
+
+and string_of_interpreted_symbol x =
+  case x of
+      UMinus => "$uminus"
+    | Sum => "$sum"
+    | Difference => "$difference"
+    | Product => "$product"
+    | Quotient => "$quotient"
+    | Quotient_E => "$quotient_e"
+    | Quotient_T => "$quotient_t"
+    | Quotient_F => "$quotient_f"
+    | Remainder_E => "$remainder_e"
+    | Remainder_T => "$remainder_t"
+    | Remainder_F => "$remainder_f"
+    | Floor => "$floor"
+    | Ceiling => "$ceiling"
+    | Truncate => "$truncate"
+    | Round => "$round"
+    | To_Int => "$to_int"
+    | To_Rat => "$to_rat"
+    | To_Real => "$to_real"
+    | Less => "$less"
+    | LessEq => "$lesseq"
+    | Greater => "$greater"
+    | GreaterEq => "$greatereq"
+    | EvalEq => "$evaleq"
+    | Is_Int => "$is_int"
+    | Is_Rat => "$is_rat"
+    | Apply => "@"
+
+and string_of_logic_symbol Equals = "="
+  | string_of_logic_symbol NEquals = "!="
+  | string_of_logic_symbol Or = "|"
+  | string_of_logic_symbol And = "&"
+  | string_of_logic_symbol Iff = "<=>"
+  | string_of_logic_symbol If = "=>"
+  | string_of_logic_symbol Fi = "<="
+  | string_of_logic_symbol Xor = "<~>"
+  | string_of_logic_symbol Nor = "~|"
+  | string_of_logic_symbol Nand = "~&"
+  | string_of_logic_symbol Not = "~"
+  | string_of_logic_symbol Op_Forall = "!!"
+  | string_of_logic_symbol Op_Exists = "??"
+  | string_of_logic_symbol True = "$true"
+  | string_of_logic_symbol False = "$false"
+
+and string_of_quantifier Forall = "!"
+  | string_of_quantifier Exists  = "?"
+  | string_of_quantifier Epsilon  = "@+"
+  | string_of_quantifier Iota  = "@-"
+  | string_of_quantifier Lambda  = "^"
+  | string_of_quantifier Dep_Prod = "!>"
+  | string_of_quantifier Dep_Sum  = "?*"
+
+and string_of_tptp_atom (TFF_Typed_Atom (symbol, tptp_type_option)) =
+    (case tptp_type_option of
+       NONE => string_of_symbol symbol
+     | SOME tptp_type =>
+         string_of_symbol symbol ^ " : " ^ string_of_tptp_type tptp_type)
+  | string_of_tptp_atom (THF_Atom_term tptp_term) = string_of_tptp_term tptp_term
+  | string_of_tptp_atom (THF_Atom_conn_term symbol) = string_of_symbol symbol
+
+and string_of_tptp_formula (Pred (symbol, tptp_term_list)) =
+      "(" ^ string_of_symbol symbol ^
+      String.concatWith " " (map string_of_tptp_term tptp_term_list) ^ ")"
+  | string_of_tptp_formula (Fmla (symbol, tptp_formula_list)) =
+      "(" ^
+      string_of_symbol symbol ^
+      String.concatWith " " (map string_of_tptp_formula tptp_formula_list) ^ ")"
+  | string_of_tptp_formula (Sequent (tptp_formula_list1, tptp_formula_list2)) = "" (*FIXME*)
+  | string_of_tptp_formula (Quant (quantifier, varlist, tptp_formula)) =
+      string_of_quantifier quantifier ^ "[" ^
+      String.concatWith ", " (map (fn (n, ty) =>
+        case ty of
+          NONE => n
+        | SOME ty => n ^ " : " ^ string_of_tptp_type ty) varlist) ^ "] : (" ^
+      string_of_tptp_formula tptp_formula ^ ")"
+  | string_of_tptp_formula (Conditional _) = "" (*FIXME*)
+  | string_of_tptp_formula (Let _) = "" (*FIXME*)
+  | string_of_tptp_formula (Atom tptp_atom) = string_of_tptp_atom tptp_atom
+  | string_of_tptp_formula (THF_type tptp_type) = string_of_tptp_type tptp_type
+  | string_of_tptp_formula (THF_typing (tptp_formula, tptp_type)) =
+      string_of_tptp_formula tptp_formula ^ " : " ^ string_of_tptp_type tptp_type
+
+and string_of_tptp_type (Prod_type (tptp_type1, tptp_type2)) =
+      string_of_tptp_type tptp_type1 ^ " * " ^ string_of_tptp_type tptp_type2
+  | string_of_tptp_type (Fn_type (tptp_type1, tptp_type2)) =
+      string_of_tptp_type tptp_type1 ^ " > " ^ string_of_tptp_type tptp_type2
+  | string_of_tptp_type (Atom_type str) = str
+  | string_of_tptp_type (Defined_type tptp_base_type) =
+      string_of_tptp_base_type tptp_base_type
+  | string_of_tptp_type (Sum_type (tptp_type1, tptp_type2)) = ""
+  | string_of_tptp_type (Fmla_type tptp_formula) = string_of_tptp_formula tptp_formula
+  | string_of_tptp_type (Subtype (symbol1, symbol2)) =
+      string_of_symbol symbol1 ^ " << " ^ string_of_symbol symbol2
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser_Test.thy	Fri Mar 09 17:24:42 2012 +0000
@@ -0,0 +1,241 @@
+(*  Title:      HOL/TPTP/TPTP_Parser_Test.thy
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+Some tests for the TPTP interface. Some of the tests rely on the Isabelle
+environment variable TPTP_PROBLEMS_PATH, which should point to the
+TPTP-vX.Y.Z/Problems directory.
+*)
+
+theory TPTP_Parser_Test
+imports TPTP_Parser
+begin
+
+ML {*
+  val warning_out = Attrib.setup_config_string @{binding "warning_out"} (K "")
+  fun S x y z = x z (y z)
+*}
+
+section "Parser tests"
+
+ML {*
+  fun payload_of (TPTP_Syntax.Annotated_Formula (_, _, _, _, fmla, _)) = fmla
+  val payloads_of = map payload_of
+*}
+ML {*
+  TPTP_Parser.parse_expression "" "fof(dt_k4_waybel34, axiom, ~ v3).";
+  TPTP_Parser.parse_expression "" "thf(dt_k4_waybel34, axiom, ~ ($true | $false)).";
+  TPTP_Parser.parse_expression ""
+   "thf(dt_k4_waybel34, axiom, ~ (! [X : $o, Y : ($o > $o)] : ( (X | (Y = Y))))).";
+  TPTP_Parser.parse_expression "" "tff(dt_k4_waybel34, axiom, ~ ($true)).";
+  payloads_of it;
+*}
+ML {*
+  TPTP_Parser.parse_expression "" "thf(bla, type, x : $o).";
+  TPTP_Parser.parse_expression ""
+   "fof(dt_k4_waybel34, axiom, ~ v3_struct_0(k4_waybel34(A))).";
+  TPTP_Parser.parse_expression ""
+   "fof(dt_k4_waybel34, axiom, (! [A] : (v1_xboole_0(A) => ( ~ v3_struct_0(k4_waybel34(A)))))).";
+*}
+ML {*
+  TPTP_Parser.parse_expression ""
+  ("fof(dt_k4_waybel34,axiom,(" ^
+    "! [A] :" ^
+      "( ~ v1_xboole_0(A)" ^
+     "=> ( ~ v3_struct_0(k4_waybel34(A))" ^
+        "& v2_altcat_1(k4_waybel34(A))" ^
+        "& v6_altcat_1(k4_waybel34(A))" ^
+        "& v11_altcat_1(k4_waybel34(A))" ^
+        "& v12_altcat_1(k4_waybel34(A))" ^
+        "& v2_yellow21(k4_waybel34(A))" ^
+        "& l2_altcat_1(k4_waybel34(A)) ) ) )).")
+*}
+
+ML {*
+open TPTP_Syntax;
+@{assert}
+  ((TPTP_Parser.parse_expression ""
+   "thf(x,axiom,^ [X] : ^ [Y] : f @ g)."
+   |> payloads_of |> hd)
+  =
+  Fmla (Interpreted_ExtraLogic Apply,
+   [Quant (Lambda, [("X", NONE)],
+      Quant (Lambda, [("Y", NONE)],
+       Atom (THF_Atom_term (Term_Func (Uninterpreted "f", []))))),
+    Atom (THF_Atom_term (Term_Func (Uninterpreted "g", [])))])
+)
+*}
+
+
+section "Source problems"
+ML {*
+  (*problem source*)
+  val thf_probs_dir =
+    Path.explode "$TPTP_PROBLEMS_PATH"
+    |> Path.expand;
+
+  (*list of files to under test*)
+  val files = TPTP_Syntax.get_file_list thf_probs_dir;
+
+(*  (*test problem-name parsing and mangling*)
+  val problem_names =
+    map (Path.base #>
+         Path.implode #>
+         TPTP_Problem_Name.parse_problem_name #>
+         TPTP_Problem_Name.mangle_problem_name)
+        files*)
+*}
+
+
+section "Supporting test functions"
+ML {*
+  fun report ctxt str =
+    let
+      val warning_out = Config.get ctxt warning_out
+    in
+      if warning_out = "" then warning str
+      else
+        let
+          val out_stream = TextIO.openAppend warning_out
+        in (TextIO.output (out_stream, str ^ "\n");
+            TextIO.flushOut out_stream;
+            TextIO.closeOut out_stream)
+        end
+    end
+
+  fun test_fn ctxt f msg default_val file_name =
+    let
+      val _ = TPTP_Syntax.debug tracing (msg ^ " " ^ Path.print file_name)
+    in
+     (f file_name; ())
+     (*otherwise report exceptions as warnings*)
+     handle exn =>
+       if Exn.is_interrupt exn then
+         reraise exn
+       else
+         (report ctxt (msg ^ " test: file " ^ Path.print file_name ^
+          " raised exception: " ^ ML_Compiler.exn_message exn);
+          default_val)
+    end
+
+  fun timed_test ctxt f =
+    let
+      fun f' x = (f x; ())
+      val time =
+        Timing.timing (List.app f') files
+        |> fst
+      val duration =
+        #elapsed time
+        |> Time.toSeconds
+        |> Real.fromLargeInt
+      val average =
+        (StringCvt.FIX (SOME 3),
+         (duration / Real.fromInt (length files)))
+        |-> Real.fmt
+    in
+      report ctxt ("Test timing: " ^ Timing.message time ^ "\n(about " ^ average ^
+       "s per problem)")
+    end
+*}
+
+
+subsection "More parser tests"
+ML {*
+  fun situate file_name = Path.append thf_probs_dir (Path.explode file_name);
+  fun parser_test ctxt = (*FIXME argument order*)
+    test_fn ctxt
+     (fn file_name =>
+        Path.implode file_name
+        |> (fn file =>
+             ((*report ctxt file; this is if you want the filename in the log*)
+              TPTP_Parser.parse_file file)))
+     "parser"
+     ()
+*}
+
+declare [[warning_out = ""]]
+
+text "Parse a specific problem."
+ML {*
+  map TPTP_Parser.parse_file
+   ["$TPTP_PROBLEMS_PATH/FLD/FLD051-1.p",
+    "$TPTP_PROBLEMS_PATH/FLD/FLD005-3.p",
+    "$TPTP_PROBLEMS_PATH/SWV/SWV567-1.015.p",
+    "$TPTP_PROBLEMS_PATH/SWV/SWV546-1.010.p"]
+*}
+ML {*
+  parser_test @{context} (situate "DAT/DAT001=1.p");
+  parser_test @{context} (situate "ALG/ALG001^5.p");
+  parser_test @{context} (situate "NUM/NUM021^1.p");
+  parser_test @{context} (situate "SYN/SYN000^1.p")
+*}
+
+text "Run the parser over all problems."
+ML {*report @{context} "Testing parser"*}
+ML {*
+(*  val _ = S timed_test parser_test @{context}*)
+*}
+
+
+subsection "Interpretation"
+
+text "Interpret a problem."
+ML {*
+  val (time, ((type_map, const_map, fmlas), thy)) =
+    Timing.timing
+    (TPTP_Interpret.interpret_file
+     false
+     (Path.dir thf_probs_dir)
+    (Path.append thf_probs_dir (Path.explode "LCL/LCL825-1.p"))
+     []
+     [])
+    @{theory}
+
+  (*also tried
+     "ALG/ALG001^5.p"
+     "COM/COM003+2.p"
+     "COM/COM003-1.p"
+     "COM/COM024^5.p"
+     "DAT/DAT017=1.p"
+     "NUM/NUM021^1.p"
+     "NUM/NUM858=1.p"
+     "SYN/SYN000^2.p"*)
+
+ (*These take too long
+    "NLP/NLP562+1.p"
+    "SWV/SWV546-1.010.p"
+    "SWV/SWV567-1.015.p"
+    "LCL/LCL680+1.020.p"*)
+*}
+
+text "... and display nicely."
+ML {*
+  List.app (Pretty.writeln o (Syntax.pretty_term @{context}) o #4) fmlas
+*}
+ML {*
+  (*Don't use TPTP_Syntax.string_of_tptp_formula, it's just for testing*)
+  List.app (writeln o TPTP_Syntax.string_of_tptp_formula o #3) fmlas
+*}
+
+
+text "Run interpretation over all problems. This works only for logics
+ for which interpretation is defined (in TPTP_Parser/tptp_interpret.ML)."
+ML {*
+  report @{context} "Interpreting all problems.";
+  fun interpretation_test timeout ctxt =
+    test_fn ctxt
+     (fn file =>
+      (writeln (Path.implode file);
+       TimeLimit.timeLimit (Time.fromSeconds timeout)
+       (TPTP_Interpret.interpret_file
+         false
+         (Path.dir thf_probs_dir)
+         file
+         []
+         [])
+         @{theory}))
+     "interpreter"
+     ()
+  val _ = S timed_test (interpretation_test 5) @{context}
+*}
+
+end
\ No newline at end of file
--- a/src/Pure/PIDE/xml.ML	Fri Mar 09 17:24:00 2012 +0000
+++ b/src/Pure/PIDE/xml.ML	Fri Mar 09 17:24:42 2012 +0000
@@ -3,7 +3,7 @@
     Author:     Stefan Berghofer
     Author:     Makarius
 
-Untyped XML trees and basic data representation.
+Untyped XML trees and representation of ML values.
 *)
 
 signature XML_DATA_OPS =
@@ -28,9 +28,9 @@
 
 signature XML =
 sig
-  type attributes = Properties.T
+  type attributes = (string * string) list
   datatype tree =
-      Elem of Markup.T * tree list
+      Elem of (string * attributes) * tree list
     | Text of string
   type body = tree list
   val add_content: tree -> Buffer.T -> Buffer.T
@@ -59,10 +59,10 @@
 
 (** XML trees **)
 
-type attributes = Properties.T;
+type attributes = (string * string) list;
 
 datatype tree =
-    Elem of Markup.T * tree list
+    Elem of (string * attributes) * tree list
   | Text of string;
 
 type body = tree list;
@@ -360,8 +360,7 @@
   | node t = raise XML_BODY [t];
 
 fun vector atts =
-  #1 (fold_map (fn (a, x) =>
-        fn i => if int_atom a = i then (x, i + 1) else raise XML_ATOM a) atts 0);
+  map_index (fn (i, (a, x)) => if int_atom a = i then x else raise XML_ATOM a) atts;
 
 fun tagged (Elem ((name, atts), ts)) = (int_atom name, (vector atts, ts))
   | tagged t = raise XML_BODY [t];
--- a/src/Pure/PIDE/xml.scala	Fri Mar 09 17:24:00 2012 +0000
+++ b/src/Pure/PIDE/xml.scala	Fri Mar 09 17:24:42 2012 +0000
@@ -215,7 +215,7 @@
     private def node(ts: XML.Body): XML.Tree = XML.Elem(Markup(":", Nil), ts)
 
     private def vector(xs: List[String]): XML.Attributes =
-      xs.zipWithIndex.map(p => (int_atom(p._2), p._1))
+      xs.zipWithIndex.map({ case (x, i) => (int_atom(i), x) })
 
     private def tagged(tag: Int, data: (List[String], XML.Body)): XML.Tree =
       XML.Elem(Markup(int_atom(tag), vector(data._1)), data._2)
@@ -293,15 +293,8 @@
       }
 
     private def vector(atts: XML.Attributes): List[String] =
-    {
-      val xs = new mutable.ListBuffer[String]
-      var i = 0
-      for ((a, x) <- atts) {
-        if (int_atom(a) == i) { xs += x; i = i + 1 }
-        else throw new XML_Atom(a)
-      }
-      xs.toList
-    }
+      atts.iterator.zipWithIndex.map(
+        { case ((a, x), i) => if (int_atom(a) == i) x else throw new XML_Atom(a) }).toList
 
     private def tagged(t: XML.Tree): (Int, (List[String], XML.Body)) =
       t match {
--- a/src/Pure/library.ML	Fri Mar 09 17:24:00 2012 +0000
+++ b/src/Pure/library.ML	Fri Mar 09 17:24:42 2012 +0000
@@ -470,9 +470,9 @@
   let
     fun get (_: int) [] = NONE
       | get i (x :: xs) =
-          case f x
-           of NONE => get (i + 1) xs
-            | SOME y => SOME (i, y)
+          (case f x of
+            NONE => get (i + 1) xs
+          | SOME y => SOME (i, y))
   in get 0 end;
 
 val flat = List.concat;
--- a/src/Tools/induct_tacs.ML	Fri Mar 09 17:24:00 2012 +0000
+++ b/src/Tools/induct_tacs.ML	Fri Mar 09 17:24:42 2012 +0000
@@ -18,12 +18,12 @@
 
 (* case analysis *)
 
-fun check_type ctxt t =
+fun check_type ctxt (t, pos) =
   let
     val u = singleton (Variable.polymorphic ctxt) t;
     val U = Term.fastype_of u;
     val _ = Term.is_TVar U andalso
-      error ("Cannot determine type of " ^ Syntax.string_of_term ctxt u);
+      error ("Cannot determine type of " ^ Syntax.string_of_term ctxt u ^ Position.str_of pos);
   in (u, U) end;
 
 fun gen_case_tac ctxt0 s opt_rule i st =
@@ -34,7 +34,8 @@
         SOME rule => rule
       | NONE =>
           (case Induct.find_casesT ctxt
-              (#2 (check_type ctxt (Proof_Context.read_term_schematic ctxt s))) of
+              (#2 (check_type ctxt (Proof_Context.read_term_schematic ctxt s,
+                #2 (Syntax.read_token s)))) of
             rule :: _ => rule
           | [] => @{thm case_split}));
     val _ = Method.trace ctxt [rule];
@@ -71,17 +72,22 @@
     fun induct_var name =
       let
         val t = Syntax.read_term ctxt name;
+        val (_, pos) = Syntax.read_token name;
         val (x, _) = Term.dest_Free t handle TERM _ =>
-          error ("Induction argument not a variable: " ^ Syntax.string_of_term ctxt t);
+          error ("Induction argument not a variable: " ^
+            Syntax.string_of_term ctxt t ^ Position.str_of pos);
         val eq_x = fn Free (y, _) => x = y | _ => false;
         val _ =
           if Term.exists_subterm eq_x concl then ()
-          else error ("No such variable in subgoal: " ^ Syntax.string_of_term ctxt t);
+          else
+            error ("No such variable in subgoal: " ^
+              Syntax.string_of_term ctxt t ^ Position.str_of pos);
         val _ =
           if (exists o Term.exists_subterm) eq_x prems then
-            warning ("Induction variable occurs also among premises: " ^ Syntax.string_of_term ctxt t)
+            warning ("Induction variable occurs also among premises: " ^
+              Syntax.string_of_term ctxt t ^ Position.str_of pos)
           else ();
-      in #1 (check_type ctxt t) end;
+      in #1 (check_type ctxt (t, pos)) end;
 
     val argss = map (map (Option.map induct_var)) varss;
     val rule =