merged
authorhuffman
Wed, 04 Apr 2012 20:45:19 +0200
changeset 47369 f483be2fecb9
parent 47368 4c522dff1f4d (current diff)
parent 47365 7b09206bb74b (diff)
child 47370 eabfb53cfca8
merged
src/HOL/Lifting.thy
--- a/src/HOL/Import/HOL_Light_Maps.thy	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/Import/HOL_Light_Maps.thy	Wed Apr 04 20:45:19 2012 +0200
@@ -93,9 +93,9 @@
   "snd = (\<lambda>p\<Colon>'A \<times> 'B. SOME y\<Colon>'B. \<exists>x\<Colon>'A. p = (x, y))"
   by auto
 
-(*lemma [import_const CURRY]:
+lemma CURRY_DEF[import_const CURRY]:
   "curry = (\<lambda>(f\<Colon>'A \<times> 'B \<Rightarrow> 'C) x y. f (x, y))"
-  using curry_def .*)
+  using curry_def .
 
 lemma [import_const ONE_ONE : Fun.inj]:
   "inj = (\<lambda>(f\<Colon>'A \<Rightarrow> 'B). \<forall>x1 x2. f x1 = f x2 \<longrightarrow> x1 = x2)"
@@ -200,16 +200,16 @@
   "\<forall>m n\<Colon>nat. if n = id 0 then m div n = id 0 \<and> m mod n = m else m = m div n * n + m mod n \<and> m mod n < n"
   by simp
 
-lemmas [import_type sum "_dest_sum" "_mk_sum"] = type_definition_sum
+lemmas [import_type sum "_dest_sum" "_mk_sum"] = type_definition_sum[where 'a="'A" and 'b="'B"]
 import_const_map INL : Sum_Type.Inl
 import_const_map INR : Sum_Type.Inr
 
 lemma sum_INDUCT:
-  "\<forall>P. (\<forall>a. P (Inl a)) \<and> (\<forall>a. P (Inr a)) \<longrightarrow> (\<forall>x. P x)"
+  "\<forall>P. (\<forall>a :: 'A. P (Inl a)) \<and> (\<forall>a :: 'B. P (Inr a)) \<longrightarrow> (\<forall>x. P x)"
   by (auto intro: sum.induct)
 
 lemma sum_RECURSION:
-  "\<forall>Inl' Inr'. \<exists>fn. (\<forall>a. fn (Inl a) = Inl' a) \<and> (\<forall>a. fn (Inr a) = Inr' a)"
+  "\<forall>Inl' Inr'. \<exists>fn. (\<forall>a :: 'A. fn (Inl a) = (Inl' a :: 'Z)) \<and> (\<forall>a :: 'B. fn (Inr a) = Inr' a)"
   by (intro allI, rule_tac x="sum_case Inl' Inr'" in exI) auto
 
 lemma OUTL[import_const "OUTL" : "Sum_Type.Projl"]:
@@ -249,7 +249,7 @@
   by simp
 
 lemma LENGTH[import_const LENGTH : List.length]:
-  "length [] = id 0 \<and> (\<forall>(h\<Colon>'A) t. length (h # t) = Suc (length t))"
+  "length ([] :: 'A list) = id 0 \<and> (\<forall>(h\<Colon>'A) t. length (h # t) = Suc (length t))"
   by simp
 
 lemma MAP[import_const MAP : List.map]:
@@ -307,7 +307,7 @@
   by simp
 
 lemma WF[import_const WF : Wellfounded.wfP]:
-  "wfP u \<longleftrightarrow> (\<forall>P. (\<exists>x :: 'A. P x) \<longrightarrow> (\<exists>x. P x \<and> (\<forall>y. u y x \<longrightarrow> \<not> P y)))"
+  "\<forall>u. wfP u \<longleftrightarrow> (\<forall>P. (\<exists>x :: 'A. P x) \<longrightarrow> (\<exists>x. P x \<and> (\<forall>y. u y x \<longrightarrow> \<not> P y)))"
 proof (intro allI iffI impI wfI_min[to_pred], elim exE wfE_min[to_pred])
   fix x :: "'a \<Rightarrow> 'a \<Rightarrow> bool" and xa :: "'a" and Q
   assume a: "xa \<in> Q"
--- a/src/HOL/Import/import_rule.ML	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/Import/import_rule.ML	Wed Apr 04 20:45:19 2012 +0200
@@ -235,7 +235,7 @@
           [SOME rept, SOME abst, SOME cP, SOME (cterm_of thy' (Free ("a", aty))),
              SOME (cterm_of thy' (Free ("r", typ_of ctT)))]
           @{thm typedef_hol2hollight}
-    val th4 = meta_mp typedef_th (#type_definition (#2 typedef_info))
+    val th4 = typedef_th OF [#type_definition (#2 typedef_info)]
   in
     (th4, thy')
   end
--- a/src/HOL/Lifting.thy	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/Lifting.thy	Wed Apr 04 20:45:19 2012 +0200
@@ -236,16 +236,17 @@
   shows "invariant P x x \<equiv> P x"
 using assms by (auto simp add: invariant_def)
 
-lemma copy_type_to_Quotient:
+lemma UNIV_typedef_to_Quotient:
   assumes "type_definition Rep Abs UNIV"
-  and T_def: "T \<equiv> (\<lambda>x y. Abs x = y)"
+  and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"
   shows "Quotient (op =) Abs Rep T"
 proof -
   interpret type_definition Rep Abs UNIV by fact
-  from Abs_inject Rep_inverse T_def show ?thesis by (auto intro!: QuotientI)
+  from Abs_inject Rep_inverse Abs_inverse T_def show ?thesis 
+    by (fastforce intro!: QuotientI fun_eq_iff)
 qed
 
-lemma copy_type_to_equivp:
+lemma UNIV_typedef_to_equivp:
   fixes Abs :: "'a \<Rightarrow> 'b"
   and Rep :: "'b \<Rightarrow> 'a"
   assumes "type_definition Rep Abs (UNIV::'a set)"
@@ -253,6 +254,28 @@
 by (rule identity_equivp)
 
 lemma typedef_to_Quotient:
+  assumes "type_definition Rep Abs S"
+  and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"
+  shows "Quotient (Lifting.invariant (\<lambda>x. x \<in> S)) Abs Rep T"
+proof -
+  interpret type_definition Rep Abs S by fact
+  from Rep Abs_inject Rep_inverse Abs_inverse T_def show ?thesis
+    by (auto intro!: QuotientI simp: invariant_def fun_eq_iff)
+qed
+
+lemma typedef_to_part_equivp:
+  assumes "type_definition Rep Abs S"
+  shows "part_equivp (Lifting.invariant (\<lambda>x. x \<in> S))"
+proof (intro part_equivpI)
+  interpret type_definition Rep Abs S by fact
+  show "\<exists>x. Lifting.invariant (\<lambda>x. x \<in> S) x x" using Rep by (auto simp: invariant_def)
+next
+  show "symp (Lifting.invariant (\<lambda>x. x \<in> S))" by (auto intro: sympI simp: invariant_def)
+next
+  show "transp (Lifting.invariant (\<lambda>x. x \<in> S))" by (auto intro: transpI simp: invariant_def)
+qed
+
+lemma open_typedef_to_Quotient:
   assumes "type_definition Rep Abs {x. P x}"
   and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"
   shows "Quotient (invariant P) Abs Rep T"
@@ -262,16 +285,7 @@
     by (auto intro!: QuotientI simp: invariant_def fun_eq_iff)
 qed
 
-lemma invariant_type_to_Quotient:
-  assumes "type_definition Rep Abs {x. P x}"
-  and T_def: "T \<equiv> (\<lambda>x y. (invariant P) x x \<and> Abs x = y)"
-  shows "Quotient (invariant P) Abs Rep T"
-proof -
-  interpret type_definition Rep Abs "{x. P x}" by fact
-  from Rep Abs_inject Rep_inverse T_def show ?thesis by (auto intro!: QuotientI simp: invariant_def)
-qed
-
-lemma invariant_type_to_part_equivp:
+lemma open_typedef_to_part_equivp:
   assumes "type_definition Rep Abs {x. P x}"
   shows "part_equivp (invariant P)"
 proof (intro part_equivpI)
--- a/src/HOL/Quotient.thy	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/Quotient.thy	Wed Apr 04 20:45:19 2012 +0200
@@ -772,22 +772,30 @@
 using assms
 by (rule OOO_quotient3) auto
 
-subsection {* Invariant *}
+subsection {* Quotient3 to Quotient *}
+
+lemma Quotient3_to_Quotient:
+assumes "Quotient3 R Abs Rep"
+and "T \<equiv> \<lambda>x y. R x x \<and> Abs x = y"
+shows "Quotient R Abs Rep T"
+using assms unfolding Quotient3_def by (intro QuotientI) blast+
 
-lemma copy_type_to_Quotient3:
-  assumes "type_definition Rep Abs UNIV"
-  shows "Quotient3 (op =) Abs Rep"
-proof -
-  interpret type_definition Rep Abs UNIV by fact
-  from Abs_inject Rep_inverse show ?thesis by (auto intro!: Quotient3I)
-qed
-
-lemma invariant_type_to_Quotient3:
-  assumes "type_definition Rep Abs {x. P x}"
-  shows "Quotient3 (Lifting.invariant P) Abs Rep"
-proof -
-  interpret type_definition Rep Abs "{x. P x}" by fact
-  from Rep Abs_inject Rep_inverse show ?thesis by (auto intro!: Quotient3I simp: invariant_def)
+lemma Quotient3_to_Quotient_equivp:
+assumes q: "Quotient3 R Abs Rep"
+and T_def: "T \<equiv> \<lambda>x y. Abs x = y"
+and eR: "equivp R"
+shows "Quotient R Abs Rep T"
+proof (intro QuotientI)
+  fix a
+  show "Abs (Rep a) = a" using q by(rule Quotient3_abs_rep)
+next
+  fix a
+  show "R (Rep a) (Rep a)" using q by(rule Quotient3_rep_reflp)
+next
+  fix r s
+  show "R r s = (R r r \<and> R s s \<and> Abs r = Abs s)" using q by(rule Quotient3_rel[symmetric])
+next
+  show "T = (\<lambda>x y. R x x \<and> Abs x = y)" using T_def equivp_reflp[OF eR] by simp
 qed
 
 subsection {* ML setup *}
--- a/src/HOL/Quotient_Examples/Lift_DList.thy	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/Quotient_Examples/Lift_DList.thy	Wed Apr 04 20:45:19 2012 +0200
@@ -54,28 +54,13 @@
 proof -
   {
     fix x y
-    have "list_all2 cr_dlist x y \<Longrightarrow> 
-      List.map Abs_dlist x = y \<and> list_all2 (Lifting.invariant distinct) x x"
+    have "list_all2 cr_dlist x y \<Longrightarrow> x = List.map list_of_dlist y"
       unfolding list_all2_def cr_dlist_def by (induction x y rule: list_induct2') auto
   }
   note cr = this
-
   fix x :: "'a list list" and y :: "'a list list"
-  assume a: "(list_all2 cr_dlist OO Lifting.invariant distinct OO (list_all2 cr_dlist)\<inverse>\<inverse>) x y"
-  from a have l_x: "list_all2 (Lifting.invariant distinct) x x" by (auto simp add: cr)
-  from a have l_y: "list_all2 (Lifting.invariant distinct) y y" by (auto simp add: cr)
-  from a have m: "(Lifting.invariant distinct) (List.map Abs_dlist x) (List.map Abs_dlist y)" 
-    by (auto simp add: cr)
-
-  have "x = y" 
-  proof -
-    have m':"List.map Abs_dlist x = List.map Abs_dlist y" using m unfolding Lifting.invariant_def by simp
-    have dist: "\<And>l. list_all2 (Lifting.invariant distinct) l l \<Longrightarrow> !x. x \<in> (set l) \<longrightarrow> distinct x"
-      unfolding list_all2_def Lifting.invariant_def by (auto simp add: zip_same)
-    from dist[OF l_x] dist[OF l_y] have "inj_on Abs_dlist (set x \<union> set y)" by (intro inj_onI) 
-      (metis CollectI UnE Abs_dlist_inverse)
-    with m' show ?thesis by (rule map_inj_on)
-  qed
+  assume "(list_all2 cr_dlist OO Lifting.invariant distinct OO (list_all2 cr_dlist)\<inverse>\<inverse>) x y"
+  then have "x = y" by (auto dest: cr simp add: Lifting.invariant_def)
   then show "?thesis x y" unfolding Lifting.invariant_def by auto
 qed
 
--- a/src/HOL/TPTP/TPTP_Parser/tptp.lex	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/TPTP/TPTP_Parser/tptp.lex	Wed Apr 04 20:45:19 2012 +0200
@@ -3,8 +3,8 @@
 
  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.
+ * Could include %posarg to ensure that we'd start counting character positions
+   from 0, but it would punish performance.
  * %s AF F COMMENT; -- could improve by making stateful.
 
  Acknowledgements:
@@ -52,55 +52,55 @@
 %header (functor TPTPLexFun(structure Tokens: TPTP_TOKENS));
 %arg (file_name:string);
 
-printable_char            = .;
-viewable_char             = [.\n];
+percentage_sign           = "%";
+double_quote              = ["];
+do_char                   = ([^"]|[\\]["\\]);
+single_quote              = ['];
+sq_char                   = ([\040-\041\043-\126]|[\\]['\\]);
+sign                      = [+-];
+dot                       = [.];
+exponent                  = [Ee];
+slash                     = [/];
+zero_numeric              = [0];
+non_zero_numeric          = [1-9];
 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                    = \$;
+printable_char            = .;
+viewable_char             = [.\n];
+
+dot_decimal               = {dot}{numeric}+;
+
 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           = "%";
+lower_word                = {lower_alpha}{alpha_numeric}*;
+upper_word                = {upper_alpha}{alpha_numeric}*;
+dollar_dollar_word        = {ddollar}{lower_word};
+dollar_word               = {dollar}{lower_word};
 
-sq_char                   = ([\040-\041\043-\126]|[\\]['\\]);
+distinct_object           = {double_quote}{do_char}+{double_quote};
 
 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};
+comment_line              = {percentage_sign}[^\n]*;
+comment_block             = [/][\*]([^\*]*[\*]+[^/\*])*[^\*]*[\*]+[/];
+comment                   = ({comment_line}|{comment_block})+;
+
+eol                       = ("\013\010"|"\010"|"\013");
 
 %%
 
@@ -127,7 +127,7 @@
 ":="            => (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.FI(!linep,!col));
 "<=>"           => (col:=yypos-(!eolpos); T.IFF(!linep,!col));
 "=>"            => (col:=yypos-(!eolpos); T.IMPLIES(!linep,!col));
 "["             => (col:=yypos-(!eolpos); T.LBRKT(!linep,!col));
@@ -170,10 +170,14 @@
 
 "$ite_f"        => (col:=yypos-(!eolpos); T.ITE_F(!linep,!col));
 "$ite_t"        => (col:=yypos-(!eolpos); T.ITE_T(!linep,!col));
+"$let_tf"        => (col:=yypos-(!eolpos); T.LET_TF(!linep,!col));
+"$let_ff"        => (col:=yypos-(!eolpos); T.LET_FF(!linep,!col));
+"$let_ft"        => (col:=yypos-(!eolpos); T.LET_FT(!linep,!col));
+"$let_tt"        => (col:=yypos-(!eolpos); T.LET_TT(!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));
+{lower_word}   => (col:=yypos-(!eolpos); T.LOWER_WORD(yytext,!linep,!col));
+{dollar_word}  => (col:=yypos-(!eolpos); T.DOLLAR_WORD(yytext,!linep,!col));
+{dollar_dollar_word}  => (col:=yypos-(!eolpos); T.DOLLAR_DOLLAR_WORD(yytext,!linep,!col));
 
 "+"           => (col:=yypos-(!eolpos); T.PLUS(!linep,!col));
 "*"           => (col:=yypos-(!eolpos); T.TIMES(!linep,!col));
--- a/src/HOL/TPTP/TPTP_Parser/tptp.yacc	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/TPTP/TPTP_Parser/tptp.yacc	Wed Apr 04 20:45:19 2012 +0200
@@ -21,10 +21,13 @@
   | "unknown" => Role_Unknown
   | thing => raise (UNRECOGNISED_ROLE thing)
 
+fun extract_quant_info (Quant (quantifier, vars, tptp_formula)) =
+  (quantifier, vars, tptp_formula)
+
 %%
 %name TPTP
 %term AMPERSAND | AT_SIGN | CARET | COLON | COMMA | EQUALS | EXCLAMATION
-  | LET | ARROW | IF | IFF | IMPLIES | INCLUDE
+  | LET | ARROW | FI | 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
@@ -36,10 +39,12 @@
   | 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
+  | DOLLAR_WORD of string | DOLLAR_DOLLAR_WORD of string
   | SUBTYPE | LET_TERM
   | THF | TFF | FOF | CNF
   | ITE_F | ITE_T
+  | LET_TF | LET_FF | LET_FT | LET_TT
+
 %nonterm
     annotations of annotation option
   | name of string
@@ -116,9 +121,6 @@
   | 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
@@ -144,8 +146,6 @@
   | 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
@@ -183,6 +183,17 @@
   | tptp_file of tptp_problem
   | tptp of tptp_problem
 
+  | thf_let_defn of tptp_let list
+  | tff_let of tptp_formula
+  | tff_let_term_defn of tptp_let list
+  | tff_let_formula_defn of tptp_let list
+  | tff_quantified_type of tptp_type
+  | tff_monotype of tptp_type
+  | tff_type_arguments of tptp_type list
+  | let_term of tptp_term
+  | atomic_defined_word of string
+  | atomic_system_word of string
+
 %pos int
 %eop EOF
 %noshift EOF
@@ -196,7 +207,7 @@
 
 %left AT_SIGN
 %nonassoc IFF XOR
-%right IMPLIES IF
+%right IMPLIES FI
 %nonassoc EQUALS NEQUALS
 %right VLINE NOR
 %left AMPERSAND NAND
@@ -218,88 +229,488 @@
 
  Parser for TPTP languages. Latest version of the language spec can
  be obtained from http://www.cs.miami.edu/~tptp/TPTP/SyntaxBNF.html
+ This implements version 5.3.0.
 *)
 
+tptp : tptp_file (( tptp_file ))
+
+tptp_file : tptp_input tptp_file (( tptp_input :: tptp_file ))
+          | COMMENT tptp_file    (( tptp_file ))
+          |                      (( [] ))
+
+tptp_input : annotated_formula (( annotated_formula ))
+           | include_           (( include_ ))
+
+annotated_formula : thf_annotated (( thf_annotated ))
+                  | tff_annotated (( tff_annotated ))
+                  | fof_annotated (( fof_annotated ))
+                  | cnf_annotated (( cnf_annotated ))
+
+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)
+))
+
 annotations : COMMA general_term optional_info (( SOME (general_term, optional_info) ))
             |                                  (( NONE ))
 
-optional_info : COMMA useful_info (( useful_info ))
-              |                   (( [] ))
+formula_role : LOWER_WORD (( classify_role LOWER_WORD ))
+
+
+(* THF formulas *)
+
+thf_formula : thf_logic_formula (( thf_logic_formula ))
+            | thf_sequent       (( thf_sequent ))
+
+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          (( Type_fmla thf_subtype ))
+
+thf_binary_formula : thf_binary_pair   (( thf_binary_pair ))
+                   | thf_binary_tuple  (( thf_binary_tuple ))
+                   | thf_binary_type   (( Type_fmla thf_binary_type ))
 
-useful_info : general_list (( general_list ))
+thf_binary_pair : thf_unitary_formula thf_pair_connective thf_unitary_formula ((
+  Fmla (thf_pair_connective, [thf_unitary_formula1, thf_unitary_formula2])
+))
+
+thf_binary_tuple : thf_or_formula    (( thf_or_formula ))
+                 | thf_and_formula   (( thf_and_formula ))
+                 | thf_apply_formula (( thf_apply_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_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_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_unitary_formula : thf_quantified_formula (( thf_quantified_formula ))
+                    | thf_unary_formula      (( thf_unary_formula ))
+                    | thf_atom               (( thf_atom ))
+                    | thf_conditional        (( thf_conditional ))
+                    | thf_let                (( thf_let ))
+                    | LPAREN thf_logic_formula RPAREN  (( thf_logic_formula ))
 
-general_list : LBRKT general_terms RBRKT (( general_terms ))
-             | LBRKT RBRKT               (( [] ))
+thf_quantified_formula : thf_quantifier LBRKT thf_variable_list RBRKT COLON thf_unitary_formula ((
+  Quant (thf_quantifier, thf_variable_list, thf_unitary_formula)
+))
+
+thf_variable_list : thf_variable                          (( [thf_variable] ))
+                  | thf_variable COMMA thf_variable_list  (( thf_variable :: thf_variable_list ))
+
+thf_variable : thf_typed_variable (( thf_typed_variable ))
+             | variable_          (( (variable_, NONE) ))
+
+thf_typed_variable : variable_ COLON thf_top_level_type (( (variable_, SOME thf_top_level_type) ))
+
+thf_unary_formula : thf_unary_connective LPAREN thf_logic_formula RPAREN ((
+  Fmla (thf_unary_connective, [thf_logic_formula])
+))
+
+thf_atom : term          (( Atom (THF_Atom_term term) ))
+         | thf_conn_term (( Atom (THF_Atom_conn_term thf_conn_term) ))
 
-general_terms : general_term COMMA general_terms (( general_term :: general_terms ))
-              | general_term                     (( [general_term] ))
+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_let : LET_TF LPAREN thf_let_defn COMMA thf_formula RPAREN ((
+  Let (thf_let_defn, thf_formula)
+))
+
+(*FIXME here could check that fmla is of right form (TPTP BNF L130-134)*)
+thf_let_defn : thf_quantified_formula ((
+  let
+    val (_, vars, fmla) = extract_quant_info thf_quantified_formula
+  in [Let_fmla (hd vars, fmla)]
+  end
+))
+
+thf_type_formula : thf_typeable_formula COLON thf_top_level_type (( (thf_typeable_formula, thf_top_level_type) ))
+
+thf_typeable_formula : thf_atom                         (( thf_atom ))
+                     | LPAREN thf_logic_formula RPAREN  (( thf_logic_formula ))
+
+thf_subtype : constant SUBTYPE constant (( Subtype(constant1, constant2) ))
+
+thf_top_level_type : thf_logic_formula (( Fmla_type thf_logic_formula ))
 
-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 ))
+thf_unitary_type : thf_unitary_formula (( Fmla_type thf_unitary_formula ))
+
+thf_binary_type : thf_mapping_type (( thf_mapping_type ))
+                | thf_xprod_type   (( thf_xprod_type ))
+                | thf_union_type   (( thf_union_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_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_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_sequent : thf_tuple GENTZEN_ARROW thf_tuple  (( Sequent(thf_tuple1, thf_tuple2) ))
+            | LPAREN thf_sequent RPAREN          (( thf_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 ))
+
+
+(* TFF Formulas *)
+
+tff_formula : tff_logic_formula  (( tff_logic_formula ))
+            | tff_typed_atom     (( Atom (TFF_Typed_Atom tff_typed_atom) ))
+            | tff_sequent        (( tff_sequent ))
+
+tff_logic_formula : tff_binary_formula   (( tff_binary_formula ))
+                  | tff_unitary_formula  (( tff_unitary_formula ))
+
+tff_binary_formula : tff_binary_nonassoc (( tff_binary_nonassoc ))
+                   | tff_binary_assoc    (( tff_binary_assoc ))
+
+tff_binary_nonassoc : tff_unitary_formula binary_connective tff_unitary_formula (( Fmla (binary_connective, [tff_unitary_formula1, tff_unitary_formula2]) ))
+
+tff_binary_assoc : tff_or_formula  (( tff_or_formula ))
+                 | tff_and_formula (( tff_and_formula ))
 
-atomic_word : LOWER_WORD    (( LOWER_WORD ))
-            | SINGLE_QUOTED (( SINGLE_QUOTED ))
-            | THF           (( "thf" ))
-            | TFF           (( "tff" ))
-            | FOF           (( "fof" ))
-            | CNF           (( "cnf" ))
-            | INCLUDE       (( "include" ))
+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_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_unitary_formula : tff_quantified_formula           (( tff_quantified_formula ))
+                    | tff_unary_formula                (( tff_unary_formula ))
+                    | atomic_formula                   (( atomic_formula ))
+                    | tff_conditional                  (( tff_conditional ))
+                    | tff_let                          (( tff_let ))
+                    | LPAREN tff_logic_formula RPAREN  (( tff_logic_formula ))
+
+tff_quantified_formula : fol_quantifier LBRKT tff_variable_list RBRKT COLON tff_unitary_formula ((
+  Quant (fol_quantifier, tff_variable_list, tff_unitary_formula)
+))
+
+tff_variable_list : tff_variable                         (( [tff_variable] ))
+                  | tff_variable COMMA tff_variable_list (( tff_variable :: tff_variable_list ))
 
-variable_ : UPPER_WORD  (( UPPER_WORD ))
+tff_variable : tff_typed_variable (( tff_typed_variable ))
+             | variable_          (( (variable_, NONE) ))
+
+tff_typed_variable : variable_ COLON tff_atomic_type (( (variable_, SOME tff_atomic_type) ))
+
+tff_unary_formula : unary_connective tff_unitary_formula (( Fmla (unary_connective, [tff_unitary_formula]) ))
+                  | fol_infix_unary                      (( fol_infix_unary ))
+
+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)
+))
 
-general_function: atomic_word LPAREN general_terms RPAREN (( Application (atomic_word, general_terms) ))
+tff_let : LET_TF LPAREN tff_let_term_defn COMMA tff_formula RPAREN ((Let (tff_let_term_defn, tff_formula) ))
+        | LET_FF LPAREN tff_let_formula_defn COMMA tff_formula RPAREN (( Let (tff_let_formula_defn, tff_formula) ))
+
+(*FIXME here could check that fmla is of right form (TPTP BNF L210-214)*)
+(*FIXME why "term" if using "Let_fmla"?*)
+tff_let_term_defn : tff_quantified_formula ((
+  let
+    val (_, vars, fmla) = extract_quant_info tff_quantified_formula
+  in [Let_fmla (hd vars, fmla)]
+  end
+))
 
-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 ))
+(*FIXME here could check that fmla is of right form (TPTP BNF L215-217)*)
+tff_let_formula_defn : tff_quantified_formula ((
+  let
+    val (_, vars, fmla) = extract_quant_info tff_quantified_formula
+  in [Let_fmla (hd vars, fmla)]
+  end
+))
+
+tff_sequent : tff_tuple GENTZEN_ARROW tff_tuple (( Sequent (tff_tuple1, tff_tuple2) ))
+            | LPAREN tff_sequent RPAREN         (( tff_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_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_untyped_atom : functor_       (( (functor_, NONE) ))
+                 | system_functor (( (system_functor, NONE) ))
+
+tff_top_level_type : tff_atomic_type     (( tff_atomic_type ))
+                   | tff_mapping_type    (( tff_mapping_type ))
+                   | tff_quantified_type (( tff_quantified_type ))
+
+tff_quantified_type : DEP_PROD LBRKT tff_variable_list RBRKT COLON tff_monotype ((
+       Fmla_type (Quant (Dep_Prod, tff_variable_list, Type_fmla tff_monotype))
+))
+                    | LPAREN tff_quantified_type RPAREN (( tff_quantified_type ))
+
+tff_monotype : tff_atomic_type                (( tff_atomic_type ))
+             | LPAREN tff_mapping_type RPAREN (( tff_mapping_type ))
+
+tff_unitary_type : tff_atomic_type               (( tff_atomic_type ))
+                 | LPAREN tff_xprod_type RPAREN  (( tff_xprod_type ))
+
+tff_atomic_type : atomic_word   (( Atom_type atomic_word ))
+                | defined_type  (( Defined_type defined_type ))
+                | atomic_word LPAREN tff_type_arguments RPAREN (( Fmla_type (Fmla (Uninterpreted atomic_word, (map Type_fmla tff_type_arguments))) ))
+                | variable_ (( Fmla_type (Pred (Interpreted_ExtraLogic Apply, [Term_Var variable_])) ))
 
-number : integer          (( (Int_num, integer) ))
-       | REAL             (( (Real_num, REAL) ))
-       | RATIONAL         (( (Rat_num, RATIONAL) ))
+tff_type_arguments : tff_atomic_type   (( [tff_atomic_type]  ))
+                   | tff_atomic_type COMMA tff_type_arguments (( tff_atomic_type :: tff_type_arguments ))
+
+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_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 ))
+
+
+(* FOF Formulas *)
+
+fof_formula : fof_logic_formula  (( fof_logic_formula ))
+            | fof_sequent        (( fof_sequent ))
+
+fof_logic_formula : fof_binary_formula   (( fof_binary_formula ))
+                  | fof_unitary_formula  (( fof_unitary_formula ))
+
+fof_binary_formula : fof_binary_nonassoc (( fof_binary_nonassoc ))
+                   | fof_binary_assoc    (( fof_binary_assoc ))
 
-integer: UNSIGNED_INTEGER (( UNSIGNED_INTEGER ))
-       | SIGNED_INTEGER   (( SIGNED_INTEGER ))
+fof_binary_nonassoc : fof_unitary_formula binary_connective fof_unitary_formula ((
+  Fmla (binary_connective, [fof_unitary_formula1, fof_unitary_formula2] )
+))
+
+fof_binary_assoc : fof_or_formula  (( fof_or_formula ))
+                 | fof_and_formula (( fof_and_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_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_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_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)
+))
 
-file_name : SINGLE_QUOTED (( SINGLE_QUOTED ))
+fof_variable_list : variable_                          (( [variable_] ))
+                  | variable_ COMMA fof_variable_list  (( variable_ :: fof_variable_list ))
+
+fof_unary_formula : unary_connective fof_unitary_formula (( Fmla (unary_connective, [fof_unitary_formula]) ))
+                  | fol_infix_unary                      (( fol_infix_unary ))
+
+fof_sequent : fof_tuple GENTZEN_ARROW fof_tuple (( Sequent (fof_tuple1, fof_tuple2) ))
+            | LPAREN fof_sequent RPAREN         (( fof_sequent ))
+
+fof_tuple : LBRKT RBRKT                 (( [] ))
+          | LBRKT fof_tuple_list RBRKT  (( fof_tuple_list ))
+
+fof_tuple_list : fof_logic_formula                      (( [fof_logic_formula] ))
+               | fof_logic_formula COMMA fof_tuple_list (( fof_logic_formula :: fof_tuple_list ))
+
+
+(* CNF Formulas *)
+
+cnf_formula : LPAREN disjunction RPAREN  (( disjunction ))
+            | disjunction                (( disjunction ))
+
+disjunction : literal                    (( literal ))
+            | disjunction VLINE literal  (( Fmla (Interpreted_Logic Or, [disjunction, literal]) ))
+
+literal : atomic_formula        (( atomic_formula ))
+        | TILDE atomic_formula  (( Fmla (Interpreted_Logic Not, [atomic_formula]) ))
+        | fol_infix_unary       (( fol_infix_unary ))
+
+
+(* Special Formulas  *)
+
+thf_conn_term : thf_pair_connective   (( thf_pair_connective ))
+              | assoc_connective      (( assoc_connective ))
+              | thf_unary_connective  (( thf_unary_connective ))
+
+fol_infix_unary : term infix_inequality term (( Pred (infix_inequality, [term1, term2]) ))
+
+
+(* Connectives - THF *)
 
-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 ))
+thf_quantifier : fol_quantifier (( fol_quantifier ))
+               | CARET          (( Lambda ))
+               | DEP_PROD       (( Dep_Prod ))
+               | DEP_SUM        (( Dep_Sum ))
+               | INDEF_CHOICE   (( Epsilon ))
+               | DEFIN_CHOICE   (( Iota ))
+
+thf_pair_connective : infix_equality    (( infix_equality ))
+                    | infix_inequality  (( infix_inequality ))
+                    | binary_connective (( binary_connective ))
+
+thf_unary_connective : unary_connective (( unary_connective ))
+                     | OPERATOR_FORALL  (( Interpreted_Logic Op_Forall ))
+                     | OPERATOR_EXISTS  (( Interpreted_Logic Op_Exists ))
+
+
+(* Connectives - THF and TFF
+instead of subtype_sign use token SUBTYPE
+*)
+
+
+(* Connectives - FOF *)
 
-system_type : ATOMIC_SYSTEM_WORD (( ATOMIC_SYSTEM_WORD ))
+fol_quantifier : EXCLAMATION  (( Forall ))
+               | QUESTION     (( Exists ))
+
+binary_connective : IFF       (( Interpreted_Logic Iff ))
+                  | IMPLIES   (( Interpreted_Logic If ))
+                  | FI        (( Interpreted_Logic Fi ))
+                  | XOR       (( Interpreted_Logic Xor ))
+                  | NOR       (( Interpreted_Logic Nor ))
+                  | NAND      (( Interpreted_Logic Nand ))
+
+assoc_connective : VLINE      (( Interpreted_Logic Or ))
+                 | AMPERSAND  (( Interpreted_Logic And ))
 
-defined_type : ATOMIC_DEFINED_WORD ((
-  case ATOMIC_DEFINED_WORD of
-    "$i" => Type_Ind
+unary_connective : TILDE (( Interpreted_Logic Not ))
+
+
+(* The sequent arrow
+use token GENTZEN_ARROW
+*)
+
+
+(* Types for THF and TFF *)
+
+defined_type : atomic_defined_word ((
+  case atomic_defined_word of
+    "$oType" => Type_Bool
   | "$o" => Type_Bool
   | "$iType" => Type_Ind
-  | "$oType" => Type_Bool
-  | "$int" => Type_Int
+  | "$i" => Type_Ind
+  | "$tType" => Type_Type
   | "$real" => Type_Real
   | "$rat" => Type_Rat
-  | "$tType" => Type_Type
+  | "$int" => Type_Int
   | thing => raise UNRECOGNISED_SYMBOL ("defined_type", thing)
 ))
 
+system_type : atomic_system_word (( atomic_system_word ))
+
+
+(* First-order atoms *)
+
+atomic_formula : plain_atomic_formula   (( plain_atomic_formula ))
+               | defined_atomic_formula (( defined_atomic_formula ))
+               | system_atomic_formula  (( system_atomic_formula ))
+
+plain_atomic_formula : plain_term (( Pred plain_term ))
+
+defined_atomic_formula : defined_plain_formula (( defined_plain_formula ))
+                       | defined_infix_formula (( defined_infix_formula ))
+
+defined_plain_formula : defined_plain_term (( Pred defined_plain_term ))
+
+(*FIXME not used*)
+defined_prop : atomic_defined_word ((
+  case atomic_defined_word of
+    "$true"  => "$true"
+  | "$false" => "$false"
+  | thing => raise UNRECOGNISED_SYMBOL ("defined_prop", thing)
+))
+
+(*FIXME not used*)
+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_infix_formula : term defined_infix_pred term ((Pred (defined_infix_pred, [term1, term2])))
+
+defined_infix_pred : infix_equality  (( infix_equality ))
+
+infix_equality : EQUALS    (( Interpreted_Logic Equals ))
+
+infix_inequality : NEQUALS (( Interpreted_Logic NEquals ))
+
+system_atomic_formula : system_term  (( Pred system_term ))
+
+
+(* First-order terms *)
+
+term : function_term     (( function_term ))
+     | variable_         (( Term_Var variable_ ))
+     | conditional_term  (( conditional_term ))
+     | let_term          (( let_term ))
+
+function_term : plain_term    (( Term_Func plain_term ))
+              | defined_term  (( defined_term ))
+              | system_term   (( Term_Func system_term ))
+
+plain_term : constant                          (( (constant, []) ))
+           | functor_ LPAREN arguments RPAREN  (( (functor_, arguments) ))
+
+constant : functor_ (( functor_ ))
+
 functor_ : atomic_word (( Uninterpreted atomic_word ))
 
-arguments : term                  (( [term] ))
-          | term COMMA arguments  (( term :: arguments ))
+defined_term : defined_atom        (( defined_atom ))
+             | defined_atomic_term (( defined_atomic_term ))
+
+defined_atom : number          (( Term_Num number ))
+             | DISTINCT_OBJECT (( Term_Distinct_Object DISTINCT_OBJECT ))
+
+defined_atomic_term : defined_plain_term (( Term_Func defined_plain_term ))
 
-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_plain_term : defined_constant                        (( (defined_constant, []) ))
+                   | defined_functor LPAREN arguments RPAREN (( (defined_functor, arguments) ))
+
+defined_constant : defined_functor (( defined_functor ))
 
-defined_functor : ATOMIC_DEFINED_WORD ((
-  case ATOMIC_DEFINED_WORD of
-    "$sum" => Interpreted_ExtraLogic Sum
+(*FIXME would be nicer to split these up*)
+defined_functor : atomic_defined_word ((
+  case atomic_defined_word of
+    "$uminus" => Interpreted_ExtraLogic UMinus
+  | "$sum" => Interpreted_ExtraLogic Sum
   | "$difference" => Interpreted_ExtraLogic Difference
   | "$product" => Interpreted_ExtraLogic Product
   | "$quotient" => Interpreted_ExtraLogic Quotient
@@ -316,7 +727,6 @@
   | "$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
@@ -339,296 +749,46 @@
   | "$is_int" => Interpreted_ExtraLogic Is_Int
   | "$is_rat" => Interpreted_ExtraLogic Is_Rat
 
+  | "$distinct" => Interpreted_ExtraLogic Distinct
+
   | thing => raise UNRECOGNISED_SYMBOL ("defined_functor", thing)
 ))
 
-defined_constant : defined_functor (( defined_functor ))
+system_term : system_constant                         (( (system_constant, []) ))
+            | system_functor LPAREN arguments RPAREN  (( (system_functor, arguments) ))
+
+system_constant : system_functor (( system_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 ))
+system_functor : atomic_system_word (( System atomic_system_word ))
+
+variable_ : UPPER_WORD  (( UPPER_WORD ))
+
+arguments : term                  (( [term] ))
+          | term COMMA arguments  (( term :: arguments ))
 
 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 ))
+let_term : LET_FT LPAREN tff_let_formula_defn COMMA term RPAREN ((Term_Let (tff_let_formula_defn, term) ))
+         | LET_TT LPAREN tff_let_term_defn COMMA term RPAREN ((Term_Let (tff_let_term_defn, term) ))
 
 
-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 ))
+(* Formula sources
+Don't currently use following non-terminals:
+source, sources, dag_source, inference_record, inference_rule, parent_list,
+parent_info, parent_details, internal_source, intro_type, external_source,
+file_source, file_info, theory, theory_name, creator_source, creator_name.
+*)
 
-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 ))
+
+(* Useful info fields *)
 
-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)
-))
+optional_info : COMMA useful_info (( useful_info ))
+              |                   (( [] ))
 
-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 ))
+useful_info : general_list (( general_list ))
 
 include_ : INCLUDE LPAREN file_name formula_selection RPAREN PERIOD ((
   Include (file_name, formula_selection)
@@ -640,14 +800,57 @@
 name_list : name COMMA name_list   (( name :: name_list ))
           | name                   (( [name] ))
 
+
+(* Non-logical data *)
+
+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 ))
+
+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 ))
+
+general_function: atomic_word LPAREN general_terms RPAREN (( Application (atomic_word, general_terms) ))
+
+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 ))
+
+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 purpose *)
+
 name : atomic_word (( atomic_word ))
      | integer     (( integer ))
 
-tptp_input : annotated_formula (( annotated_formula ))
-           | include_           (( include_ ))
+atomic_word : LOWER_WORD    (( LOWER_WORD ))
+            | SINGLE_QUOTED (( SINGLE_QUOTED ))
+            | THF           (( "thf" ))
+            | TFF           (( "tff" ))
+            | FOF           (( "fof" ))
+            | CNF           (( "cnf" ))
+            | INCLUDE       (( "include" ))
+
+atomic_defined_word : DOLLAR_WORD (( DOLLAR_WORD ))
 
-tptp_file : tptp_input tptp_file (( tptp_input :: tptp_file ))
-          | COMMENT tptp_file    (( tptp_file ))
-          |                      (( [] ))
+atomic_system_word : DOLLAR_DOLLAR_WORD (( DOLLAR_DOLLAR_WORD ))
+
+integer: UNSIGNED_INTEGER (( UNSIGNED_INTEGER ))
+       | SIGNED_INTEGER   (( SIGNED_INTEGER ))
 
-tptp : tptp_file (( tptp_file ))
+number : integer          (( (Int_num, integer) ))
+       | REAL             (( (Real_num, REAL) ))
+       | RATIONAL         (( (Rat_num, RATIONAL) ))
+
+file_name : SINGLE_QUOTED (( SINGLE_QUOTED ))
--- a/src/HOL/TPTP/TPTP_Parser/tptp_interpret.ML	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_interpret.ML	Wed Apr 04 20:45:19 2012 +0200
@@ -13,9 +13,9 @@
   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
+  (*mapping from TPTP types to Isabelle/HOL types. This map works over all
+  base types (i.e. TPTP functions must be interpreted as Isabelle/HOL functions).
+  The map must be total wrt TPTP types. Later on there could be a configuration
   option to make a map extensible.*)
   type type_map = (TPTP_Syntax.tptp_type * typ) list
 
@@ -33,7 +33,7 @@
   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
+  while interpreting that statement, and a map from TPTP 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
@@ -53,7 +53,7 @@
      generative_type_interpretation : bool,
      generative_const_interpretation : bool*)}
 
-  (*map types form THF to Isabelle/HOL*)
+  (*map types form TPTP to Isabelle/HOL*)
   val interpret_type : config -> theory -> type_map ->
     TPTP_Syntax.tptp_type -> typ
 
@@ -68,11 +68,11 @@
   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
+    type_map = mapping of TPTP-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.
+      imported TPTP problem in Isar.
     const_map = as previous, but for constants.
-    path_prefix = path where THF problems etc are located (to help the "include"
+    path_prefix = path where TPTP problems etc are located (to help the "include"
       directive find the files.
     thy = theory where interpreted info will be stored.
   *)
@@ -93,8 +93,8 @@
     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
+        This is "true" if we are running this over a fresh TPTP problem, but
+        "false" if we are calling this _during_ the interpretation of a TPTP file
         (i.e. when interpreting an "include" directive.
       config = config
       path_prefix = " "
@@ -118,13 +118,13 @@
   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.*)
+  (*Generates a fresh Isabelle/HOL type for interpreting a TPTP 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.*)
+  all TPTP problems.*)
   val get_file_list : Path.T -> Path.T list
 
   type manifest = TPTP_Problem_Name.problem_name * tptp_general_meaning
@@ -245,13 +245,13 @@
       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))) =
+  | fmlatype_to_type (Type_fmla (Fn_type (ty1, ty2))) =
       let
         val ty1' = case ty1 of
-            Fn_type _ => fmlatype_to_type (THF_type ty1)
+            Fn_type _ => fmlatype_to_type (Type_fmla ty1)
           | Fmla_type ty1 => fmlatype_to_type ty1
         val ty2' = case ty2 of
-            Fn_type _ => fmlatype_to_type (THF_type ty2)
+            Fn_type _ => fmlatype_to_type (Type_fmla ty2)
           | Fmla_type ty2 => fmlatype_to_type ty2
       in Fn_type (ty1', ty2') end
 
@@ -327,7 +327,7 @@
   (Const (str, Term.dummyT) $ Bound 0 $ Bound 1)
    |> (Term.absdummy Term.dummyT #> Term.absdummy Term.dummyT)
 
-fun dummy_THF () =
+fun dummy_term () =
   HOLogic.choice_const Term.dummyT $ Term.absdummy Term.dummyT @{term "True"}
 
 fun interpret_symbol config thy language const_map symb =
@@ -357,7 +357,7 @@
         | 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 ()
+        | _ => dummy_term ()
         )
    | Interpreted_Logic logic_symbol =>
        (case logic_symbol of
@@ -523,7 +523,7 @@
         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*)
+          else dummy_term () (*FIXME: not yet supporting rationals and reals*)
       in (num_term, thy) end
   | Term_Distinct_Object str =>
       let
@@ -651,14 +651,19 @@
         | 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*)
-  *)
+    | Conditional (fmla1, fmla2, fmla3) =>
+        let
+          val interp = interpret_formula config language const_map var_types type_map
+          val (((fmla1', fmla2'), fmla3'), thy') =
+            interp fmla1 thy
+            ||>> interp fmla2
+            ||>> interp fmla3
+        in (HOLogic.mk_conj (HOLogic.mk_imp (fmla1', fmla2'),
+                             HOLogic.mk_imp (HOLogic.mk_not fmla1', fmla3')),
+            thy')
+        end
+    | Let (tptp_let_list, tptp_formula) => (*FIXME not yet implemented*)
+        raise MISINTERPRET_FORMULA ("'Let' not yet implemented", tptp_fmla)
     | Atom tptp_atom =>
         (case tptp_atom of
           TFF_Typed_Atom (symbol, tptp_type_opt) =>
@@ -669,7 +674,7 @@
              type_map tptp_term
         | THF_Atom_conn_term symbol =>
             (interpret_symbol config thy language const_map symbol, thy))
-    | THF_type _ => (*FIXME unsupported*)
+    | Type_fmla _ => (*FIXME unsupported*)
          raise MISINTERPRET_FORMULA
           ("Cannot interpret types as formulas", tptp_fmla)
     | THF_typing (tptp_formula, _) => (*FIXME ignoring type info*)
@@ -748,7 +753,7 @@
            in
              case tptp_ty of
                Defined_type Type_Type => (*add new type*)
-                 (*generate an Isabelle/HOL type to interpret this THF type,
+                 (*generate an Isabelle/HOL type to interpret this TPTP type,
                  and add it to both the Isabelle/HOL theory and to the type_map*)
                   let
                     val (type_map', thy') =
--- a/src/HOL/TPTP/TPTP_Parser/tptp_lexyacc.ML	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_lexyacc.ML	Wed Apr 04 20:45:19 2012 +0200
@@ -13,6 +13,10 @@
 sig
 type ('a,'b) token
 type svalue
+val LET_TT:  'a * 'a -> (svalue,'a) token
+val LET_FT:  'a * 'a -> (svalue,'a) token
+val LET_FF:  'a * 'a -> (svalue,'a) token
+val LET_TF:  'a * 'a -> (svalue,'a) token
 val ITE_T:  'a * 'a -> (svalue,'a) token
 val ITE_F:  'a * 'a -> (svalue,'a) token
 val CNF:  'a * 'a -> (svalue,'a) token
@@ -21,8 +25,8 @@
 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 DOLLAR_DOLLAR_WORD: (string) *  'a * 'a -> (svalue,'a) token
+val DOLLAR_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
@@ -76,7 +80,7 @@
 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 FI:  '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
@@ -103,8 +107,8 @@
 
  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.
+ * Could include %posarg to ensure that we'd start counting character positions
+   from 0, but it would punish performance.
  * %s AF F COMMENT; -- could improve by making stateful.
 
  Acknowledgements:
@@ -170,9 +174,9 @@
 \\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\
+"\000\000\000\000\000\000\000\000\000\144\146\000\000\145\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\144\140\134\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\
@@ -843,11 +847,11 @@
  (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\
+\\000\000\000\000\132\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\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\129\103\103\125\103\103\119\103\103\109\103\103\103\
 \\103\103\103\103\104\103\103\103\103\103\103\000\000\000\000\000\
 \\000"
 ),
@@ -902,8 +906,8 @@
 \\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\103\103\103\103\110\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"
 ),
  (110, 
@@ -913,8 +917,8 @@
 \\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\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\111\103\103\103\103\103\103\000\000\000\000\000\
 \\000"
 ),
  (111, 
@@ -935,19 +939,19 @@
 \\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\
+\\000\103\103\103\103\103\116\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, 
+ (113, 
 "\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\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\103\103\103\103\103\115\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\114\103\103\103\103\103\103\000\000\000\000\000\
 \\000"
 ),
  (116, 
@@ -968,8 +972,8 @@
 \\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\103\103\103\103\103\103\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\120\103\103\103\103\103\103\000\000\000\000\000\
 \\000"
 ),
  (120, 
@@ -979,7 +983,18 @@
 \\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\
+\\000\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\103\000\000\000\000\000\
+\\000"
+),
+ (121, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\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\122\
+\\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"
 ),
@@ -987,72 +1002,127 @@
 "\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\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\
+\\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\124\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\123\103\103\103\103\103\103\000\000\000\000\000\
 \\000"
 ),
- (123, 
+ (125, 
 "\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\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\
+\\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\126\
+\\103\103\103\103\103\103\103\103\103\103\103\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"
+ (126, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\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\128\103\103\103\103\103\103\103\103\103\
+\\103\103\103\103\127\103\103\103\103\103\103\000\000\000\000\000\
+\\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\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"
+ (129, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\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\130\103\
+\\103\103\103\103\103\103\103\103\103\103\103\000\000\000\000\000\
+\\000"
 ),
  (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\
+\\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\131\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"
+),
+ (132, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\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\133\133\133\133\133\133\133\133\133\133\133\133\133\133\
+\\133\133\133\133\133\133\133\133\133\133\133\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\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\133\133\133\133\133\133\133\133\133\133\000\000\000\000\000\000\
+\\000\133\133\133\133\133\133\133\133\133\133\133\133\133\133\133\
+\\133\133\133\133\133\133\133\133\133\133\133\000\000\000\000\133\
+\\000\133\133\133\133\133\133\133\133\133\133\133\133\133\133\133\
+\\133\133\133\133\133\133\133\133\133\133\133\000\000\000\000\000\
+\\000"
+),
+ (134, 
+"\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\000\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\139\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135"
+),
+ (135, 
+"\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\138\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\136\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135"
+),
+ (136, 
+"\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\137\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\136\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135"
+),
+ (140, 
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\143\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\142\141\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\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\
+ (144, 
+"\000\000\000\000\000\000\000\000\000\144\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\144\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
@@ -1060,8 +1130,8 @@
 \\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\
+ (145, 
+"\000\000\000\000\000\000\000\000\000\000\146\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
@@ -1086,25 +1156,25 @@
 {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 283)], trans = 7},
+{fin = [(N 283)], trans = 8},
+{fin = [(N 283)], trans = 9},
+{fin = [(N 186),(N 283)], trans = 7},
+{fin = [(N 283)], trans = 11},
+{fin = [(N 198),(N 283)], trans = 7},
+{fin = [(N 283)], trans = 13},
+{fin = [(N 283)], trans = 14},
+{fin = [(N 283)], trans = 15},
+{fin = [(N 283)], trans = 16},
+{fin = [(N 283)], trans = 17},
+{fin = [(N 283)], trans = 18},
+{fin = [(N 206),(N 283)], trans = 7},
+{fin = [(N 283)], trans = 20},
+{fin = [(N 283)], trans = 21},
+{fin = [(N 190),(N 283)], trans = 7},
+{fin = [(N 283)], trans = 23},
+{fin = [(N 283)], trans = 24},
+{fin = [(N 194),(N 283)], trans = 7},
 {fin = [(N 25)], trans = 0},
 {fin = [(N 80)], trans = 0},
 {fin = [(N 50)], trans = 0},
@@ -1114,7 +1184,7 @@
 {fin = [(N 12)], trans = 0},
 {fin = [(N 78)], trans = 33},
 {fin = [(N 21)], trans = 0},
-{fin = [(N 283)], trans = 0},
+{fin = [(N 315)], trans = 0},
 {fin = [(N 38)], trans = 0},
 {fin = [(N 31)], trans = 37},
 {fin = [(N 48)], trans = 0},
@@ -1123,10 +1193,10 @@
 {fin = [(N 68)], trans = 0},
 {fin = [(N 41)], trans = 42},
 {fin = [(N 45)], trans = 0},
-{fin = [(N 277)], trans = 0},
+{fin = [(N 309)], trans = 0},
 {fin = [(N 27)], trans = 45},
 {fin = [(N 36)], trans = 0},
-{fin = [(N 286)], trans = 0},
+{fin = [(N 318)], trans = 0},
 {fin = [(N 126)], trans = 48},
 {fin = [], trans = 49},
 {fin = [(N 104)], trans = 49},
@@ -1155,11 +1225,11 @@
 {fin = [(N 55)], trans = 0},
 {fin = [(N 123)], trans = 74},
 {fin = [(N 58)], trans = 75},
-{fin = [(N 274)], trans = 0},
+{fin = [(N 306)], trans = 0},
 {fin = [(N 29)], trans = 0},
-{fin = [(N 268)], trans = 78},
+{fin = [(N 300)], trans = 78},
 {fin = [(N 76)], trans = 0},
-{fin = [(N 270)], trans = 0},
+{fin = [(N 302)], trans = 0},
 {fin = [(N 82)], trans = 0},
 {fin = [(N 52)], trans = 0},
 {fin = [], trans = 83},
@@ -1182,39 +1252,49 @@
 {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 290)], trans = 103},
+{fin = [(N 290)], trans = 104},
+{fin = [(N 290)], trans = 105},
+{fin = [(N 211),(N 290)], trans = 103},
+{fin = [(N 290)], trans = 107},
+{fin = [(N 231),(N 290)], trans = 103},
+{fin = [(N 290)], trans = 109},
+{fin = [(N 290)], trans = 110},
+{fin = [(N 290)], trans = 111},
+{fin = [(N 290)], trans = 112},
+{fin = [(N 290)], trans = 113},
+{fin = [(N 277),(N 290)], trans = 103},
+{fin = [(N 253),(N 290)], trans = 103},
+{fin = [(N 290)], trans = 116},
+{fin = [(N 269),(N 290)], trans = 103},
+{fin = [(N 261),(N 290)], trans = 103},
+{fin = [(N 290)], trans = 119},
+{fin = [(N 290)], trans = 120},
+{fin = [(N 290)], trans = 121},
+{fin = [(N 290)], trans = 122},
+{fin = [(N 245),(N 290)], trans = 103},
+{fin = [(N 238),(N 290)], trans = 103},
+{fin = [(N 290)], trans = 125},
+{fin = [(N 290)], trans = 126},
+{fin = [(N 226),(N 290)], trans = 103},
+{fin = [(N 216),(N 290)], trans = 103},
+{fin = [(N 290)], trans = 129},
+{fin = [(N 290)], trans = 130},
+{fin = [(N 221),(N 290)], trans = 103},
+{fin = [], trans = 132},
+{fin = [(N 298)], trans = 133},
+{fin = [], trans = 134},
+{fin = [], trans = 135},
+{fin = [], trans = 136},
+{fin = [(N 95)], trans = 135},
 {fin = [(N 95)], trans = 0},
-{fin = [], trans = 126},
-{fin = [(N 33)], trans = 130},
-{fin = [(N 280)], trans = 0},
+{fin = [], trans = 136},
+{fin = [(N 33)], trans = 140},
+{fin = [(N 312)], trans = 0},
 {fin = [(N 64)], trans = 0},
 {fin = [(N 18)], trans = 0},
-{fin = [(N 2)], trans = 134},
-{fin = [(N 7)], trans = 135},
+{fin = [(N 2)], trans = 144},
+{fin = [(N 7)], trans = 145},
 {fin = [(N 7)], trans = 0}])
 end
 structure StartStates =
@@ -1284,23 +1364,27 @@
 | 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))
+| 253 => (col:=yypos-(!eolpos); T.LET_TF(!linep,!col))
+| 261 => (col:=yypos-(!eolpos); T.LET_FF(!linep,!col))
+| 269 => (col:=yypos-(!eolpos); T.LET_FT(!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))
+| 277 => (col:=yypos-(!eolpos); T.LET_TT(!linep,!col))
+| 283 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.LOWER_WORD(yytext,!linep,!col) end
 | 29 => (col:=yypos-(!eolpos); T.COMMA(!linep,!col))
+| 290 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.DOLLAR_WORD(yytext,!linep,!col) end
+| 298 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.DOLLAR_DOLLAR_WORD(yytext,!linep,!col) end
+| 300 => (col:=yypos-(!eolpos); T.PLUS(!linep,!col))
+| 302 => (col:=yypos-(!eolpos); T.TIMES(!linep,!col))
+| 306 => (col:=yypos-(!eolpos); T.GENTZEN_ARROW(!linep,!col))
+| 309 => (col:=yypos-(!eolpos); T.SUBTYPE(!linep,!col))
 | 31 => (col:=yypos-(!eolpos); T.EQUALS(!linep,!col))
+| 312 => (col:=yypos-(!eolpos); T.DEP_PROD(!linep,!col))
+| 315 => (col:=yypos-(!eolpos); T.DEP_SUM(!linep,!col))
+| 318 => (col:=yypos-(!eolpos); T.LET_TERM(!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))
+| 41 => (col:=yypos-(!eolpos); T.FI(!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))
@@ -1392,6 +1476,9 @@
   | "unknown" => Role_Unknown
   | thing => raise (UNRECOGNISED_ROLE thing)
 
+fun extract_quant_info (Quant (quantifier, vars, tptp_formula)) =
+  (quantifier, vars, tptp_formula)
+
 
 end
 structure LrTable = Token.LrTable
@@ -1399,93 +1486,94 @@
 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\001\000\052\002\002\000\052\002\004\000\069\002\005\000\052\002\
+\\006\000\052\002\009\000\052\002\010\000\052\002\011\000\052\002\
+\\012\000\052\002\019\000\052\002\020\000\052\002\021\000\052\002\
+\\022\000\052\002\026\000\052\002\027\000\052\002\037\000\052\002\
+\\059\000\052\002\060\000\052\002\000\000\
+\\001\000\001\000\055\002\002\000\055\002\004\000\070\002\005\000\055\002\
+\\006\000\055\002\009\000\055\002\010\000\055\002\011\000\055\002\
+\\012\000\055\002\019\000\055\002\020\000\055\002\021\000\055\002\
+\\022\000\055\002\026\000\055\002\027\000\055\002\037\000\055\002\
+\\059\000\055\002\060\000\055\002\000\000\
+\\001\000\001\000\219\002\005\000\219\002\006\000\234\002\010\000\219\002\
+\\011\000\219\002\012\000\219\002\019\000\219\002\020\000\234\002\
+\\021\000\219\002\022\000\219\002\026\000\219\002\027\000\219\002\
+\\037\000\219\002\000\000\
+\\001\000\001\000\222\002\005\000\222\002\006\000\245\002\010\000\222\002\
+\\011\000\222\002\012\000\222\002\019\000\222\002\020\000\245\002\
+\\021\000\222\002\022\000\222\002\026\000\222\002\027\000\222\002\
+\\037\000\222\002\000\000\
+\\001\000\001\000\229\002\005\000\229\002\006\000\236\002\010\000\229\002\
+\\011\000\229\002\012\000\229\002\019\000\229\002\020\000\236\002\
+\\021\000\229\002\022\000\229\002\026\000\229\002\027\000\229\002\
+\\037\000\229\002\000\000\
+\\001\000\001\000\239\002\004\000\130\002\005\000\239\002\006\000\239\002\
+\\010\000\239\002\011\000\239\002\012\000\239\002\016\000\222\000\
+\\019\000\239\002\020\000\239\002\021\000\239\002\022\000\239\002\
+\\027\000\239\002\037\000\239\002\000\000\
+\\001\000\001\000\252\002\004\000\131\002\005\000\252\002\006\000\252\002\
+\\010\000\252\002\011\000\252\002\012\000\252\002\016\000\217\000\
+\\019\000\252\002\020\000\252\002\021\000\252\002\022\000\252\002\
+\\027\000\252\002\037\000\252\002\000\000\
+\\001\000\001\000\211\000\003\000\210\000\006\000\209\000\007\000\124\000\
+\\010\000\208\000\011\000\207\000\012\000\206\000\013\000\035\000\
+\\015\000\205\000\016\000\204\000\019\000\203\000\020\000\202\000\
+\\021\000\201\000\022\000\200\000\025\000\121\000\028\000\120\000\
+\\037\000\199\000\044\000\101\000\045\000\100\000\046\000\034\000\
+\\047\000\033\000\049\000\032\000\050\000\099\000\051\000\031\000\
+\\053\000\098\000\055\000\198\000\056\000\197\000\057\000\196\000\
+\\058\000\195\000\062\000\194\000\063\000\193\000\064\000\097\000\
+\\065\000\096\000\068\000\030\000\069\000\029\000\070\000\028\000\
+\\071\000\027\000\072\000\192\000\073\000\095\000\074\000\191\000\
+\\076\000\094\000\077\000\093\000\000\000\
+\\001\000\001\000\211\000\003\000\210\000\006\000\209\000\007\000\124\000\
+\\010\000\208\000\011\000\207\000\012\000\206\000\013\000\035\000\
+\\016\000\033\001\019\000\203\000\020\000\202\000\021\000\201\000\
+\\022\000\200\000\025\000\121\000\026\000\032\001\028\000\120\000\
+\\037\000\199\000\044\000\101\000\045\000\100\000\046\000\034\000\
+\\047\000\033\000\049\000\032\000\050\000\099\000\051\000\031\000\
+\\053\000\098\000\055\000\198\000\056\000\197\000\057\000\196\000\
+\\058\000\195\000\062\000\194\000\063\000\193\000\064\000\097\000\
+\\065\000\096\000\068\000\030\000\069\000\029\000\070\000\028\000\
+\\071\000\027\000\072\000\192\000\073\000\095\000\074\000\191\000\
+\\076\000\094\000\077\000\093\000\000\000\
+\\001\000\001\000\211\000\003\000\210\000\006\000\209\000\007\000\124\000\
+\\010\000\208\000\011\000\207\000\012\000\206\000\013\000\035\000\
+\\016\000\033\001\019\000\203\000\020\000\202\000\021\000\201\000\
+\\022\000\200\000\025\000\121\000\028\000\120\000\037\000\199\000\
+\\044\000\101\000\045\000\100\000\046\000\034\000\047\000\033\000\
+\\049\000\032\000\050\000\099\000\051\000\031\000\053\000\098\000\
+\\055\000\198\000\056\000\197\000\057\000\196\000\058\000\195\000\
+\\062\000\194\000\063\000\193\000\064\000\097\000\065\000\096\000\
+\\068\000\030\000\069\000\029\000\070\000\028\000\071\000\027\000\
+\\072\000\192\000\073\000\095\000\074\000\191\000\076\000\094\000\
+\\077\000\093\000\000\000\
+\\001\000\001\000\211\000\003\000\210\000\006\000\209\000\007\000\124\000\
+\\010\000\208\000\011\000\207\000\012\000\206\000\013\000\035\000\
+\\016\000\110\001\019\000\203\000\020\000\202\000\021\000\201\000\
+\\022\000\200\000\025\000\121\000\028\000\120\000\037\000\199\000\
+\\044\000\101\000\045\000\100\000\046\000\034\000\047\000\033\000\
+\\049\000\032\000\050\000\099\000\051\000\031\000\053\000\098\000\
+\\055\000\198\000\056\000\197\000\057\000\196\000\058\000\195\000\
+\\062\000\194\000\063\000\193\000\064\000\097\000\065\000\096\000\
+\\068\000\030\000\069\000\029\000\070\000\028\000\071\000\027\000\
+\\072\000\192\000\073\000\095\000\074\000\191\000\076\000\094\000\
+\\077\000\093\000\000\000\
+\\001\000\001\000\015\001\002\000\014\001\005\000\034\002\006\000\209\000\
+\\009\000\073\002\010\000\208\000\011\000\207\000\012\000\206\000\
+\\019\000\203\000\020\000\202\000\021\000\201\000\022\000\200\000\
+\\026\000\034\002\027\000\034\002\037\000\013\001\059\000\073\002\
+\\060\000\073\002\000\000\
+\\001\000\003\000\210\000\007\000\124\000\025\000\121\000\055\000\198\000\
+\\056\000\197\000\062\000\194\000\063\000\193\000\000\000\
+\\001\000\004\000\250\000\000\000\
+\\001\000\004\000\016\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\004\000\217\001\000\000\
+\\001\000\004\000\224\001\000\000\
+\\001\000\004\000\255\001\000\000\
+\\001\000\005\000\132\002\009\000\139\002\027\000\132\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\
@@ -1494,200 +1582,201 @@
 \\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\
+\\001\000\005\000\158\001\000\000\
+\\001\000\005\000\159\001\000\000\
+\\001\000\005\000\160\001\000\000\
+\\001\000\005\000\177\001\000\000\
+\\001\000\005\000\178\001\000\000\
+\\001\000\005\000\179\001\000\000\
+\\001\000\005\000\187\001\000\000\
+\\001\000\005\000\188\001\000\000\
+\\001\000\005\000\238\001\000\000\
+\\001\000\005\000\249\001\000\000\
+\\001\000\005\000\252\001\000\000\
+\\001\000\006\000\209\000\000\000\
+\\001\000\006\000\209\000\020\000\202\000\000\000\
+\\001\000\007\000\124\000\013\000\035\000\015\000\123\000\016\000\122\000\
+\\025\000\121\000\028\000\120\000\044\000\101\000\045\000\100\000\
+\\046\000\034\000\047\000\033\000\049\000\032\000\050\000\099\000\
+\\051\000\031\000\053\000\098\000\064\000\097\000\065\000\096\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\
+\\073\000\095\000\076\000\094\000\077\000\093\000\000\000\
+\\001\000\007\000\124\000\013\000\035\000\015\000\151\000\016\000\150\000\
+\\025\000\121\000\028\000\120\000\044\000\101\000\045\000\100\000\
+\\046\000\034\000\047\000\033\000\049\000\032\000\050\000\099\000\
+\\051\000\031\000\053\000\098\000\064\000\097\000\065\000\096\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\
+\\072\000\149\000\073\000\095\000\074\000\148\000\075\000\147\000\
+\\076\000\094\000\077\000\093\000\000\000\
+\\001\000\007\000\124\000\013\000\035\000\016\000\238\000\025\000\121\000\
+\\026\000\243\000\028\000\120\000\044\000\101\000\045\000\100\000\
+\\046\000\034\000\047\000\033\000\049\000\032\000\050\000\099\000\
+\\051\000\031\000\053\000\098\000\064\000\097\000\065\000\096\000\
+\\068\000\030\000\069\000\029\000\070\000\028\000\071\000\027\000\
+\\073\000\095\000\076\000\094\000\077\000\093\000\000\000\
+\\001\000\007\000\124\000\013\000\035\000\016\000\238\000\025\000\121\000\
+\\028\000\120\000\044\000\101\000\045\000\100\000\046\000\034\000\
+\\047\000\033\000\049\000\032\000\050\000\099\000\051\000\031\000\
+\\053\000\098\000\064\000\097\000\065\000\096\000\068\000\030\000\
+\\069\000\029\000\070\000\028\000\071\000\027\000\073\000\095\000\
+\\076\000\094\000\077\000\093\000\000\000\
+\\001\000\007\000\124\000\013\000\035\000\016\000\254\000\025\000\121\000\
+\\026\000\007\001\028\000\120\000\044\000\101\000\045\000\100\000\
+\\046\000\034\000\047\000\033\000\049\000\032\000\050\000\099\000\
+\\051\000\031\000\053\000\098\000\064\000\097\000\065\000\096\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\
+\\072\000\149\000\073\000\095\000\074\000\148\000\075\000\147\000\
+\\076\000\094\000\077\000\093\000\000\000\
+\\001\000\007\000\124\000\013\000\035\000\016\000\254\000\025\000\121\000\
+\\028\000\120\000\044\000\101\000\045\000\100\000\046\000\034\000\
+\\047\000\033\000\049\000\032\000\050\000\099\000\051\000\031\000\
+\\053\000\098\000\064\000\097\000\065\000\096\000\068\000\030\000\
+\\069\000\029\000\070\000\028\000\071\000\027\000\072\000\149\000\
+\\073\000\095\000\074\000\148\000\075\000\147\000\076\000\094\000\
+\\077\000\093\000\000\000\
+\\001\000\007\000\124\000\025\000\121\000\000\000\
+\\001\000\009\000\140\002\027\000\151\002\060\000\151\002\000\000\
+\\001\000\009\000\019\001\059\000\018\001\060\000\017\001\000\000\
+\\001\000\009\000\166\001\000\000\
+\\001\000\013\000\035\000\015\000\050\001\026\000\153\001\039\000\049\001\
+\\040\000\048\001\041\000\047\001\042\000\046\001\043\000\045\001\
+\\044\000\101\000\045\000\100\000\046\000\034\000\047\000\033\000\
+\\049\000\032\000\050\000\099\000\051\000\031\000\053\000\044\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\
+\\001\000\013\000\035\000\015\000\050\001\039\000\049\001\040\000\048\001\
+\\041\000\047\001\042\000\046\001\043\000\045\001\044\000\101\000\
+\\045\000\100\000\046\000\034\000\047\000\033\000\049\000\032\000\
+\\050\000\099\000\051\000\031\000\053\000\044\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\
+\\001\000\013\000\035\000\016\000\103\000\028\000\102\000\044\000\101\000\
+\\045\000\100\000\046\000\034\000\047\000\033\000\049\000\032\000\
+\\050\000\099\000\051\000\031\000\053\000\098\000\064\000\097\000\
+\\065\000\096\000\068\000\030\000\069\000\029\000\070\000\028\000\
+\\071\000\027\000\073\000\095\000\076\000\094\000\077\000\093\000\000\000\
+\\001\000\013\000\035\000\016\000\093\001\049\000\032\000\050\000\099\000\
+\\051\000\031\000\063\000\092\001\064\000\097\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\016\000\173\001\049\000\032\000\050\000\099\000\
+\\051\000\031\000\063\000\092\001\064\000\097\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\016\000\005\002\049\000\032\000\050\000\099\000\
+\\051\000\031\000\064\000\097\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\016\000\010\002\049\000\032\000\050\000\099\000\
+\\051\000\031\000\064\000\097\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\016\000\012\002\049\000\032\000\050\000\099\000\
+\\051\000\031\000\064\000\097\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\028\000\102\000\044\000\101\000\045\000\100\000\
+\\046\000\034\000\047\000\033\000\049\000\032\000\050\000\099\000\
+\\051\000\031\000\053\000\098\000\064\000\097\000\065\000\096\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\
+\\073\000\095\000\076\000\094\000\077\000\093\000\000\000\
+\\001\000\013\000\035\000\044\000\101\000\045\000\100\000\046\000\034\000\
+\\047\000\033\000\049\000\032\000\050\000\099\000\051\000\031\000\
+\\053\000\098\000\064\000\097\000\065\000\096\000\068\000\030\000\
+\\069\000\029\000\070\000\028\000\071\000\027\000\073\000\095\000\
+\\076\000\094\000\077\000\093\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\050\000\099\000\051\000\031\000\
+\\064\000\097\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\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\015\000\123\000\000\000\
+\\001\000\015\000\151\000\000\000\
+\\001\000\015\000\205\000\000\000\
+\\001\000\015\000\236\000\000\000\
+\\001\000\015\000\252\000\000\000\
+\\001\000\015\000\023\001\000\000\
+\\001\000\015\000\050\001\000\000\
+\\001\000\015\000\168\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\016\000\223\000\000\000\
+\\001\000\016\000\224\000\000\000\
+\\001\000\016\000\225\000\000\000\
+\\001\000\016\000\255\000\000\000\
+\\001\000\016\000\000\001\000\000\
+\\001\000\016\000\001\001\000\000\
+\\001\000\016\000\026\001\000\000\
+\\001\000\016\000\027\001\000\000\
+\\001\000\016\000\146\001\000\000\
+\\001\000\016\000\147\001\000\000\
+\\001\000\016\000\148\001\000\000\
+\\001\000\016\000\149\001\000\000\
+\\001\000\016\000\150\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\023\000\141\001\000\000\
+\\001\000\023\000\161\001\000\000\
+\\001\000\023\000\165\001\000\000\
+\\001\000\023\000\181\001\000\000\
+\\001\000\026\000\212\000\000\000\
+\\001\000\026\000\076\001\000\000\
+\\001\000\026\000\106\001\000\000\
+\\001\000\026\000\140\001\000\000\
+\\001\000\026\000\162\001\000\000\
+\\001\000\026\000\174\001\000\000\
+\\001\000\026\000\183\001\000\000\
+\\001\000\026\000\200\001\000\000\
+\\001\000\026\000\242\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\035\001\000\000\
+\\001\000\027\000\063\001\037\000\216\000\000\000\
+\\001\000\027\000\064\001\000\000\
+\\001\000\027\000\073\001\000\000\
+\\001\000\027\000\074\001\000\000\
+\\001\000\027\000\077\001\000\000\
+\\001\000\027\000\102\001\000\000\
+\\001\000\027\000\103\001\000\000\
+\\001\000\027\000\104\001\000\000\
+\\001\000\027\000\107\001\000\000\
+\\001\000\027\000\137\001\000\000\
+\\001\000\027\000\138\001\000\000\
+\\001\000\027\000\154\001\000\000\
+\\001\000\027\000\156\001\000\000\
+\\001\000\027\000\157\001\000\000\
+\\001\000\027\000\186\001\000\000\
+\\001\000\027\000\211\001\000\000\
+\\001\000\027\000\213\001\000\000\
+\\001\000\027\000\215\001\060\000\214\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\229\001\000\000\
+\\001\000\027\000\230\001\000\000\
+\\001\000\027\000\231\001\000\000\
+\\001\000\027\000\232\001\000\000\
+\\001\000\027\000\233\001\000\000\
+\\001\000\027\000\234\001\000\000\
+\\001\000\027\000\236\001\000\000\
+\\001\000\027\000\237\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\027\000\245\001\060\000\214\001\000\000\
+\\001\000\027\000\247\001\000\000\
+\\001\000\027\000\248\001\000\000\
+\\001\000\027\000\251\001\000\000\
+\\001\000\027\000\002\002\000\000\
+\\001\000\027\000\006\002\000\000\
+\\001\000\027\000\007\002\000\000\
+\\001\000\027\000\011\002\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\050\000\099\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\
+\\001\000\061\000\235\000\000\000\
+\\001\000\061\000\251\000\000\000\
+\\001\000\061\000\022\001\000\000\
 \\014\002\000\000\
 \\015\002\000\000\
 \\016\002\000\000\
-\\017\002\000\000\
+\\017\002\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\
 \\018\002\000\000\
 \\019\002\000\000\
 \\020\002\000\000\
@@ -1696,27 +1785,25 @@
 \\023\002\000\000\
 \\024\002\000\000\
 \\025\002\000\000\
+\\026\002\000\000\
 \\027\002\000\000\
 \\028\002\000\000\
-\\029\002\005\000\144\001\000\000\
+\\029\002\005\000\215\000\000\000\
 \\030\002\000\000\
 \\031\002\000\000\
-\\032\002\016\000\212\000\000\000\
+\\032\002\000\000\
 \\033\002\000\000\
-\\034\002\000\000\
 \\035\002\000\000\
-\\036\002\016\000\213\000\000\000\
+\\036\002\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\
+\\041\002\037\000\009\001\000\000\
+\\042\002\001\000\010\001\000\000\
+\\043\002\002\000\011\001\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\
@@ -1725,25 +1812,28 @@
 \\051\002\000\000\
 \\052\002\000\000\
 \\053\002\000\000\
+\\054\002\000\000\
 \\055\002\000\000\
 \\056\002\000\000\
-\\057\002\000\000\
+\\057\002\005\000\184\001\000\000\
 \\058\002\000\000\
+\\059\002\000\000\
+\\060\002\004\000\185\001\000\000\
+\\061\002\000\000\
 \\062\002\000\000\
 \\063\002\000\000\
+\\064\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\
+\\075\002\060\000\020\001\000\000\
+\\076\002\059\000\021\001\000\000\
+\\077\002\009\000\019\001\000\000\
 \\078\002\000\000\
 \\079\002\000\000\
 \\080\002\000\000\
@@ -1753,21 +1843,22 @@
 \\084\002\000\000\
 \\085\002\000\000\
 \\086\002\000\000\
-\\087\002\000\000\
+\\087\002\005\000\139\001\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\
+\\093\002\001\000\249\000\010\000\208\000\011\000\207\000\012\000\206\000\
+\\019\000\203\000\021\000\201\000\022\000\200\000\037\000\248\000\000\000\
 \\094\002\000\000\
 \\095\002\000\000\
 \\096\002\000\000\
-\\097\002\000\000\
-\\098\002\000\000\
+\\097\002\037\000\245\000\000\000\
+\\098\002\001\000\246\000\000\000\
 \\099\002\000\000\
-\\100\002\037\000\211\000\000\000\
-\\101\002\005\000\063\001\000\000\
+\\100\002\000\000\
+\\101\002\000\000\
 \\102\002\000\000\
 \\103\002\000\000\
 \\104\002\000\000\
@@ -1776,10 +1867,10 @@
 \\107\002\000\000\
 \\108\002\000\000\
 \\109\002\000\000\
-\\110\002\005\000\150\001\000\000\
+\\110\002\005\000\175\001\000\000\
 \\111\002\000\000\
 \\112\002\000\000\
-\\113\002\000\000\
+\\113\002\004\000\176\001\000\000\
 \\114\002\000\000\
 \\115\002\000\000\
 \\116\002\000\000\
@@ -1787,79 +1878,80 @@
 \\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\
+\\121\002\000\000\
+\\122\002\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\
+\\127\002\005\000\105\001\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\
+\\133\002\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\
+\\139\002\060\000\212\001\000\000\
+\\140\002\000\000\
+\\141\002\016\000\167\001\000\000\
 \\142\002\000\000\
 \\143\002\000\000\
 \\144\002\000\000\
+\\145\002\005\000\241\001\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\
+\\150\002\000\000\
+\\152\002\000\000\
 \\153\002\000\000\
+\\154\002\000\000\
+\\155\002\001\000\234\000\010\000\208\000\011\000\207\000\012\000\206\000\
+\\019\000\203\000\021\000\201\000\022\000\200\000\037\000\233\000\000\000\
 \\156\002\000\000\
 \\157\002\000\000\
 \\158\002\000\000\
-\\159\002\000\000\
-\\160\002\000\000\
+\\159\002\037\000\230\000\000\000\
+\\160\002\001\000\231\000\000\000\
 \\161\002\000\000\
-\\162\002\004\000\160\001\000\000\
-\\163\002\005\000\159\001\000\000\
+\\162\002\000\000\
+\\163\002\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\
+\\170\002\005\000\163\001\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\
+\\177\002\000\000\
+\\178\002\005\000\075\001\000\000\
 \\179\002\000\000\
 \\180\002\000\000\
-\\181\002\000\000\
+\\181\002\037\000\216\000\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\
+\\183\002\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\
+\\189\002\016\000\024\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\
+\\196\002\000\000\
 \\197\002\000\000\
 \\198\002\000\000\
 \\199\002\000\000\
@@ -1868,216 +1960,254 @@
 \\202\002\000\000\
 \\203\002\000\000\
 \\204\002\000\000\
-\\205\002\009\000\011\001\000\000\
+\\205\002\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\
+\\208\002\000\000\
+\\209\002\000\000\
 \\210\002\000\000\
 \\211\002\000\000\
 \\212\002\000\000\
-\\215\002\000\000\
+\\213\002\000\000\
+\\214\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\
+\\220\002\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\
+\\235\002\000\000\
+\\236\002\000\000\
+\\237\002\000\000\
+\\237\002\066\000\025\001\000\000\
 \\238\002\000\000\
 \\239\002\000\000\
+\\239\002\016\000\222\000\000\000\
 \\240\002\000\000\
 \\241\002\000\000\
 \\242\002\000\000\
+\\243\002\000\000\
 \\244\002\000\000\
 \\245\002\000\000\
 \\246\002\000\000\
 \\247\002\000\000\
-\\248\002\000\000\
+\\248\002\016\000\218\000\000\000\
 \\249\002\000\000\
 \\250\002\000\000\
 \\251\002\000\000\
-\\252\002\000\000\
+\\252\002\016\000\217\000\000\000\
 \\253\002\000\000\
 \\254\002\000\000\
-\\255\002\000\000\
+\\255\002\005\000\155\001\000\000\
 \\000\003\000\000\
 \\001\003\000\000\
 \\002\003\000\000\
-\\003\003\005\000\046\000\000\000\
+\\003\003\000\000\
 \\004\003\000\000\
-\\005\003\005\000\208\000\000\000\
+\\005\003\005\000\145\001\000\000\
 \\006\003\000\000\
 \\007\003\000\000\
 \\008\003\000\000\
-\\009\003\000\000\
+\\009\003\005\000\046\000\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\
+\\011\003\005\000\213\000\000\000\
+\\012\003\004\000\142\001\000\000\
 \\013\003\000\000\
+\\014\003\000\000\
+\\015\003\016\000\143\001\000\000\
+\\016\003\000\000\
+\\017\003\000\000\
+\\018\003\000\000\
+\\019\003\000\000\
+\\020\003\000\000\
+\\021\003\000\000\
+\\022\003\000\000\
+\\023\003\000\000\
+\\024\003\000\000\
+\\025\003\000\000\
+\\026\003\000\000\
+\\027\003\000\000\
+\\028\003\000\000\
+\\029\003\000\000\
+\\030\003\005\000\199\001\000\000\
+\\031\003\000\000\
+\\032\003\000\000\
+\\033\003\000\000\
+\\034\003\000\000\
+\\035\003\000\000\
+\\036\003\000\000\
+\\037\003\000\000\
+\\038\003\000\000\
+\\039\003\000\000\
+\\040\003\000\000\
+\\041\003\000\000\
+\\042\003\000\000\
+\\043\003\000\000\
+\\044\003\000\000\
+\\045\003\000\000\
+\\046\003\000\000\
+\\047\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"
+"\153\000\150\000\153\000\155\000\
+\\154\000\156\000\157\000\158\000\
+\\159\000\073\000\074\000\075\000\
+\\076\000\153\000\077\000\151\000\
+\\061\000\061\000\061\000\061\000\
+\\152\000\144\000\158\001\157\001\
+\\020\000\164\001\163\001\162\001\
+\\161\001\159\001\160\001\168\001\
+\\169\001\165\001\021\000\022\000\
+\\023\000\135\001\173\001\146\000\
+\\146\000\146\000\146\000\105\000\
+\\064\000\024\000\166\000\025\000\
+\\026\000\027\000\091\000\061\000\
+\\053\000\041\000\042\000\007\000\
+\\133\001\096\000\137\001\123\001\
+\\119\001\101\001\165\000\055\001\
+\\056\001\060\001\058\001\089\001\
+\\090\001\092\001\093\001\091\001\
+\\100\001\098\001\002\000\105\001\
+\\103\001\111\001\112\001\003\000\
+\\116\001\004\000\120\001\122\001\
+\\118\001\040\000\109\001\170\001\
+\\113\001\099\001\110\001\078\000\
+\\079\000\080\000\167\001\166\001\
+\\114\001\124\001\172\001\171\001\
+\\060\000\059\000\165\000\026\001\
+\\028\001\030\001\031\001\033\001\
+\\034\001\029\001\039\001\040\001\
+\\027\001\147\000\047\001\068\000\
+\\044\000\041\001\087\001\078\001\
+\\041\000\043\000\077\001\240\000\
+\\165\000\222\000\225\000\227\000\
+\\228\000\230\000\231\000\226\000\
+\\236\000\237\000\223\000\013\000\
+\\239\000\224\000\148\000\249\000\
+\\069\000\046\000\238\000\006\000\
+\\005\000\081\000\082\000\083\000\
+\\042\000\045\000\165\000\167\000\
+\\169\000\172\000\173\000\176\000\
+\\177\000\178\000\011\000\185\000\
+\\186\000\170\000\014\000\171\000\
+\\049\000\174\000\207\000\208\000\
+\\209\000\000\000\189\000\188\000\
+\\168\000\149\000\199\000\070\000\
+\\061\001\063\001\065\001\073\001\
+\\062\001\074\001\072\001\071\001\
+\\102\001\106\001\115\001\104\001\
+\\198\000\084\000\085\000\067\001\
+\\068\001\076\001\075\001\070\001\
+\\069\001\085\001\083\001\082\001\
+\\097\001\084\001\007\000\008\000\
+\\080\001\079\001\081\001\096\001\
+\\066\001\086\001\134\001\061\000\
+\\106\000\052\000\059\000\060\000\
+\\060\000\060\000\060\000\095\001\
+\\060\000\047\000\047\000\046\000\
+\\059\001\039\000\107\000\108\000\
+\\044\000\044\000\044\000\044\000\
+\\044\000\065\000\145\000\046\001\
+\\044\000\109\000\110\000\052\001\
+\\097\000\050\001\111\000\046\000\
+\\046\000\046\000\046\000\046\000\
+\\054\000\066\000\145\000\248\000\
+\\046\000\047\000\047\000\046\000\
+\\112\000\113\000\114\000\004\001\
+\\098\000\001\001\115\000\010\000\
+\\010\000\010\000\010\000\010\000\
+\\010\000\010\000\009\000\010\000\
+\\010\000\010\000\010\000\010\000\
+\\067\000\145\000\009\000\063\000\
+\\012\000\009\000\116\000\117\000\
+\\220\000\099\000\218\000\009\000\
+\\136\001\092\000\142\001\146\001\
+\\144\001\143\001\138\001\141\001\
+\\131\001\140\001\145\001\086\000\
+\\087\000\088\000\089\000\090\000\
+\\051\000\057\001\118\000\125\001\
+\\119\000\094\001\064\001\120\000\
+\\028\000\253\000\029\000\254\000\
+\\030\000\054\001\093\000\036\001\
+\\038\001\032\001\035\001\037\001\
+\\048\001\100\000\044\001\042\001\
+\\049\001\044\000\051\001\094\000\
+\\233\000\235\000\229\000\232\000\
+\\234\000\088\001\008\001\005\001\
+\\050\000\019\000\007\001\017\001\
+\\019\001\016\001\072\000\055\000\
+\\255\000\101\000\243\000\245\000\
+\\246\000\031\000\032\000\033\000\
+\\241\000\006\001\000\001\046\000\
+\\002\001\095\000\180\000\187\000\
+\\009\000\182\000\184\000\175\000\
+\\179\000\183\000\181\000\205\000\
+\\203\000\206\000\212\000\214\000\
+\\210\000\211\000\213\000\215\000\
+\\216\000\102\000\192\000\194\000\
+\\195\000\121\000\204\000\108\001\
+\\034\000\202\000\035\000\001\000\
+\\217\000\009\000\219\000\163\000\
+\\052\000\052\000\164\000\071\000\
+\\042\000\060\000\053\000\041\000\
+\\007\000\156\001\103\000\154\001\
+\\121\001\060\000\117\001\107\001\
+\\060\000\060\000\060\000\162\000\
+\\015\000\145\000\053\001\161\000\
+\\062\000\062\000\145\000\122\000\
+\\014\001\123\000\124\000\055\000\
+\\016\000\145\000\062\000\042\000\
+\\042\000\046\000\003\001\160\000\
+\\125\000\017\000\145\000\009\000\
+\\197\000\007\000\009\000\221\000\
+\\139\001\126\000\130\001\132\001\
+\\127\000\128\000\129\000\130\000\
+\\131\000\052\000\153\001\126\001\
+\\132\000\133\000\036\000\044\000\
+\\045\001\022\001\134\000\020\001\
+\\104\000\010\001\062\000\023\001\
+\\062\000\015\001\135\000\046\000\
+\\244\000\247\000\136\000\137\000\
+\\037\000\190\000\010\000\193\000\
+\\196\000\138\000\038\000\147\001\
+\\149\001\152\001\151\001\150\001\
+\\148\001\155\001\129\001\128\001\
+\\060\000\043\001\018\001\062\000\
+\\018\000\024\001\025\001\048\000\
+\\242\000\252\000\251\000\046\000\
+\\191\000\201\000\009\000\139\000\
+\\021\001\056\000\140\000\141\000\
+\\127\001\009\001\011\001\057\000\
+\\250\000\200\000\013\001\142\000\
+\\058\000\012\001\058\000\143\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\
+\\128\000\008\000\129\000\007\000\130\000\006\000\131\000\005\000\
+\\132\000\004\000\133\000\003\000\134\000\002\000\135\000\001\000\
+\\136\000\011\002\000\000\
+\\000\000\
+\\128\000\008\000\129\000\007\000\130\000\006\000\131\000\005\000\
+\\132\000\004\000\133\000\003\000\134\000\002\000\135\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\
+\\128\000\008\000\129\000\007\000\130\000\006\000\131\000\005\000\
+\\132\000\004\000\133\000\003\000\134\000\002\000\135\000\020\000\000\000\
 \\000\000\
 \\000\000\
 \\002\000\024\000\009\000\023\000\014\000\022\000\000\000\
@@ -2103,10 +2233,10 @@
 \\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\
+\\127\000\045\000\000\000\
+\\127\000\047\000\000\000\
+\\127\000\048\000\000\000\
+\\127\000\049\000\000\000\
 \\000\000\
 \\000\000\
 \\000\000\
@@ -2116,1094 +2246,1191 @@
 \\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\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\066\000\055\000\065\000\057\000\064\000\058\000\063\000\
+\\059\000\062\000\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\061\000\113\000\062\000\112\000\063\000\111\000\065\000\110\000\
+\\066\000\109\000\067\000\108\000\068\000\107\000\069\000\106\000\
+\\070\000\105\000\071\000\104\000\072\000\103\000\073\000\102\000\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\144\000\020\000\085\000\022\000\084\000\023\000\143\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\074\000\138\000\076\000\137\000\077\000\136\000\083\000\135\000\
+\\084\000\134\000\085\000\133\000\089\000\132\000\090\000\131\000\
+\\091\000\130\000\092\000\129\000\093\000\128\000\094\000\127\000\
+\\095\000\126\000\096\000\125\000\097\000\124\000\138\000\123\000\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\098\000\173\000\100\000\172\000\101\000\171\000\
+\\102\000\170\000\103\000\169\000\104\000\168\000\105\000\167\000\
+\\106\000\166\000\107\000\165\000\108\000\164\000\110\000\163\000\
+\\111\000\162\000\112\000\161\000\113\000\160\000\117\000\159\000\
+\\118\000\158\000\119\000\157\000\120\000\156\000\121\000\155\000\
+\\122\000\154\000\123\000\153\000\124\000\152\000\125\000\151\000\
+\\126\000\150\000\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\001\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\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\036\000\219\000\037\000\218\000\038\000\217\000\000\000\
+\\000\000\
+\\000\000\
+\\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\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\225\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\224\000\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\066\000\055\000\065\000\057\000\064\000\058\000\226\000\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\001\000\227\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\050\000\230\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\063\000\111\000\065\000\110\000\066\000\235\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\061\000\113\000\062\000\238\000\063\000\111\000\065\000\110\000\
+\\066\000\109\000\067\000\108\000\068\000\107\000\069\000\106\000\
+\\070\000\105\000\071\000\104\000\072\000\237\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\060\000\240\000\063\000\111\000\065\000\110\000\066\000\109\000\
+\\067\000\108\000\068\000\107\000\069\000\106\000\070\000\105\000\
+\\071\000\104\000\072\000\239\000\144\000\061\000\145\000\060\000\
+\\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\001\000\242\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\050\000\245\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\077\000\136\000\085\000\133\000\089\000\132\000\090\000\251\000\
+\\138\000\123\000\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\144\000\020\000\085\000\022\000\084\000\023\000\143\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\074\000\138\000\076\000\002\001\077\000\136\000\083\000\135\000\
+\\084\000\001\001\085\000\133\000\089\000\132\000\090\000\131\000\
+\\091\000\130\000\092\000\129\000\093\000\128\000\094\000\127\000\
+\\095\000\126\000\096\000\000\001\138\000\123\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\075\000\004\001\077\000\136\000\085\000\133\000\089\000\132\000\
+\\090\000\131\000\091\000\130\000\092\000\129\000\093\000\128\000\
+\\094\000\127\000\095\000\126\000\096\000\003\001\138\000\123\000\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\001\000\006\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\036\000\183\000\037\000\182\000\050\000\179\000\053\000\010\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\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\098\000\173\000\100\000\027\001\101\000\171\000\
+\\102\000\170\000\103\000\169\000\104\000\168\000\105\000\167\000\
+\\106\000\166\000\107\000\165\000\108\000\164\000\110\000\163\000\
+\\111\000\162\000\112\000\161\000\113\000\160\000\117\000\159\000\
+\\118\000\158\000\119\000\157\000\120\000\156\000\121\000\155\000\
+\\122\000\154\000\123\000\153\000\124\000\152\000\125\000\026\001\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\099\000\029\001\101\000\171\000\102\000\170\000\
+\\103\000\169\000\104\000\168\000\105\000\167\000\106\000\166\000\
+\\107\000\165\000\108\000\164\000\110\000\163\000\111\000\162\000\
+\\112\000\161\000\113\000\160\000\117\000\159\000\118\000\158\000\
+\\119\000\157\000\120\000\156\000\121\000\155\000\122\000\154\000\
+\\123\000\153\000\124\000\152\000\125\000\028\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\002\000\058\000\003\000\032\001\009\000\023\000\014\000\022\000\000\000\
+\\000\000\
+\\006\000\041\001\008\000\040\001\009\000\039\001\010\000\038\001\
+\\011\000\037\001\012\000\036\001\013\000\035\001\014\000\087\000\
+\\016\000\034\001\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\066\000\055\000\065\000\057\000\049\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\051\001\021\000\050\001\022\000\084\000\
+\\023\000\083\000\024\000\082\000\025\000\187\000\026\000\080\000\
+\\027\000\186\000\028\000\078\000\029\000\077\000\030\000\076\000\
+\\031\000\075\000\032\000\184\000\033\000\073\000\034\000\072\000\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\051\001\021\000\052\001\022\000\084\000\
+\\023\000\083\000\024\000\082\000\025\000\187\000\026\000\080\000\
+\\027\000\186\000\028\000\078\000\029\000\077\000\030\000\076\000\
+\\031\000\075\000\032\000\184\000\033\000\073\000\034\000\072\000\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\053\001\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\054\001\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\051\001\021\000\055\001\022\000\084\000\
+\\023\000\083\000\024\000\082\000\025\000\187\000\026\000\080\000\
+\\027\000\186\000\028\000\078\000\029\000\077\000\030\000\076\000\
+\\031\000\075\000\032\000\184\000\033\000\073\000\034\000\072\000\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\051\000\140\000\089\000\057\001\139\000\056\001\000\000\
+\\051\000\140\000\089\000\059\001\140\000\058\001\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\077\000\136\000\085\000\133\000\089\000\132\000\090\000\131\000\
+\\091\000\130\000\092\000\129\000\093\000\128\000\094\000\127\000\
+\\095\000\126\000\096\000\060\001\138\000\123\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\036\000\219\000\038\000\217\000\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\063\000\111\000\065\000\110\000\066\000\063\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\063\000\111\000\065\000\110\000\066\000\064\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\063\000\111\000\065\000\110\000\066\000\065\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\063\000\111\000\065\000\110\000\066\000\066\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\063\000\111\000\065\000\110\000\066\000\067\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\061\000\068\001\000\000\
+\\011\000\070\001\064\000\069\001\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\063\000\111\000\065\000\110\000\066\000\109\000\067\000\108\000\
+\\068\000\107\000\069\000\106\000\070\000\105\000\071\000\104\000\
+\\072\000\237\000\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\077\000\136\000\085\000\133\000\089\000\132\000\090\000\076\001\
+\\138\000\123\000\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\077\000\136\000\085\000\133\000\089\000\132\000\090\000\077\001\
+\\138\000\123\000\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\077\000\136\000\085\000\133\000\089\000\132\000\090\000\078\001\
+\\138\000\123\000\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\077\000\136\000\085\000\133\000\089\000\132\000\090\000\079\001\
+\\138\000\123\000\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\077\000\136\000\085\000\133\000\089\000\132\000\090\000\080\001\
+\\138\000\123\000\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\089\001\011\000\088\001\047\000\087\001\079\000\086\001\
+\\080\000\085\001\081\000\084\001\082\000\083\001\141\000\082\001\
+\\145\000\081\001\000\000\
+\\074\000\092\001\000\000\
+\\011\000\096\001\086\000\095\001\087\000\094\001\088\000\093\001\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\077\000\136\000\085\000\133\000\089\000\132\000\090\000\131\000\
+\\091\000\130\000\092\000\129\000\093\000\128\000\094\000\127\000\
+\\095\000\126\000\096\000\000\001\138\000\123\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\051\000\140\000\089\000\059\001\140\000\097\001\000\000\
+\\051\000\140\000\089\000\057\001\139\000\098\001\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\077\000\136\000\085\000\133\000\089\000\132\000\090\000\131\000\
+\\091\000\130\000\092\000\129\000\093\000\128\000\094\000\127\000\
+\\095\000\126\000\096\000\099\001\138\000\123\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\113\000\160\000\117\000\159\000\118\000\106\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\113\000\160\000\117\000\159\000\118\000\109\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\113\000\160\000\117\000\159\000\118\000\110\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\113\000\160\000\117\000\159\000\118\000\111\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\113\000\160\000\117\000\159\000\118\000\112\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\113\000\160\000\117\000\159\000\118\000\113\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\113\000\160\000\117\000\159\000\118\000\114\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\169\000\
+\\104\000\168\000\105\000\167\000\106\000\166\000\107\000\165\000\
+\\108\000\164\000\109\000\116\001\110\000\163\000\111\000\162\000\
+\\112\000\161\000\113\000\160\000\117\000\159\000\118\000\158\000\
+\\119\000\157\000\120\000\156\000\121\000\155\000\122\000\154\000\
+\\123\000\153\000\124\000\152\000\125\000\115\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\108\000\118\001\113\000\160\000\117\000\159\000\118\000\117\001\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\108\000\119\001\113\000\160\000\117\000\159\000\118\000\117\001\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\106\000\121\001\108\000\120\001\113\000\160\000\117\000\159\000\
+\\118\000\117\001\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\108\000\122\001\113\000\160\000\117\000\159\000\118\000\117\001\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\108\000\123\001\113\000\160\000\117\000\159\000\118\000\117\001\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\098\000\124\001\000\000\
+\\011\000\128\001\114\000\127\001\115\000\126\001\116\000\125\001\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\169\000\
+\\104\000\168\000\105\000\167\000\106\000\166\000\107\000\165\000\
+\\108\000\164\000\110\000\163\000\111\000\162\000\112\000\161\000\
+\\113\000\160\000\117\000\159\000\118\000\158\000\119\000\157\000\
+\\120\000\156\000\121\000\155\000\122\000\154\000\123\000\153\000\
+\\124\000\152\000\125\000\129\001\144\000\061\000\145\000\060\000\
+\\146\000\059\000\000\000\
+\\009\000\090\000\019\000\131\001\031\000\130\001\000\000\
+\\051\000\178\000\054\000\175\000\117\000\133\001\137\000\132\001\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\169\000\
+\\104\000\168\000\105\000\167\000\106\000\166\000\107\000\165\000\
+\\108\000\164\000\110\000\163\000\111\000\162\000\112\000\161\000\
+\\113\000\160\000\117\000\159\000\118\000\158\000\119\000\157\000\
+\\120\000\156\000\121\000\155\000\122\000\154\000\123\000\153\000\
+\\124\000\152\000\125\000\134\001\144\000\061\000\145\000\060\000\
+\\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\169\000\
+\\104\000\168\000\105\000\167\000\106\000\166\000\107\000\165\000\
+\\108\000\164\000\110\000\163\000\111\000\162\000\112\000\161\000\
+\\113\000\160\000\117\000\159\000\118\000\158\000\119\000\157\000\
+\\120\000\156\000\121\000\155\000\122\000\154\000\123\000\153\000\
+\\124\000\152\000\125\000\026\001\144\000\061\000\145\000\060\000\
+\\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\005\000\142\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\006\000\041\001\007\000\150\001\008\000\149\001\009\000\039\001\
+\\010\000\038\001\011\000\037\001\012\000\036\001\013\000\035\001\
+\\014\000\087\000\016\000\034\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\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\060\000\162\001\063\000\111\000\065\000\110\000\066\000\109\000\
+\\067\000\108\000\068\000\107\000\069\000\106\000\070\000\105\000\
+\\071\000\104\000\072\000\239\000\144\000\061\000\145\000\060\000\
+\\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\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\089\001\011\000\088\001\047\000\087\001\078\000\170\001\
+\\079\000\169\001\080\000\168\001\081\000\084\001\141\000\167\001\
+\\145\000\081\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\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\075\000\178\001\077\000\136\000\085\000\133\000\089\000\132\000\
+\\090\000\131\000\091\000\130\000\092\000\129\000\093\000\128\000\
+\\094\000\127\000\095\000\126\000\096\000\003\001\138\000\123\000\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\169\000\
+\\104\000\168\000\105\000\167\000\106\000\166\000\107\000\165\000\
+\\108\000\164\000\110\000\163\000\111\000\162\000\112\000\161\000\
+\\113\000\160\000\117\000\159\000\118\000\158\000\119\000\157\000\
+\\120\000\156\000\121\000\155\000\122\000\154\000\123\000\153\000\
+\\124\000\152\000\125\000\180\001\144\000\061\000\145\000\060\000\
+\\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\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\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\099\000\187\001\101\000\171\000\102\000\170\000\
+\\103\000\169\000\104\000\168\000\105\000\167\000\106\000\166\000\
+\\107\000\165\000\108\000\164\000\110\000\163\000\111\000\162\000\
+\\112\000\161\000\113\000\160\000\117\000\159\000\118\000\158\000\
+\\119\000\157\000\120\000\156\000\121\000\155\000\122\000\154\000\
+\\123\000\153\000\124\000\152\000\125\000\028\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\006\000\041\001\008\000\188\001\009\000\039\001\010\000\038\001\
+\\011\000\037\001\012\000\036\001\013\000\035\001\014\000\087\000\
+\\016\000\034\001\000\000\
+\\006\000\041\001\007\000\189\001\008\000\149\001\009\000\039\001\
+\\010\000\038\001\011\000\037\001\012\000\036\001\013\000\035\001\
+\\014\000\087\000\016\000\034\001\000\000\
+\\000\000\
+\\006\000\191\001\017\000\190\001\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\144\000\020\000\085\000\022\000\084\000\023\000\143\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\074\000\138\000\076\000\137\000\077\000\136\000\083\000\135\000\
+\\084\000\134\000\085\000\133\000\089\000\132\000\090\000\131\000\
+\\091\000\130\000\092\000\129\000\093\000\128\000\094\000\127\000\
+\\095\000\126\000\096\000\125\000\097\000\192\001\138\000\123\000\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\193\001\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\066\000\055\000\065\000\057\000\064\000\058\000\063\000\
+\\059\000\194\001\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\061\000\113\000\062\000\112\000\063\000\111\000\065\000\110\000\
+\\066\000\109\000\067\000\108\000\068\000\107\000\069\000\106\000\
+\\070\000\105\000\071\000\104\000\072\000\103\000\073\000\195\001\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\098\000\173\000\100\000\172\000\101\000\171\000\
+\\102\000\170\000\103\000\169\000\104\000\168\000\105\000\167\000\
+\\106\000\166\000\107\000\165\000\108\000\164\000\110\000\163\000\
+\\111\000\162\000\112\000\161\000\113\000\160\000\117\000\159\000\
+\\118\000\158\000\119\000\157\000\120\000\156\000\121\000\155\000\
+\\122\000\154\000\123\000\153\000\124\000\152\000\125\000\151\000\
+\\126\000\196\001\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\051\001\021\000\199\001\022\000\084\000\
+\\023\000\083\000\024\000\082\000\025\000\187\000\026\000\080\000\
+\\027\000\186\000\028\000\078\000\029\000\077\000\030\000\076\000\
+\\031\000\075\000\032\000\184\000\033\000\073\000\034\000\072\000\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\200\001\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\201\001\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\202\001\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\011\000\070\001\064\000\204\001\000\000\
+\\000\000\
+\\000\000\
+\\009\000\089\001\011\000\088\001\047\000\087\001\080\000\205\001\
+\\145\000\081\001\000\000\
+\\009\000\089\001\011\000\088\001\047\000\087\001\080\000\207\001\
+\\143\000\206\001\145\000\081\001\000\000\
+\\011\000\096\001\086\000\095\001\087\000\094\001\088\000\208\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\089\001\011\000\088\001\047\000\087\001\078\000\214\001\
+\\079\000\169\001\080\000\168\001\081\000\084\001\141\000\167\001\
+\\145\000\081\001\000\000\
+\\000\000\
+\\011\000\096\001\086\000\095\001\087\000\094\001\088\000\216\001\000\000\
+\\009\000\089\001\011\000\088\001\047\000\087\001\080\000\217\001\
+\\145\000\081\001\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\144\000\020\000\085\000\022\000\084\000\023\000\143\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\074\000\138\000\076\000\137\000\077\000\136\000\083\000\135\000\
+\\084\000\134\000\085\000\133\000\089\000\132\000\090\000\131\000\
+\\091\000\130\000\092\000\129\000\093\000\128\000\094\000\127\000\
+\\095\000\126\000\096\000\125\000\097\000\218\001\138\000\123\000\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\144\000\020\000\085\000\022\000\084\000\023\000\143\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\074\000\138\000\076\000\137\000\077\000\136\000\083\000\135\000\
+\\084\000\134\000\085\000\133\000\089\000\132\000\090\000\131\000\
+\\091\000\130\000\092\000\129\000\093\000\128\000\094\000\127\000\
+\\095\000\126\000\096\000\125\000\097\000\219\001\138\000\123\000\
+\\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\077\000\136\000\085\000\133\000\089\000\132\000\090\000\131\000\
+\\091\000\130\000\092\000\129\000\093\000\128\000\094\000\127\000\
+\\095\000\126\000\096\000\220\001\138\000\123\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\011\000\128\001\114\000\127\001\115\000\126\001\116\000\223\001\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\169\000\
+\\104\000\168\000\105\000\167\000\106\000\166\000\107\000\165\000\
+\\108\000\164\000\109\000\224\001\110\000\163\000\111\000\162\000\
+\\112\000\161\000\113\000\160\000\117\000\159\000\118\000\158\000\
+\\119\000\157\000\120\000\156\000\121\000\155\000\122\000\154\000\
+\\123\000\153\000\124\000\152\000\125\000\115\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\098\000\173\000\100\000\172\000\101\000\171\000\
+\\102\000\170\000\103\000\169\000\104\000\168\000\105\000\167\000\
+\\106\000\166\000\107\000\165\000\108\000\164\000\110\000\163\000\
+\\111\000\162\000\112\000\161\000\113\000\160\000\117\000\159\000\
+\\118\000\158\000\119\000\157\000\120\000\156\000\121\000\155\000\
+\\122\000\154\000\123\000\153\000\124\000\152\000\125\000\151\000\
+\\126\000\225\001\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\169\000\
+\\104\000\168\000\105\000\167\000\106\000\166\000\107\000\165\000\
+\\108\000\164\000\110\000\163\000\111\000\162\000\112\000\161\000\
+\\113\000\160\000\117\000\159\000\118\000\158\000\119\000\157\000\
+\\120\000\156\000\121\000\155\000\122\000\154\000\123\000\153\000\
+\\124\000\152\000\125\000\226\001\144\000\061\000\145\000\060\000\
+\\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\006\000\041\001\007\000\233\001\008\000\149\001\009\000\039\001\
+\\010\000\038\001\011\000\037\001\012\000\036\001\013\000\035\001\
+\\014\000\087\000\016\000\034\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\117\000\046\000\116\000\051\000\115\000\055\000\114\000\
+\\063\000\111\000\065\000\110\000\066\000\237\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\089\001\011\000\088\001\047\000\087\001\080\000\241\001\
+\\145\000\081\001\000\000\
+\\000\000\
+\\009\000\089\001\011\000\088\001\047\000\087\001\080\000\242\001\
+\\145\000\081\001\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\077\000\136\000\085\000\133\000\089\000\132\000\090\000\244\001\
+\\138\000\123\000\144\000\061\000\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\107\001\
+\\113\000\160\000\117\000\159\000\118\000\248\001\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\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\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\251\001\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\009\000\089\001\011\000\088\001\047\000\087\001\080\000\207\001\
+\\143\000\252\001\145\000\081\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\085\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\081\000\026\000\080\000\027\000\079\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\075\000\
+\\032\000\074\000\033\000\073\000\034\000\072\000\035\000\071\000\
+\\039\000\070\000\042\000\069\000\043\000\068\000\044\000\067\000\
+\\045\000\142\000\046\000\141\000\051\000\140\000\055\000\139\000\
+\\077\000\136\000\085\000\133\000\089\000\132\000\090\000\131\000\
+\\091\000\130\000\092\000\129\000\093\000\128\000\094\000\127\000\
+\\095\000\126\000\096\000\254\001\138\000\123\000\144\000\061\000\
+\\145\000\060\000\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\009\000\090\000\011\000\089\000\012\000\088\000\014\000\087\000\
+\\019\000\086\000\020\000\188\000\022\000\084\000\023\000\083\000\
+\\024\000\082\000\025\000\187\000\026\000\080\000\027\000\186\000\
+\\028\000\078\000\029\000\077\000\030\000\076\000\031\000\185\000\
+\\032\000\184\000\033\000\073\000\034\000\072\000\036\000\183\000\
+\\037\000\182\000\046\000\181\000\049\000\180\000\050\000\179\000\
+\\051\000\178\000\052\000\177\000\053\000\176\000\054\000\175\000\
+\\056\000\174\000\101\000\171\000\102\000\170\000\103\000\169\000\
+\\104\000\168\000\105\000\167\000\106\000\166\000\107\000\165\000\
+\\108\000\164\000\110\000\163\000\111\000\162\000\112\000\161\000\
+\\113\000\160\000\117\000\159\000\118\000\158\000\119\000\157\000\
+\\120\000\156\000\121\000\155\000\122\000\154\000\123\000\153\000\
+\\124\000\152\000\125\000\255\001\144\000\061\000\145\000\060\000\
+\\146\000\059\000\000\000\
+\\000\000\
+\\000\000\
+\\009\000\089\001\011\000\088\001\047\000\087\001\080\000\002\002\
+\\142\000\001\002\145\000\081\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\089\001\011\000\088\001\047\000\087\001\079\000\007\002\
+\\080\000\006\002\081\000\084\001\145\000\081\001\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\000\000\
+\\009\000\089\001\011\000\088\001\047\000\087\001\078\000\170\001\
+\\079\000\169\001\080\000\168\001\081\000\084\001\145\000\081\001\000\000\
+\\000\000\
+\\009\000\089\001\011\000\088\001\047\000\087\001\078\000\214\001\
+\\079\000\169\001\080\000\168\001\081\000\084\001\145\000\081\001\000\000\
 \\000\000\
 \"
-val numstates = 498
-val numrules = 282
+val numstates = 524
+val numrules = 290
 val s = Unsynchronized.ref "" and index = Unsynchronized.ref 0
 val string_to_int = fn () => 
 let val i = !index
@@ -3266,12 +3493,18 @@
 structure MlyValue = 
 struct
 datatype svalue = VOID | ntVOID of unit
- | ATOMIC_SYSTEM_WORD of  (string) | ATOMIC_DEFINED_WORD of  (string)
+ | DOLLAR_DOLLAR_WORD of  (string) | DOLLAR_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)
+ | RATIONAL of  (string) | REAL of  (string)
+ | atomic_system_word of  (string) | atomic_defined_word of  (string)
+ | let_term of  (tptp_term) | tff_type_arguments of  (tptp_type list)
+ | tff_monotype of  (tptp_type) | tff_quantified_type of  (tptp_type)
+ | tff_let_formula_defn of  (tptp_let list)
+ | tff_let_term_defn of  (tptp_let list) | tff_let of  (tptp_formula)
+ | thf_let_defn of  (tptp_let list) | 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)
@@ -3296,8 +3529,7 @@
  | 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_let of  (tptp_formula) | 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)
@@ -3318,9 +3550,7 @@
  | 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_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)
@@ -3397,7 +3627,7 @@
   | (T 6) => "EXCLAMATION"
   | (T 7) => "LET"
   | (T 8) => "ARROW"
-  | (T 9) => "IF"
+  | (T 9) => "FI"
   | (T 10) => "IFF"
   | (T 11) => "IMPLIES"
   | (T 12) => "INCLUDE"
@@ -3451,8 +3681,8 @@
   | (T 60) => "GENTZEN_ARROW"
   | (T 61) => "DEP_SUM"
   | (T 62) => "DEP_PROD"
-  | (T 63) => "ATOMIC_DEFINED_WORD"
-  | (T 64) => "ATOMIC_SYSTEM_WORD"
+  | (T 63) => "DOLLAR_WORD"
+  | (T 64) => "DOLLAR_DOLLAR_WORD"
   | (T 65) => "SUBTYPE"
   | (T 66) => "LET_TERM"
   | (T 67) => "THF"
@@ -3461,21 +3691,26 @@
   | (T 70) => "CNF"
   | (T 71) => "ITE_F"
   | (T 72) => "ITE_T"
+  | (T 73) => "LET_TF"
+  | (T 74) => "LET_FF"
+  | (T 75) => "LET_FT"
+  | (T 76) => "LET_TT"
   | _ => "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
+ $$ (T 76) $$ (T 75) $$ (T 74) $$ (T 73) $$ (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
@@ -3484,292 +3719,1675 @@
 fn (i392,defaultPos,stack,
     (file_name):arg) =>
 case (i392,stack)
-of  ( 0, ( ( _, ( MlyValue.optional_info optional_info, _, 
+of  ( 0, ( ( _, ( MlyValue.tptp_file tptp_file, tptp_file1left, 
+tptp_file1right)) :: rest671)) => let val  result = MlyValue.tptp (
+( tptp_file ))
+ in ( LrTable.NT 135, ( result, tptp_file1left, tptp_file1right), 
+rest671)
+end
+|  ( 1, ( ( _, ( 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 134, ( result, tptp_input1left, tptp_file1right), 
+rest671)
+end
+|  ( 2, ( ( _, ( MlyValue.tptp_file tptp_file, _, tptp_file1right)) ::
+ ( _, ( _, COMMENT1left, _)) :: rest671)) => let val  result = 
+MlyValue.tptp_file (( tptp_file ))
+ in ( LrTable.NT 134, ( result, COMMENT1left, tptp_file1right), 
+rest671)
+end
+|  ( 3, ( rest671)) => let val  result = MlyValue.tptp_file (( [] ))
+ in ( LrTable.NT 134, ( result, defaultPos, defaultPos), rest671)
+end
+|  ( 4, ( ( _, ( MlyValue.annotated_formula annotated_formula, 
+annotated_formula1left, annotated_formula1right)) :: rest671)) => let
+ val  result = MlyValue.tptp_input (( annotated_formula ))
+ in ( LrTable.NT 133, ( result, annotated_formula1left, 
+annotated_formula1right), rest671)
+end
+|  ( 5, ( ( _, ( MlyValue.include_ include_, include_1left, 
+include_1right)) :: rest671)) => let val  result = MlyValue.tptp_input
+ (( include_ ))
+ in ( LrTable.NT 133, ( result, include_1left, include_1right), 
+rest671)
+end
+|  ( 6, ( ( _, ( MlyValue.thf_annotated thf_annotated, 
+thf_annotated1left, thf_annotated1right)) :: rest671)) => let val  
+result = MlyValue.annotated_formula (( thf_annotated ))
+ in ( LrTable.NT 131, ( result, thf_annotated1left, 
+thf_annotated1right), rest671)
+end
+|  ( 7, ( ( _, ( MlyValue.tff_annotated tff_annotated, 
+tff_annotated1left, tff_annotated1right)) :: rest671)) => let val  
+result = MlyValue.annotated_formula (( tff_annotated ))
+ in ( LrTable.NT 131, ( result, tff_annotated1left, 
+tff_annotated1right), rest671)
+end
+|  ( 8, ( ( _, ( MlyValue.fof_annotated fof_annotated, 
+fof_annotated1left, fof_annotated1right)) :: rest671)) => let val  
+result = MlyValue.annotated_formula (( fof_annotated ))
+ in ( LrTable.NT 131, ( result, fof_annotated1left, 
+fof_annotated1right), rest671)
+end
+|  ( 9, ( ( _, ( MlyValue.cnf_annotated cnf_annotated, 
+cnf_annotated1left, cnf_annotated1right)) :: rest671)) => let val  
+result = MlyValue.annotated_formula (( cnf_annotated ))
+ in ( LrTable.NT 131, ( result, cnf_annotated1left, 
+cnf_annotated1right), rest671)
+end
+|  ( 10, ( ( _, ( _, _, 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 130, ( result, THF1left, PERIOD1right), rest671)
+end
+|  ( 11, ( ( _, ( _, _, 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 129, ( result, TFF1left, PERIOD1right), rest671)
+end
+|  ( 12, ( ( _, ( _, _, 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 128, ( result, FOF1left, PERIOD1right), rest671)
+end
+|  ( 13, ( ( _, ( _, _, 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 127, ( result, CNF1left, PERIOD1right), rest671)
+end
+|  ( 14, ( ( _, ( 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 (
+|  ( 15, ( 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)
+|  ( 16, ( ( _, ( MlyValue.LOWER_WORD LOWER_WORD, LOWER_WORD1left, 
+LOWER_WORD1right)) :: rest671)) => let val  result = 
+MlyValue.formula_role (( classify_role LOWER_WORD ))
+ in ( LrTable.NT 126, ( result, LOWER_WORD1left, LOWER_WORD1right), 
+rest671)
+end
+|  ( 17, ( ( _, ( 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 125, ( result, thf_logic_formula1left, 
+thf_logic_formula1right), rest671)
+end
+|  ( 18, ( ( _, ( MlyValue.thf_sequent thf_sequent, thf_sequent1left, 
+thf_sequent1right)) :: rest671)) => let val  result = 
+MlyValue.thf_formula (( thf_sequent ))
+ in ( LrTable.NT 125, ( result, thf_sequent1left, thf_sequent1right), 
+rest671)
+end
+|  ( 19, ( ( _, ( 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 124, ( result, thf_binary_formula1left, 
+thf_binary_formula1right), rest671)
+end
+|  ( 20, ( ( _, ( 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 124, ( result, thf_unitary_formula1left, 
+thf_unitary_formula1right), rest671)
+end
+|  ( 21, ( ( _, ( 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 124, ( result, thf_type_formula1left, 
+thf_type_formula1right), rest671)
+end
+|  ( 22, ( ( _, ( MlyValue.thf_subtype thf_subtype, thf_subtype1left, 
+thf_subtype1right)) :: rest671)) => let val  result = 
+MlyValue.thf_logic_formula (( Type_fmla thf_subtype ))
+ in ( LrTable.NT 124, ( result, thf_subtype1left, thf_subtype1right), 
+rest671)
+end
+|  ( 23, ( ( _, ( 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 123, ( result, thf_binary_pair1left, 
+thf_binary_pair1right), rest671)
+end
+|  ( 24, ( ( _, ( 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 123, ( result, thf_binary_tuple1left, 
+thf_binary_tuple1right), rest671)
+end
+|  ( 25, ( ( _, ( MlyValue.thf_binary_type thf_binary_type, 
+thf_binary_type1left, thf_binary_type1right)) :: rest671)) => let val 
+ result = MlyValue.thf_binary_formula (( Type_fmla thf_binary_type ))
+ in ( LrTable.NT 123, ( result, thf_binary_type1left, 
+thf_binary_type1right), rest671)
+end
+|  ( 26, ( ( _, ( 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 122, ( result, thf_unitary_formula1left, 
+thf_unitary_formula2right), rest671)
+end
+|  ( 27, ( ( _, ( 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 121, ( result, thf_or_formula1left, 
+thf_or_formula1right), rest671)
+end
+|  ( 28, ( ( _, ( 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 121, ( result, thf_and_formula1left, 
+thf_and_formula1right), rest671)
+end
+|  ( 29, ( ( _, ( 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 121, ( result, thf_apply_formula1left, 
+thf_apply_formula1right), rest671)
+end
+|  ( 30, ( ( _, ( 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 120, ( result, thf_unitary_formula1left, 
+thf_unitary_formula2right), rest671)
+end
+|  ( 31, ( ( _, ( 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 120, ( result, thf_or_formula1left, 
+thf_unitary_formula1right), rest671)
+end
+|  ( 32, ( ( _, ( 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 119, ( result, thf_unitary_formula1left, 
+thf_unitary_formula2right), rest671)
+end
+|  ( 33, ( ( _, ( 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 119, ( result, thf_and_formula1left, 
+thf_unitary_formula1right), rest671)
+end
+|  ( 34, ( ( _, ( 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 118, ( result, thf_unitary_formula1left, 
+thf_unitary_formula2right), rest671)
+end
+|  ( 35, ( ( _, ( 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 118, ( result, thf_apply_formula1left, 
+thf_unitary_formula1right), rest671)
+end
+|  ( 36, ( ( _, ( 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 117, ( result, thf_quantified_formula1left, 
+thf_quantified_formula1right), rest671)
+end
+|  ( 37, ( ( _, ( 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 117, ( result, thf_unary_formula1left, 
+thf_unary_formula1right), rest671)
+end
+|  ( 38, ( ( _, ( MlyValue.thf_atom thf_atom, thf_atom1left, 
+thf_atom1right)) :: rest671)) => let val  result = 
+MlyValue.thf_unitary_formula (( thf_atom ))
+ in ( LrTable.NT 117, ( result, thf_atom1left, thf_atom1right), 
+rest671)
+end
+|  ( 39, ( ( _, ( MlyValue.thf_conditional thf_conditional, 
+thf_conditional1left, thf_conditional1right)) :: rest671)) => let val 
+ result = MlyValue.thf_unitary_formula (( thf_conditional ))
+ in ( LrTable.NT 117, ( result, thf_conditional1left, 
+thf_conditional1right), rest671)
+end
+|  ( 40, ( ( _, ( MlyValue.thf_let thf_let, thf_let1left, 
+thf_let1right)) :: rest671)) => let val  result = 
+MlyValue.thf_unitary_formula (( thf_let ))
+ in ( LrTable.NT 117, ( result, thf_let1left, thf_let1right), 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), 
+|  ( 41, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.thf_logic_formula thf_logic_formula, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.thf_unitary_formula (( thf_logic_formula ))
+ in ( LrTable.NT 117, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 42, ( ( _, ( 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 116, ( result, thf_quantifier1left, 
+thf_unitary_formula1right), rest671)
+end
+|  ( 43, ( ( _, ( MlyValue.thf_variable thf_variable, 
+thf_variable1left, thf_variable1right)) :: rest671)) => let val  
+result = MlyValue.thf_variable_list (( [thf_variable] ))
+ in ( LrTable.NT 115, ( result, thf_variable1left, thf_variable1right)
+, rest671)
+end
+|  ( 44, ( ( _, ( 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 115, ( result, thf_variable1left, 
+thf_variable_list1right), rest671)
+end
+|  ( 45, ( ( _, ( 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 114, ( result, thf_typed_variable1left, 
+thf_typed_variable1right), rest671)
+end
+|  ( 46, ( ( _, ( MlyValue.variable_ variable_, variable_1left, 
+variable_1right)) :: rest671)) => let val  result = 
+MlyValue.thf_variable (( (variable_, NONE) ))
+ in ( LrTable.NT 114, ( result, variable_1left, variable_1right), 
 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), 
+|  ( 47, ( ( _, ( 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 113, ( result, variable_1left, 
+thf_top_level_type1right), rest671)
+end
+|  ( 48, ( ( _, ( _, _, 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 112, ( result, thf_unary_connective1left, 
+RPAREN1right), rest671)
+end
+|  ( 49, ( ( _, ( MlyValue.term term, term1left, term1right)) :: 
+rest671)) => let val  result = MlyValue.thf_atom (
+( Atom (THF_Atom_term term) ))
+ in ( LrTable.NT 102, ( result, term1left, term1right), rest671)
+end
+|  ( 50, ( ( _, ( 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 102, ( result, thf_conn_term1left, 
+thf_conn_term1right), rest671)
+end
+|  ( 51, ( ( _, ( _, _, 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 100, ( result, ITE_F1left, RPAREN1right), rest671)
+
+end
+|  ( 52, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.thf_formula 
+thf_formula, _, _)) :: _ :: ( _, ( MlyValue.thf_let_defn thf_let_defn,
+ _, _)) :: _ :: ( _, ( _, LET_TF1left, _)) :: rest671)) => let val  
+result = MlyValue.thf_let ((
+  Let (thf_let_defn, thf_formula)
+))
+ in ( LrTable.NT 101, ( result, LET_TF1left, RPAREN1right), rest671)
+
+end
+|  ( 53, ( ( _, ( MlyValue.thf_quantified_formula 
+thf_quantified_formula, thf_quantified_formula1left, 
+thf_quantified_formula1right)) :: rest671)) => let val  result = 
+MlyValue.thf_let_defn (
+(
+  let
+    val (_, vars, fmla) = extract_quant_info thf_quantified_formula
+  in [Let_fmla (hd vars, fmla)]
+  end
+)
+)
+ in ( LrTable.NT 136, ( result, thf_quantified_formula1left, 
+thf_quantified_formula1right), rest671)
+end
+|  ( 54, ( ( _, ( 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 111, ( result, thf_typeable_formula1left, 
+thf_top_level_type1right), rest671)
+end
+|  ( 55, ( ( _, ( MlyValue.thf_atom thf_atom, thf_atom1left, 
+thf_atom1right)) :: rest671)) => let val  result = 
+MlyValue.thf_typeable_formula (( thf_atom ))
+ in ( LrTable.NT 110, ( result, thf_atom1left, thf_atom1right), 
 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), 
+|  ( 56, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.thf_logic_formula thf_logic_formula, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.thf_typeable_formula (( thf_logic_formula ))
+ in ( LrTable.NT 110, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 57, ( ( _, ( MlyValue.constant constant2, _, constant2right)) ::
+ _ :: ( _, ( MlyValue.constant constant1, constant1left, _)) :: 
+rest671)) => let val  result = MlyValue.thf_subtype (
+( Subtype(constant1, constant2) ))
+ in ( LrTable.NT 109, ( result, constant1left, constant2right), 
 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), 
+|  ( 58, ( ( _, ( 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 108, ( result, thf_logic_formula1left, 
+thf_logic_formula1right), rest671)
+end
+|  ( 59, ( ( _, ( 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 107, ( result, thf_unitary_formula1left, 
+thf_unitary_formula1right), rest671)
+end
+|  ( 60, ( ( _, ( 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 106, ( result, thf_mapping_type1left, 
+thf_mapping_type1right), rest671)
+end
+|  ( 61, ( ( _, ( 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 106, ( result, thf_xprod_type1left, 
+thf_xprod_type1right), rest671)
+end
+|  ( 62, ( ( _, ( 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 106, ( result, thf_union_type1left, 
+thf_union_type1right), rest671)
+end
+|  ( 63, ( ( _, ( 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 105, ( result, thf_unitary_type1left, 
+thf_unitary_type2right), rest671)
+end
+|  ( 64, ( ( _, ( 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 105, ( result, thf_unitary_type1left, 
+thf_mapping_type1right), rest671)
+end
+|  ( 65, ( ( _, ( 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 104, ( result, thf_unitary_type1left, 
+thf_unitary_type2right), rest671)
+end
+|  ( 66, ( ( _, ( 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 104, ( result, thf_xprod_type1left, 
+thf_unitary_type1right), rest671)
+end
+|  ( 67, ( ( _, ( 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 103, ( result, thf_unitary_type1left, 
+thf_unitary_type2right), rest671)
+end
+|  ( 68, ( ( _, ( 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 103, ( result, thf_union_type1left, 
+thf_unitary_type1right), rest671)
+end
+|  ( 69, ( ( _, ( 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 99, ( result, thf_tuple1left, thf_tuple2right), 
+rest671)
+end
+|  ( 70, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.thf_sequent 
+thf_sequent, _, _)) :: ( _, ( _, LPAREN1left, _)) :: rest671)) => let
+ val  result = MlyValue.thf_sequent (( thf_sequent ))
+ in ( LrTable.NT 99, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 71, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( _, LBRKT1left, _)) :: 
+rest671)) => let val  result = MlyValue.thf_tuple (( [] ))
+ in ( LrTable.NT 97, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 72, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( 
+MlyValue.thf_tuple_list thf_tuple_list, _, _)) :: ( _, ( _, LBRKT1left
+, _)) :: rest671)) => let val  result = MlyValue.thf_tuple (
+( thf_tuple_list ))
+ in ( LrTable.NT 97, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 73, ( ( _, ( 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 98, ( result, thf_logic_formula1left, 
+thf_logic_formula1right), rest671)
+end
+|  ( 74, ( ( _, ( 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 98, ( result, thf_logic_formula1left, 
+thf_tuple_list1right), rest671)
+end
+|  ( 75, ( ( _, ( 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 96, ( result, tff_logic_formula1left, 
+tff_logic_formula1right), rest671)
+end
+|  ( 76, ( ( _, ( 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 96, ( result, tff_typed_atom1left, 
+tff_typed_atom1right), rest671)
+end
+|  ( 77, ( ( _, ( MlyValue.tff_sequent tff_sequent, tff_sequent1left, 
+tff_sequent1right)) :: rest671)) => let val  result = 
+MlyValue.tff_formula (( tff_sequent ))
+ in ( LrTable.NT 96, ( result, tff_sequent1left, tff_sequent1right), 
 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)
+|  ( 78, ( ( _, ( 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 95, ( result, tff_binary_formula1left, 
+tff_binary_formula1right), rest671)
+end
+|  ( 79, ( ( _, ( 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 95, ( result, tff_unitary_formula1left, 
+tff_unitary_formula1right), rest671)
+end
+|  ( 80, ( ( _, ( 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 94, ( result, tff_binary_nonassoc1left, 
+tff_binary_nonassoc1right), rest671)
+end
+|  ( 81, ( ( _, ( 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 94, ( result, tff_binary_assoc1left, 
+tff_binary_assoc1right), rest671)
+end
+|  ( 82, ( ( _, ( 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 93, ( result, tff_unitary_formula1left, 
+tff_unitary_formula2right), rest671)
+end
+|  ( 83, ( ( _, ( 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 92, ( result, tff_or_formula1left, 
+tff_or_formula1right), rest671)
+end
+|  ( 84, ( ( _, ( 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 92, ( result, tff_and_formula1left, 
+tff_and_formula1right), rest671)
+end
+|  ( 85, ( ( _, ( 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 91, ( result, tff_unitary_formula1left, 
+tff_unitary_formula2right), rest671)
+end
+|  ( 86, ( ( _, ( 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 91, ( result, tff_or_formula1left, 
+tff_unitary_formula1right), rest671)
+end
+|  ( 87, ( ( _, ( 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 90, ( result, tff_unitary_formula1left, 
+tff_unitary_formula2right), rest671)
+end
+|  ( 88, ( ( _, ( 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 90, ( result, tff_and_formula1left, 
+tff_unitary_formula1right), rest671)
+end
+|  ( 89, ( ( _, ( 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 89, ( result, tff_quantified_formula1left, 
+tff_quantified_formula1right), rest671)
+end
+|  ( 90, ( ( _, ( 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 89, ( result, tff_unary_formula1left, 
+tff_unary_formula1right), rest671)
+end
+|  ( 91, ( ( _, ( MlyValue.atomic_formula atomic_formula, 
+atomic_formula1left, atomic_formula1right)) :: rest671)) => let val  
+result = MlyValue.tff_unitary_formula (( atomic_formula ))
+ in ( LrTable.NT 89, ( result, atomic_formula1left, 
+atomic_formula1right), rest671)
+end
+|  ( 92, ( ( _, ( MlyValue.tff_conditional tff_conditional, 
+tff_conditional1left, tff_conditional1right)) :: rest671)) => let val 
+ result = MlyValue.tff_unitary_formula (( tff_conditional ))
+ in ( LrTable.NT 89, ( result, tff_conditional1left, 
+tff_conditional1right), rest671)
+end
+|  ( 93, ( ( _, ( MlyValue.tff_let tff_let, tff_let1left, 
+tff_let1right)) :: rest671)) => let val  result = 
+MlyValue.tff_unitary_formula (( tff_let ))
+ in ( LrTable.NT 89, ( result, tff_let1left, tff_let1right), 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), 
+|  ( 94, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_logic_formula tff_logic_formula, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_unitary_formula (( tff_logic_formula ))
+ in ( LrTable.NT 89, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 95, ( ( _, ( 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 88, ( result, fol_quantifier1left, 
+tff_unitary_formula1right), rest671)
+end
+|  ( 96, ( ( _, ( MlyValue.tff_variable tff_variable, 
+tff_variable1left, tff_variable1right)) :: rest671)) => let val  
+result = MlyValue.tff_variable_list (( [tff_variable] ))
+ in ( LrTable.NT 87, ( result, tff_variable1left, tff_variable1right),
+ rest671)
+end
+|  ( 97, ( ( _, ( 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 87, ( result, tff_variable1left, 
+tff_variable_list1right), rest671)
+end
+|  ( 98, ( ( _, ( 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 86, ( result, tff_typed_variable1left, 
+tff_typed_variable1right), rest671)
+end
+|  ( 99, ( ( _, ( MlyValue.variable_ variable_, variable_1left, 
+variable_1right)) :: rest671)) => let val  result = 
+MlyValue.tff_variable (( (variable_, NONE) ))
+ in ( LrTable.NT 86, ( result, variable_1left, variable_1right), 
 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), 
+|  ( 100, ( ( _, ( 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 85, ( result, variable_1left, tff_atomic_type1right),
+ rest671)
+end
+|  ( 101, ( ( _, ( 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 84, ( result, unary_connective1left, 
+tff_unitary_formula1right), rest671)
+end
+|  ( 102, ( ( _, ( 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 84, ( result, fol_infix_unary1left, 
+fol_infix_unary1right), rest671)
+end
+|  ( 103, ( ( _, ( _, _, 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
+|  ( 104, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.tff_formula
+ tff_formula, _, _)) :: _ :: ( _, ( MlyValue.tff_let_term_defn 
+tff_let_term_defn, _, _)) :: _ :: ( _, ( _, LET_TF1left, _)) :: 
+rest671)) => let val  result = MlyValue.tff_let (
+(Let (tff_let_term_defn, tff_formula) ))
+ in ( LrTable.NT 137, ( result, LET_TF1left, RPAREN1right), rest671)
+
+end
+|  ( 105, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.tff_formula
+ tff_formula, _, _)) :: _ :: ( _, ( MlyValue.tff_let_formula_defn 
+tff_let_formula_defn, _, _)) :: _ :: ( _, ( _, LET_FF1left, _)) :: 
+rest671)) => let val  result = MlyValue.tff_let (
+( Let (tff_let_formula_defn, tff_formula) ))
+ in ( LrTable.NT 137, ( result, LET_FF1left, RPAREN1right), rest671)
+
+end
+|  ( 106, ( ( _, ( MlyValue.tff_quantified_formula 
+tff_quantified_formula, tff_quantified_formula1left, 
+tff_quantified_formula1right)) :: rest671)) => let val  result = 
+MlyValue.tff_let_term_defn (
+(
+  let
+    val (_, vars, fmla) = extract_quant_info tff_quantified_formula
+  in [Let_fmla (hd vars, fmla)]
+  end
+)
+)
+ in ( LrTable.NT 138, ( result, tff_quantified_formula1left, 
+tff_quantified_formula1right), rest671)
+end
+|  ( 107, ( ( _, ( MlyValue.tff_quantified_formula 
+tff_quantified_formula, tff_quantified_formula1left, 
+tff_quantified_formula1right)) :: rest671)) => let val  result = 
+MlyValue.tff_let_formula_defn (
+(
+  let
+    val (_, vars, fmla) = extract_quant_info tff_quantified_formula
+  in [Let_fmla (hd vars, fmla)]
+  end
+)
+)
+ in ( LrTable.NT 139, ( result, tff_quantified_formula1left, 
+tff_quantified_formula1right), rest671)
+end
+|  ( 108, ( ( _, ( 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
-|  ( 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), 
+|  ( 109, ( ( _, ( _, _, 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
+|  ( 110, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( _, LBRKT1left, _)) ::
+ rest671)) => let val  result = MlyValue.tff_tuple (( [] ))
+ in ( LrTable.NT 73, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 111, ( ( _, ( _, _, 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
+|  ( 112, ( ( _, ( 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
+|  ( 113, ( ( _, ( 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
+|  ( 114, ( ( _, ( 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 83, ( result, tff_untyped_atom1left, 
+tff_top_level_type1right), rest671)
+end
+|  ( 115, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_typed_atom tff_typed_atom, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_typed_atom (( tff_typed_atom ))
+ in ( LrTable.NT 83, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 116, ( ( _, ( MlyValue.functor_ functor_, functor_1left, 
+functor_1right)) :: rest671)) => let val  result = 
+MlyValue.tff_untyped_atom (( (functor_, NONE) ))
+ in ( LrTable.NT 82, ( result, functor_1left, functor_1right), rest671
+)
+end
+|  ( 117, ( ( _, ( MlyValue.system_functor system_functor, 
+system_functor1left, system_functor1right)) :: rest671)) => let val  
+result = MlyValue.tff_untyped_atom (( (system_functor, NONE) ))
+ in ( LrTable.NT 82, ( result, system_functor1left, 
+system_functor1right), rest671)
+end
+|  ( 118, ( ( _, ( 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 81, ( result, tff_atomic_type1left, 
+tff_atomic_type1right), rest671)
+end
+|  ( 119, ( ( _, ( 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 81, ( result, tff_mapping_type1left, 
+tff_mapping_type1right), rest671)
+end
+|  ( 120, ( ( _, ( MlyValue.tff_quantified_type tff_quantified_type, 
+tff_quantified_type1left, tff_quantified_type1right)) :: rest671)) =>
+ let val  result = MlyValue.tff_top_level_type (
+( tff_quantified_type ))
+ in ( LrTable.NT 81, ( result, tff_quantified_type1left, 
+tff_quantified_type1right), rest671)
+end
+|  ( 121, ( ( _, ( MlyValue.tff_monotype tff_monotype, _, 
+tff_monotype1right)) :: _ :: _ :: ( _, ( MlyValue.tff_variable_list 
+tff_variable_list, _, _)) :: _ :: ( _, ( _, DEP_PROD1left, _)) :: 
+rest671)) => let val  result = MlyValue.tff_quantified_type (
+(
+       Fmla_type (Quant (Dep_Prod, tff_variable_list, Type_fmla tff_monotype))
+)
+)
+ in ( LrTable.NT 140, ( result, DEP_PROD1left, tff_monotype1right), 
 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)
+|  ( 122, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_quantified_type tff_quantified_type, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_quantified_type (( tff_quantified_type ))
+ in ( LrTable.NT 140, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 123, ( ( _, ( MlyValue.tff_atomic_type tff_atomic_type, 
+tff_atomic_type1left, tff_atomic_type1right)) :: rest671)) => let val 
+ result = MlyValue.tff_monotype (( tff_atomic_type ))
+ in ( LrTable.NT 141, ( result, tff_atomic_type1left, 
+tff_atomic_type1right), rest671)
+end
+|  ( 124, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_mapping_type tff_mapping_type, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_monotype (( tff_mapping_type ))
+ in ( LrTable.NT 141, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 125, ( ( _, ( 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 80, ( result, tff_atomic_type1left, 
+tff_atomic_type1right), rest671)
+end
+|  ( 126, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_xprod_type tff_xprod_type, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_unitary_type (( tff_xprod_type ))
+ in ( LrTable.NT 80, ( result, LPAREN1left, RPAREN1right), 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
+|  ( 127, ( ( _, ( MlyValue.atomic_word atomic_word, atomic_word1left,
+ atomic_word1right)) :: rest671)) => let val  result = 
+MlyValue.tff_atomic_type (( Atom_type atomic_word ))
+ in ( LrTable.NT 79, ( result, atomic_word1left, atomic_word1right), 
+rest671)
+end
+|  ( 128, ( ( _, ( MlyValue.defined_type defined_type, 
+defined_type1left, defined_type1right)) :: rest671)) => let val  
+result = MlyValue.tff_atomic_type (( Defined_type defined_type ))
+ in ( LrTable.NT 79, ( result, defined_type1left, defined_type1right),
+ rest671)
+end
+|  ( 129, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_type_arguments tff_type_arguments, _, _)) :: _ :: ( _, ( 
+MlyValue.atomic_word atomic_word, atomic_word1left, _)) :: rest671))
+ => let val  result = MlyValue.tff_atomic_type (
+( Fmla_type (Fmla (Uninterpreted atomic_word, (map Type_fmla tff_type_arguments))) )
+)
+ in ( LrTable.NT 79, ( result, atomic_word1left, RPAREN1right), 
+rest671)
+end
+|  ( 130, ( ( _, ( MlyValue.variable_ variable_, variable_1left, 
+variable_1right)) :: rest671)) => let val  result = 
+MlyValue.tff_atomic_type (
+( Fmla_type (Pred (Interpreted_ExtraLogic Apply, [Term_Var variable_])) )
+)
+ in ( LrTable.NT 79, ( result, variable_1left, variable_1right), 
+rest671)
+end
+|  ( 131, ( ( _, ( MlyValue.tff_atomic_type tff_atomic_type, 
+tff_atomic_type1left, tff_atomic_type1right)) :: rest671)) => let val 
+ result = MlyValue.tff_type_arguments (( [tff_atomic_type]  ))
+ in ( LrTable.NT 142, ( result, tff_atomic_type1left, 
+tff_atomic_type1right), rest671)
+end
+|  ( 132, ( ( _, ( MlyValue.tff_type_arguments tff_type_arguments, _, 
+tff_type_arguments1right)) :: _ :: ( _, ( MlyValue.tff_atomic_type 
+tff_atomic_type, tff_atomic_type1left, _)) :: rest671)) => let val  
+result = MlyValue.tff_type_arguments (
+( tff_atomic_type :: tff_type_arguments ))
+ in ( LrTable.NT 142, ( result, tff_atomic_type1left, 
+tff_type_arguments1right), rest671)
+end
+|  ( 133, ( ( _, ( 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 78, ( result, tff_unitary_type1left, 
+tff_atomic_type1right), rest671)
+end
+|  ( 134, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_mapping_type tff_mapping_type, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_mapping_type (( tff_mapping_type ))
+ in ( LrTable.NT 78, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 135, ( ( _, ( 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 77, ( result, tff_atomic_type1left, 
+tff_atomic_type2right), rest671)
+end
+|  ( 136, ( ( _, ( 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 77, ( result, tff_xprod_type1left, 
+tff_atomic_type1right), rest671)
+end
+|  ( 137, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( 
+MlyValue.tff_xprod_type tff_xprod_type, _, _)) :: ( _, ( _, 
+LPAREN1left, _)) :: rest671)) => let val  result = 
+MlyValue.tff_xprod_type (( tff_xprod_type ))
+ in ( LrTable.NT 77, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 138, ( ( _, ( 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
+|  ( 139, ( ( _, ( 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
+|  ( 140, ( ( _, ( 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
+|  ( 141, ( ( _, ( 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
+|  ( 142, ( ( _, ( 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
+|  ( 143, ( ( _, ( 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
+|  ( 144, ( ( _, ( 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] )
+)
 )
-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
+ in ( LrTable.NT 69, ( result, fof_unitary_formula1left, 
+fof_unitary_formula2right), rest671)
+end
+|  ( 145, ( ( _, ( 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
+|  ( 146, ( ( _, ( 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
+|  ( 147, ( ( _, ( 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
+|  ( 148, ( ( _, ( 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
+|  ( 149, ( ( _, ( 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
+|  ( 150, ( ( _, ( 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
+|  ( 151, ( ( _, ( 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
+|  ( 152, ( ( _, ( 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
+|  ( 153, ( ( _, ( 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
+|  ( 154, ( ( _, ( _, _, 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
+|  ( 155, ( ( _, ( 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
+|  ( 156, ( ( _, ( 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
+|  ( 157, ( ( _, ( 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
-|  ( 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)) =>
+|  ( 158, ( ( _, ( 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
+|  ( 159, ( ( _, ( 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
+|  ( 160, ( ( _, ( 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
+|  ( 161, ( ( _, ( _, _, 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
+|  ( 162, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( _, LBRKT1left, _)) ::
+ rest671)) => let val  result = MlyValue.fof_tuple (( [] ))
+ in ( LrTable.NT 60, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 163, ( ( _, ( _, _, 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
+|  ( 164, ( ( _, ( 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
+|  ( 165, ( ( _, ( 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
+|  ( 166, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.disjunction
+ disjunction, _, _)) :: ( _, ( _, LPAREN1left, _)) :: rest671)) => let
+ val  result = MlyValue.cnf_formula (( disjunction ))
+ in ( LrTable.NT 58, ( result, LPAREN1left, RPAREN1right), rest671)
+
+end
+|  ( 167, ( ( _, ( MlyValue.disjunction disjunction, disjunction1left,
+ disjunction1right)) :: rest671)) => let val  result = 
+MlyValue.cnf_formula (( disjunction ))
+ in ( LrTable.NT 58, ( result, disjunction1left, disjunction1right), 
+rest671)
+end
+|  ( 168, ( ( _, ( MlyValue.literal literal, literal1left, 
+literal1right)) :: rest671)) => let val  result = MlyValue.disjunction
+ (( literal ))
+ in ( LrTable.NT 57, ( result, literal1left, literal1right), rest671)
+
+end
+|  ( 169, ( ( _, ( 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
+|  ( 170, ( ( _, ( 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
+|  ( 171, ( ( _, ( 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
+|  ( 172, ( ( _, ( 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
+|  ( 173, ( ( _, ( 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
+|  ( 174, ( ( _, ( 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
+|  ( 175, ( ( _, ( 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
+|  ( 176, ( ( _, ( 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
+|  ( 177, ( ( _, ( 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
+|  ( 178, ( ( _, ( _, CARET1left, CARET1right)) :: rest671)) => let
+ val  result = MlyValue.thf_quantifier (( Lambda ))
+ in ( LrTable.NT 53, ( result, CARET1left, CARET1right), rest671)
+end
+|  ( 179, ( ( _, ( _, DEP_PROD1left, DEP_PROD1right)) :: rest671)) =>
+ let val  result = MlyValue.thf_quantifier (( Dep_Prod ))
+ in ( LrTable.NT 53, ( result, DEP_PROD1left, DEP_PROD1right), rest671
+)
+end
+|  ( 180, ( ( _, ( _, DEP_SUM1left, DEP_SUM1right)) :: rest671)) =>
+ let val  result = MlyValue.thf_quantifier (( Dep_Sum ))
+ in ( LrTable.NT 53, ( result, DEP_SUM1left, DEP_SUM1right), rest671)
+
+end
+|  ( 181, ( ( _, ( _, INDEF_CHOICE1left, INDEF_CHOICE1right)) :: 
+rest671)) => let val  result = MlyValue.thf_quantifier (( Epsilon ))
+ in ( LrTable.NT 53, ( result, INDEF_CHOICE1left, INDEF_CHOICE1right),
+ rest671)
+end
+|  ( 182, ( ( _, ( _, DEFIN_CHOICE1left, DEFIN_CHOICE1right)) :: 
+rest671)) => let val  result = MlyValue.thf_quantifier (( Iota ))
+ in ( LrTable.NT 53, ( result, DEFIN_CHOICE1left, DEFIN_CHOICE1right),
+ rest671)
+end
+|  ( 183, ( ( _, ( 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
+|  ( 184, ( ( _, ( 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
+|  ( 185, ( ( _, ( 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
+|  ( 186, ( ( _, ( 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
+|  ( 187, ( ( _, ( _, 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
+|  ( 188, ( ( _, ( _, 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
+|  ( 189, ( ( _, ( _, EXCLAMATION1left, EXCLAMATION1right)) :: rest671
+)) => let val  result = MlyValue.fol_quantifier (( Forall ))
+ in ( LrTable.NT 50, ( result, EXCLAMATION1left, EXCLAMATION1right), 
+rest671)
+end
+|  ( 190, ( ( _, ( _, QUESTION1left, QUESTION1right)) :: rest671)) =>
+ let val  result = MlyValue.fol_quantifier (( Exists ))
+ in ( LrTable.NT 50, ( result, QUESTION1left, QUESTION1right), rest671
+)
+end
+|  ( 191, ( ( _, ( _, IFF1left, IFF1right)) :: rest671)) => let val  
+result = MlyValue.binary_connective (( Interpreted_Logic Iff ))
+ in ( LrTable.NT 49, ( result, IFF1left, IFF1right), rest671)
+end
+|  ( 192, ( ( _, ( _, IMPLIES1left, IMPLIES1right)) :: rest671)) =>
+ let val  result = MlyValue.binary_connective (
+( Interpreted_Logic If ))
+ in ( LrTable.NT 49, ( result, IMPLIES1left, IMPLIES1right), rest671)
+
+end
+|  ( 193, ( ( _, ( _, FI1left, FI1right)) :: rest671)) => let val  
+result = MlyValue.binary_connective (( Interpreted_Logic Fi ))
+ in ( LrTable.NT 49, ( result, FI1left, FI1right), rest671)
+end
+|  ( 194, ( ( _, ( _, XOR1left, XOR1right)) :: rest671)) => let val  
+result = MlyValue.binary_connective (( Interpreted_Logic Xor ))
+ in ( LrTable.NT 49, ( result, XOR1left, XOR1right), rest671)
+end
+|  ( 195, ( ( _, ( _, NOR1left, NOR1right)) :: rest671)) => let val  
+result = MlyValue.binary_connective (( Interpreted_Logic Nor ))
+ in ( LrTable.NT 49, ( result, NOR1left, NOR1right), rest671)
+end
+|  ( 196, ( ( _, ( _, NAND1left, NAND1right)) :: rest671)) => let val 
+ result = MlyValue.binary_connective (( Interpreted_Logic Nand ))
+ in ( LrTable.NT 49, ( result, NAND1left, NAND1right), rest671)
+end
+|  ( 197, ( ( _, ( _, VLINE1left, VLINE1right)) :: rest671)) => let
+ val  result = MlyValue.assoc_connective (( Interpreted_Logic Or ))
+ in ( LrTable.NT 48, ( result, VLINE1left, VLINE1right), rest671)
+end
+|  ( 198, ( ( _, ( _, AMPERSAND1left, AMPERSAND1right)) :: rest671))
+ => let val  result = MlyValue.assoc_connective (
+( Interpreted_Logic And ))
+ in ( LrTable.NT 48, ( result, AMPERSAND1left, AMPERSAND1right), 
+rest671)
+end
+|  ( 199, ( ( _, ( _, TILDE1left, TILDE1right)) :: rest671)) => let
+ val  result = MlyValue.unary_connective (( Interpreted_Logic Not ))
+ in ( LrTable.NT 45, ( result, TILDE1left, TILDE1right), rest671)
+end
+|  ( 200, ( ( _, ( 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
+  case atomic_defined_word of
+    "$oType" => Type_Bool
   | "$o" => Type_Bool
   | "$iType" => Type_Ind
-  | "$oType" => Type_Bool
-  | "$int" => Type_Int
+  | "$i" => Type_Ind
+  | "$tType" => Type_Type
   | "$real" => Type_Real
   | "$rat" => Type_Rat
-  | "$tType" => Type_Type
+  | "$int" => Type_Int
   | 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 = 
+ in ( LrTable.NT 46, ( result, atomic_defined_word1left, 
+atomic_defined_word1right), rest671)
+end
+|  ( 201, ( ( _, ( 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
+|  ( 202, ( ( _, ( 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
+|  ( 203, ( ( _, ( 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
+|  ( 204, ( ( _, ( 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
+|  ( 205, ( ( _, ( 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
+|  ( 206, ( ( _, ( 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
+|  ( 207, ( ( _, ( 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
+|  ( 208, ( ( _, ( 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
+|  ( 209, ( ( _, ( 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
+|  ( 210, ( ( _, ( 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
+|  ( 211, ( ( _, ( 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
+|  ( 212, ( ( _, ( 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
+|  ( 213, ( ( _, ( _, EQUALS1left, EQUALS1right)) :: rest671)) => let
+ val  result = MlyValue.infix_equality (( Interpreted_Logic Equals ))
+ in ( LrTable.NT 35, ( result, EQUALS1left, EQUALS1right), rest671)
+
+end
+|  ( 214, ( ( _, ( _, NEQUALS1left, NEQUALS1right)) :: rest671)) =>
+ let val  result = MlyValue.infix_inequality (
+( Interpreted_Logic NEquals ))
+ in ( LrTable.NT 36, ( result, NEQUALS1left, NEQUALS1right), rest671)
+
+end
+|  ( 215, ( ( _, ( 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
+|  ( 216, ( ( _, ( 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
+|  ( 217, ( ( _, ( 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
+|  ( 218, ( ( _, ( 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
+|  ( 219, ( ( _, ( MlyValue.let_term let_term, let_term1left, 
+let_term1right)) :: rest671)) => let val  result = MlyValue.term (
+( let_term ))
+ in ( LrTable.NT 19, ( result, let_term1left, let_term1right), rest671
+)
+end
+|  ( 220, ( ( _, ( 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
+|  ( 221, ( ( _, ( 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
+|  ( 222, ( ( _, ( 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
+|  ( 223, ( ( _, ( MlyValue.constant constant, constant1left, 
+constant1right)) :: rest671)) => let val  result = MlyValue.plain_term
+ (( (constant, []) ))
+ in ( LrTable.NT 31, ( result, constant1left, constant1right), rest671
+)
+end
+|  ( 224, ( ( _, ( _, _, 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
+|  ( 225, ( ( _, ( 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
+|  ( 226, ( ( _, ( 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)
+|  ( 227, ( ( _, ( 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
+|  ( 228, ( ( _, ( 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
+|  ( 229, ( ( _, ( MlyValue.number number, number1left, number1right))
+ :: rest671)) => let val  result = MlyValue.defined_atom (
+( Term_Num number ))
+ in ( LrTable.NT 28, ( result, number1left, number1right), 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), 
+|  ( 230, ( ( _, ( 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
+|  ( 231, ( ( _, ( 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
+|  ( 232, ( ( _, ( 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
+|  ( 233, ( ( _, ( _, _, 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
-|  ( 47, ( ( _, ( MlyValue.ATOMIC_DEFINED_WORD ATOMIC_DEFINED_WORD, 
-ATOMIC_DEFINED_WORD1left, ATOMIC_DEFINED_WORD1right)) :: rest671)) =>
+|  ( 234, ( ( _, ( 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
+|  ( 235, ( ( _, ( 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
+  case atomic_defined_word of
+    "$uminus" => Interpreted_ExtraLogic UMinus
+  | "$sum" => Interpreted_ExtraLogic Sum
   | "$difference" => Interpreted_ExtraLogic Difference
   | "$product" => Interpreted_ExtraLogic Product
   | "$quotient" => Interpreted_ExtraLogic Quotient
@@ -3786,7 +5404,6 @@
   | "$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
@@ -3809,103 +5426,57 @@
   | "$is_int" => Interpreted_ExtraLogic Is_Int
   | "$is_rat" => Interpreted_ExtraLogic Is_Rat
 
+  | "$distinct" => Interpreted_ExtraLogic Distinct
+
   | 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), 
+ in ( LrTable.NT 21, ( result, atomic_defined_word1left, 
+atomic_defined_word1right), rest671)
+end
+|  ( 236, ( ( _, ( 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
+|  ( 237, ( ( _, ( _, _, 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
-|  ( 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)
+|  ( 238, ( ( _, ( 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
+|  ( 239, ( ( _, ( 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
+|  ( 240, ( ( _, ( 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
+|  ( 241, ( ( _, ( MlyValue.term term, term1left, term1right)) :: 
+rest671)) => let val  result = MlyValue.arguments (( [term] ))
+ in ( LrTable.NT 20, ( result, term1left, term1right), rest671)
+end
+|  ( 242, ( ( _, ( 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
-|  ( 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,
+|  ( 243, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.term term2,
  _, _)) :: _ :: ( _, ( MlyValue.term term1, _, _)) :: _ :: ( _, ( 
 MlyValue.tff_logic_formula tff_logic_formula, _, _)) :: _ :: ( _, ( _,
  ITE_T1left, _)) :: rest671)) => let val  result = 
@@ -3915,1522 +5486,271 @@
 ))
  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)
+|  ( 244, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.term term,
+ _, _)) :: _ :: ( _, ( MlyValue.tff_let_formula_defn 
+tff_let_formula_defn, _, _)) :: _ :: ( _, ( _, LET_FT1left, _)) :: 
+rest671)) => let val  result = MlyValue.let_term (
+(Term_Let (tff_let_formula_defn, term) ))
+ in ( LrTable.NT 143, ( result, LET_FT1left, 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)
+|  ( 245, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.term term,
+ _, _)) :: _ :: ( _, ( MlyValue.tff_let_term_defn tff_let_term_defn, _
+, _)) :: _ :: ( _, ( _, LET_TT1left, _)) :: rest671)) => let val  
+result = MlyValue.let_term ((Term_Let (tff_let_term_defn, term) ))
+ in ( LrTable.NT 143, ( result, LET_TT1left, 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)
+|  ( 246, ( ( _, ( 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
-|  ( 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)) :: _ :: ( _, ( 
+|  ( 247, ( rest671)) => let val  result = MlyValue.optional_info (
+( [] ))
+ in ( LrTable.NT 4, ( result, defaultPos, defaultPos), rest671)
+end
+|  ( 248, ( ( _, ( 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
+|  ( 249, ( ( _, ( _, _, 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)
+ in ( LrTable.NT 132, ( result, INCLUDE1left, PERIOD1right), rest671)
 
 end
-|  ( 270, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( MlyValue.name_list 
+|  ( 250, ( ( _, ( _, _, 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
+|  ( 251, ( rest671)) => let val  result = MlyValue.formula_selection
  (( [] ))
  in ( LrTable.NT 3, ( result, defaultPos, defaultPos), rest671)
 end
-|  ( 272, ( ( _, ( MlyValue.name_list name_list, _, name_list1right))
+|  ( 252, ( ( _, ( 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)) :: 
+|  ( 253, ( ( _, ( 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,
+|  ( 254, ( ( _, ( 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
+|  ( 255, ( ( _, ( 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
+|  ( 256, ( ( _, ( 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
+|  ( 257, ( ( _, ( 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
+|  ( 258, ( ( _, ( 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
+|  ( 259, ( ( _, ( 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
+|  ( 260, ( ( _, ( MlyValue.number number, number1left, number1right))
+ :: rest671)) => let val  result = MlyValue.general_data (
+( Number number ))
+ in ( LrTable.NT 9, ( result, number1left, number1right), rest671)
+end
+|  ( 261, ( ( _, ( 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
+|  ( 262, ( ( _, ( 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
+|  ( 263, ( ( _, ( _, _, 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
+|  ( 264, ( ( _, ( _, _, 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
+|  ( 265, ( ( _, ( _, _, 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
+|  ( 266, ( ( _, ( _, _, 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
+|  ( 267, ( ( _, ( _, _, 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
+|  ( 268, ( ( _, ( _, _, RPAREN1right)) :: ( _, ( MlyValue.term term,
+ _, _)) :: _ :: ( _, ( _, DFOT1left, _)) :: rest671)) => let val  
+result = MlyValue.formula_data (( Term_Data term ))
+ in ( LrTable.NT 12, ( result, DFOT1left, RPAREN1right), rest671)
+end
+|  ( 269, ( ( _, ( _, _, 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
+|  ( 270, ( ( _, ( _, _, RBRKT1right)) :: ( _, ( _, LBRKT1left, _)) ::
+ rest671)) => let val  result = MlyValue.general_list (( [] ))
+ in ( LrTable.NT 5, ( result, LBRKT1left, RBRKT1right), rest671)
+end
+|  ( 271, ( ( _, ( 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
+|  ( 272, ( ( _, ( 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
+|  ( 273, ( ( _, ( 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, 
+|  ( 274, ( ( _, ( 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), 
+|  ( 275, ( ( _, ( 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
+|  ( 276, ( ( _, ( 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
+|  ( 277, ( ( _, ( _, THF1left, THF1right)) :: rest671)) => let val  
+result = MlyValue.atomic_word (( "thf" ))
+ in ( LrTable.NT 8, ( result, THF1left, THF1right), rest671)
+end
+|  ( 278, ( ( _, ( _, TFF1left, TFF1right)) :: rest671)) => let val  
+result = MlyValue.atomic_word (( "tff" ))
+ in ( LrTable.NT 8, ( result, TFF1left, TFF1right), rest671)
+end
+|  ( 279, ( ( _, ( _, FOF1left, FOF1right)) :: rest671)) => let val  
+result = MlyValue.atomic_word (( "fof" ))
+ in ( LrTable.NT 8, ( result, FOF1left, FOF1right), rest671)
+end
+|  ( 280, ( ( _, ( _, CNF1left, CNF1right)) :: rest671)) => let val  
+result = MlyValue.atomic_word (( "cnf" ))
+ in ( LrTable.NT 8, ( result, CNF1left, CNF1right), rest671)
+end
+|  ( 281, ( ( _, ( _, INCLUDE1left, INCLUDE1right)) :: rest671)) =>
+ let val  result = MlyValue.atomic_word (( "include" ))
+ in ( LrTable.NT 8, ( result, INCLUDE1left, INCLUDE1right), rest671)
+
+end
+|  ( 282, ( ( _, ( MlyValue.DOLLAR_WORD DOLLAR_WORD, DOLLAR_WORD1left,
+ DOLLAR_WORD1right)) :: rest671)) => let val  result = 
+MlyValue.atomic_defined_word (( DOLLAR_WORD ))
+ in ( LrTable.NT 144, ( result, DOLLAR_WORD1left, DOLLAR_WORD1right), 
 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)
+|  ( 283, ( ( _, ( MlyValue.DOLLAR_DOLLAR_WORD DOLLAR_DOLLAR_WORD, 
+DOLLAR_DOLLAR_WORD1left, DOLLAR_DOLLAR_WORD1right)) :: rest671)) =>
+ let val  result = MlyValue.atomic_system_word (( DOLLAR_DOLLAR_WORD )
+)
+ in ( LrTable.NT 145, ( result, DOLLAR_DOLLAR_WORD1left, 
+DOLLAR_DOLLAR_WORD1right), rest671)
+end
+|  ( 284, ( ( _, ( 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
+|  ( 285, ( ( _, ( 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
+|  ( 286, ( ( _, ( MlyValue.integer integer, integer1left, 
+integer1right)) :: rest671)) => let val  result = MlyValue.number (
+( (Int_num, integer) ))
+ in ( LrTable.NT 11, ( result, integer1left, integer1right), rest671)
+
+end
+|  ( 287, ( ( _, ( MlyValue.REAL REAL, REAL1left, REAL1right)) :: 
+rest671)) => let val  result = MlyValue.number (( (Real_num, REAL) ))
+ in ( LrTable.NT 11, ( result, REAL1left, REAL1right), rest671)
+end
+|  ( 288, ( ( _, ( MlyValue.RATIONAL RATIONAL, RATIONAL1left, 
+RATIONAL1right)) :: rest671)) => let val  result = MlyValue.number (
+( (Rat_num, RATIONAL) ))
+ in ( LrTable.NT 11, ( result, RATIONAL1left, RATIONAL1right), rest671
+)
+end
+|  ( 289, ( ( _, ( 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
 | _ => raise (mlyAction i392)
 end
@@ -5462,7 +5782,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,(
+fun FI (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))
@@ -5570,10 +5890,10 @@
 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 DOLLAR_WORD (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 63,(
+ParserData.MlyValue.DOLLAR_WORD i,p1,p2))
+fun DOLLAR_DOLLAR_WORD (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 
+64,(ParserData.MlyValue.DOLLAR_DOLLAR_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,(
@@ -5590,5 +5910,13 @@
 ParserData.MlyValue.VOID,p1,p2))
 fun ITE_T (p1,p2) = Token.TOKEN (ParserData.LrTable.T 72,(
 ParserData.MlyValue.VOID,p1,p2))
-end
-end
+fun LET_TF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 73,(
+ParserData.MlyValue.VOID,p1,p2))
+fun LET_FF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 74,(
+ParserData.MlyValue.VOID,p1,p2))
+fun LET_FT (p1,p2) = Token.TOKEN (ParserData.LrTable.T 75,(
+ParserData.MlyValue.VOID,p1,p2))
+fun LET_TT (p1,p2) = Token.TOKEN (ParserData.LrTable.T 76,(
+ParserData.MlyValue.VOID,p1,p2))
+end
+end
--- a/src/HOL/TPTP/TPTP_Parser/tptp_syntax.ML	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_syntax.ML	Wed Apr 04 20:45:19 2012 +0200
@@ -53,13 +53,13 @@
     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*)
+    (*FIXME 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*)
+    Distinct | Apply
 
   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*)
+    (*FIXME these should be in defined_pred, but that's not being used in TPTP*)
     True | False
 
   and quantifier = (*interpreted binders*)
@@ -86,6 +86,7 @@
     | Term_Conditional of tptp_formula * tptp_term * tptp_term
     | Term_Num of number_kind * string
     | Term_Distinct_Object of string
+    | Term_Let of tptp_let list * tptp_term (*FIXME remove list?*)
 
   and tptp_atom =
       TFF_Typed_Atom of symbol * tptp_type option (*only TFF*)
@@ -98,14 +99,14 @@
     | 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
+    | Let of tptp_let list * tptp_formula (*FIXME remove list?*)
     | Atom of tptp_atom
-    | THF_type of tptp_type
+    | Type_fmla 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*)
+    | Let_term of (string * tptp_type option) * tptp_term (*only TFF*)
 
   and tptp_type =
       Prod_type of tptp_type * tptp_type
@@ -113,7 +114,7 @@
     | 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*)
+    | Fmla_type of tptp_formula
     | Subtype of symbol * symbol (*only THF*)
 
   type general_list = general_term list
@@ -128,7 +129,8 @@
   type position = string * int * int
 
   datatype tptp_line =
-      Annotated_Formula of position * language * string * role * tptp_formula * annotation option
+      Annotated_Formula of position * language * string * role *
+        tptp_formula * annotation option
    |  Include of string * string list
 
   type tptp_problem = tptp_line list
@@ -196,13 +198,12 @@
     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*)
+    Distinct |
+    Apply
 
   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*)
@@ -229,6 +230,7 @@
     | Term_Conditional of tptp_formula * tptp_term * tptp_term
     | Term_Num of number_kind * string
     | Term_Distinct_Object of string
+    | Term_Let of tptp_let list * tptp_term (*FIXME remove list?*)
 
   and tptp_atom =
       TFF_Typed_Atom of symbol * tptp_type option (*only TFF*)
@@ -243,21 +245,21 @@
     | Conditional of tptp_formula * tptp_formula * tptp_formula
     | Let of tptp_let list * tptp_formula
     | Atom of tptp_atom
-    | THF_type of tptp_type
+    | Type_fmla 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*)
+      Let_fmla of (string * tptp_type option) * tptp_formula
+    | Let_term of (string * tptp_type option) * tptp_term
 
   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*)
+    | Sum_type of tptp_type * tptp_type
+    | Fmla_type of tptp_formula
+    | Subtype of symbol * symbol
 
 type general_list = general_term list
 type parent_details = general_list
@@ -269,13 +271,6 @@
 
 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 =
@@ -470,7 +465,7 @@
   | 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 (Type_fmla 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
 
--- a/src/HOL/Tools/Lifting/lifting_def.ML	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/Tools/Lifting/lifting_def.ML	Wed Apr 04 20:45:19 2012 +0200
@@ -177,7 +177,7 @@
     val ((_, (_ , def_thm)), lthy') = 
       Local_Theory.define (var, ((Thm.def_binding (#1 var), []), newrhs)) lthy
 
-    val transfer_thm = def_thm RS (rsp_thm RS (quotient_thm RS @{thm Quotient_to_transfer}))
+    val transfer_thm = [quotient_thm, rsp_thm, def_thm] MRSL @{thm Quotient_to_transfer}
 
     fun qualify defname suffix = Binding.name suffix
       |> Binding.qualify true defname
--- a/src/HOL/Tools/Lifting/lifting_setup.ML	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/Tools/Lifting/lifting_setup.ML	Wed Apr 04 20:45:19 2012 +0200
@@ -22,26 +22,11 @@
 
 exception SETUP_LIFTING_INFR of string
 
-fun define_cr_rel equiv_thm abs_fun lthy =
+fun define_cr_rel rep_fun lthy =
   let
-    fun force_type_of_rel rel forced_ty =
-      let
-        val thy = Proof_Context.theory_of lthy
-        val rel_ty = (domain_type o fastype_of) rel
-        val ty_inst = Sign.typ_match thy (rel_ty, forced_ty) Vartab.empty
-      in
-        Envir.subst_term_types ty_inst rel
-      end
-
-    val (rty, qty) = (dest_funT o fastype_of) abs_fun
-    val abs_fun_graph = HOLogic.mk_eq(abs_fun $ Bound 1, Bound 0)
-    val Abs_body = (case (HOLogic.dest_Trueprop o prop_of) equiv_thm of
-      Const (@{const_name equivp}, _) $ _ => abs_fun_graph
-      | Const (@{const_name part_equivp}, _) $ rel => 
-        HOLogic.mk_conj (force_type_of_rel rel rty $ Bound 1 $ Bound 1, abs_fun_graph)
-      | _ => raise SETUP_LIFTING_INFR "unsupported equivalence theorem"
-      )
-    val def_term = Abs ("x", rty, Abs ("y", qty, Abs_body));
+    val (qty, rty) = (dest_funT o fastype_of) rep_fun
+    val rep_fun_graph = (HOLogic.eq_const rty) $ Bound 1 $ (rep_fun $ Bound 0)
+    val def_term = Abs ("x", rty, Abs ("y", qty, rep_fun_graph));
     val qty_name = (fst o dest_Type) qty
     val cr_rel_name = Binding.prefix_name "cr_" (Binding.qualified_name qty_name)  
     val (fixed_def_term, lthy') = yield_singleton (Variable.importT_terms) def_term lthy
@@ -82,9 +67,9 @@
   let
     fun derive_quot_and_equiv_thm to_quot_thm to_equiv_thm typedef_thm lthy =
       let
-        val (_ $ _ $ abs_fun $ _) = (HOLogic.dest_Trueprop o prop_of) typedef_thm
+        val (_ $ rep_fun $ _ $ _) = (HOLogic.dest_Trueprop o prop_of) typedef_thm
         val equiv_thm = typedef_thm RS to_equiv_thm
-        val (T_def, lthy') = define_cr_rel equiv_thm abs_fun lthy
+        val (T_def, lthy') = define_cr_rel rep_fun lthy
         val quot_thm =  [typedef_thm, T_def] MRSL to_quot_thm
       in
         (quot_thm, equiv_thm, lthy')
@@ -93,12 +78,13 @@
     val typedef_set = (snd o dest_comb o HOLogic.dest_Trueprop o prop_of) typedef_thm
     val (quot_thm, equiv_thm, lthy') = (case typedef_set of
       Const ("Orderings.top_class.top", _) => 
-        derive_quot_and_equiv_thm @{thm copy_type_to_Quotient} @{thm copy_type_to_equivp} 
+        derive_quot_and_equiv_thm @{thm UNIV_typedef_to_Quotient} @{thm UNIV_typedef_to_equivp} 
           typedef_thm lthy
       | Const (@{const_name "Collect"}, _) $ Abs (_, _, _) => 
-        derive_quot_and_equiv_thm @{thm invariant_type_to_Quotient} @{thm invariant_type_to_part_equivp} 
+        derive_quot_and_equiv_thm @{thm open_typedef_to_Quotient} @{thm open_typedef_to_part_equivp} 
           typedef_thm lthy
-      | _ => raise SETUP_LIFTING_INFR "unsupported typedef theorem")
+      | _ => derive_quot_and_equiv_thm @{thm typedef_to_Quotient} @{thm typedef_to_part_equivp} 
+          typedef_thm lthy)
   in
     setup_lifting_infr quot_thm equiv_thm lthy'
   end
@@ -109,5 +95,4 @@
   Outer_Syntax.local_theory @{command_spec "setup_lifting"}
     "Setup lifting infrastructure" 
       (Parse_Spec.xthm >> (fn xthm => setup_by_typedef_thm_aux xthm))
-
 end;
--- a/src/HOL/Tools/Quotient/quotient_type.ML	Wed Apr 04 16:48:39 2012 +0200
+++ b/src/HOL/Tools/Quotient/quotient_type.ML	Wed Apr 04 20:45:19 2012 +0200
@@ -99,6 +99,54 @@
   else
     lthy
 
+infix 0 MRSL
+
+fun ants MRSL thm = fold (fn rl => fn thm => rl RS thm) ants thm
+
+fun define_cr_rel equiv_thm abs_fun lthy =
+  let
+    fun force_type_of_rel rel forced_ty =
+      let
+        val thy = Proof_Context.theory_of lthy
+        val rel_ty = (domain_type o fastype_of) rel
+        val ty_inst = Sign.typ_match thy (rel_ty, forced_ty) Vartab.empty
+      in
+        Envir.subst_term_types ty_inst rel
+      end
+  
+    val (rty, qty) = (dest_funT o fastype_of) abs_fun
+    val abs_fun_graph = HOLogic.mk_eq(abs_fun $ Bound 1, Bound 0)
+    val Abs_body = (case (HOLogic.dest_Trueprop o prop_of) equiv_thm of
+      Const (@{const_name equivp}, _) $ _ => abs_fun_graph
+      | Const (@{const_name part_equivp}, _) $ rel => 
+        HOLogic.mk_conj (force_type_of_rel rel rty $ Bound 1 $ Bound 1, abs_fun_graph)
+      | _ => error "unsupported equivalence theorem"
+      )
+    val def_term = Abs ("x", rty, Abs ("y", qty, Abs_body));
+    val qty_name = (fst o dest_Type) qty
+    val cr_rel_name = Binding.prefix_name "cr_" (Binding.qualified_name qty_name)  
+    val (fixed_def_term, lthy') = yield_singleton (Variable.importT_terms) def_term lthy
+    val ((_, (_ , def_thm)), lthy'') =
+      Local_Theory.define ((cr_rel_name, NoSyn), ((Thm.def_binding cr_rel_name, []), fixed_def_term)) lthy'
+  in
+    (def_thm, lthy'')
+  end;
+
+fun setup_lifting_package quot3_thm equiv_thm lthy =
+  let
+    val (_ $ _ $ abs_fun $ _) = (HOLogic.dest_Trueprop o prop_of) quot3_thm
+    val (T_def, lthy') = define_cr_rel equiv_thm abs_fun lthy
+    val quot_thm = (case (HOLogic.dest_Trueprop o prop_of) equiv_thm of
+      Const (@{const_name equivp}, _) $ _ => 
+        [quot3_thm, T_def, equiv_thm] MRSL @{thm Quotient3_to_Quotient_equivp}
+      | Const (@{const_name part_equivp}, _) $ _ => 
+        [quot3_thm, T_def] MRSL @{thm Quotient3_to_Quotient}
+      | _ => error "unsupported equivalence theorem"
+      )
+  in
+    Lifting_Setup.setup_lifting_infr quot_thm equiv_thm lthy'
+  end
+
 fun init_quotient_infr quot_thm equiv_thm lthy =
   let
     val (_ $ rel $ abs $ rep) = (HOLogic.dest_Trueprop o prop_of) quot_thm
@@ -115,6 +163,7 @@
         (fn phi => Quotient_Info.update_quotients qty_full_name (quot_info phi)
           #> Quotient_Info.update_abs_rep qty_full_name (abs_rep_info phi))
       |> define_abs_type quot_thm
+      |> setup_lifting_package quot_thm equiv_thm
   end
 
 (* main function for constructing a quotient type *)