src/HOL/Tools/res_clause.ML
author paulson
Wed, 19 Apr 2006 10:42:45 +0200
changeset 19447 d4f3f8f9c0a5
parent 19443 e32a4703d834
child 19521 cfdab6a91332
permissions -rw-r--r--
the "th" field of type "clause"
     1 (*  Author: Jia Meng, Cambridge University Computer Laboratory
     2     ID: $Id$
     3     Copyright 2004 University of Cambridge
     4 
     5 ML data structure for storing/printing FOL clauses and arity clauses.
     6 Typed equality is treated differently.
     7 *)
     8 
     9 signature RES_CLAUSE =
    10   sig
    11   exception CLAUSE of string * term
    12   type clause and arityClause and classrelClause
    13   type fol_type
    14   type typ_var
    15   type type_literal
    16   val add_typs_aux : (typ_var * string list) list -> type_literal list * type_literal list
    17   val arity_clause_thy: theory -> arityClause list 
    18   val ascii_of : string -> string
    19   val bracket_pack : string list -> string
    20   val check_var_pairs: ''a * ''b -> (''a * ''b) list -> int
    21   val classrel_clauses_thy: theory -> classrelClause list 
    22   val clause_eq : clause * clause -> bool
    23   val clause_prefix : string 
    24   val clause2tptp : clause -> string * string list
    25   val const_prefix : string
    26   val dfg_write_file: thm list -> string -> (clause list * classrelClause list * arityClause list) -> unit
    27   val fixed_var_prefix : string
    28   val gen_tptp_cls : int * string * string * string -> string
    29   val gen_tptp_type_cls : int * string * string * string * int -> string
    30   val get_axiomName : clause ->  string
    31   val hash_clause : clause -> int
    32   val init : theory -> unit
    33   val isMeta : string -> bool
    34   val isTaut : clause -> bool
    35   val keep_types : bool ref
    36   val list_ord : ('a * 'b -> order) -> 'a list * 'b list -> order
    37   val make_axiom_clause : thm -> string * int -> clause option
    38   val make_conjecture_clauses : thm list -> clause list
    39   val make_fixed_const : string -> string		
    40   val make_fixed_type_const : string -> string   
    41   val make_fixed_type_var : string -> string
    42   val make_fixed_var : string -> string
    43   val make_schematic_type_var : string * int -> string
    44   val make_schematic_var : string * int -> string
    45   val make_type_class : string -> string
    46   val mk_fol_type: string * string * fol_type list -> fol_type
    47   val mk_typ_var_sort : Term.typ -> typ_var * sort
    48   val paren_pack : string list -> string
    49   val schematic_var_prefix : string
    50   val special_equal : bool ref
    51   val string_of_fol_type : fol_type -> string
    52   val tconst_prefix : string 
    53   val tfree_prefix : string
    54   val tptp_arity_clause : arityClause -> string
    55   val tptp_classrelClause : classrelClause -> string
    56   val tptp_of_typeLit : type_literal -> string
    57   val tptp_tfree_clause : string -> string
    58   val tptp_write_file: thm list -> string -> ((thm * (string * int)) list * classrelClause list * arityClause list) -> unit
    59   val tvar_prefix : string
    60   val types_eq: fol_type list * fol_type list -> (string*string) list * (string*string) list -> bool * ((string*string) list * (string*string) list)
    61   val types_ord : fol_type list * fol_type list -> order
    62   val union_all : ''a list list -> ''a list
    63   val writeln_strs: TextIO.outstream -> TextIO.vector list -> unit
    64   end;
    65 
    66 structure ResClause : RES_CLAUSE =
    67 struct
    68 
    69 (* Added for typed equality *)
    70 val special_equal = ref false; (* by default,equality does not carry type information *)
    71 val eq_typ_wrapper = "typeinfo"; (* default string *)
    72 
    73 
    74 val schematic_var_prefix = "V_";
    75 val fixed_var_prefix = "v_";
    76 
    77 val tvar_prefix = "T_";
    78 val tfree_prefix = "t_";
    79 
    80 val clause_prefix = "cls_"; 
    81 val arclause_prefix = "clsarity_" 
    82 val clrelclause_prefix = "clsrel_";
    83 
    84 val const_prefix = "c_";
    85 val tconst_prefix = "tc_"; 
    86 val class_prefix = "class_"; 
    87 
    88 fun union_all xss = foldl (op union) [] xss;
    89 
    90 (*Provide readable names for the more common symbolic functions*)
    91 val const_trans_table =
    92       Symtab.make [("op =", "equal"),
    93 	  	   ("Orderings.less_eq", "lessequals"),
    94 		   ("Orderings.less", "less"),
    95 		   ("op &", "and"),
    96 		   ("op |", "or"),
    97 		   ("HOL.plus", "plus"),
    98 		   ("HOL.minus", "minus"),
    99 		   ("HOL.times", "times"),
   100 		   ("Divides.op div", "div"),
   101 		   ("HOL.divide", "divide"),
   102 		   ("op -->", "implies"),
   103 		   ("{}", "emptyset"),
   104 		   ("op :", "in"),
   105 		   ("op Un", "union"),
   106 		   ("op Int", "inter"),
   107 		   ("List.op @", "append")];
   108 
   109 val type_const_trans_table =
   110       Symtab.make [("*", "prod"),
   111 	  	   ("+", "sum"),
   112 		   ("~=>", "map")];
   113 
   114 (*Escaping of special characters.
   115   Alphanumeric characters are left unchanged.
   116   The character _ goes to __
   117   Characters in the range ASCII space to / go to _A to _P, respectively.
   118   Other printing characters go to _NNN where NNN is the decimal ASCII code.*)
   119 local
   120 
   121 val A_minus_space = Char.ord #"A" - Char.ord #" ";
   122 
   123 fun ascii_of_c c =
   124   if Char.isAlphaNum c then String.str c
   125   else if c = #"_" then "__"
   126   else if #" " <= c andalso c <= #"/" 
   127        then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
   128   else if Char.isPrint c then ("_" ^ Int.toString (Char.ord c))
   129   else ""
   130 
   131 in
   132 
   133 val ascii_of = String.translate ascii_of_c;
   134 
   135 end;
   136 
   137 (* convert a list of strings into one single string; surrounded by brackets *)
   138 fun paren_pack [] = ""   (*empty argument list*)
   139   | paren_pack strings = "(" ^ commas strings ^ ")";
   140 
   141 fun bracket_pack strings = "[" ^ commas strings ^ "]";
   142 
   143 
   144 (*Remove the initial ' character from a type variable, if it is present*)
   145 fun trim_type_var s =
   146   if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
   147   else error ("trim_type: Malformed type variable encountered: " ^ s);
   148 
   149 fun ascii_of_indexname (v,0) = ascii_of v
   150   | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
   151 
   152 fun make_schematic_var v = schematic_var_prefix ^ (ascii_of_indexname v);
   153 fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x);
   154 
   155 fun make_schematic_type_var (x,i) = 
   156       tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
   157 fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
   158 
   159 fun lookup_const c =
   160     case Symtab.lookup const_trans_table c of
   161         SOME c' => c'
   162       | NONE => ascii_of c;
   163 
   164 fun lookup_type_const c = 
   165     case Symtab.lookup type_const_trans_table c of
   166         SOME c' => c'
   167       | NONE => ascii_of c;
   168 
   169 fun make_fixed_const "op =" = "equal"   (*MUST BE "equal" because it's built-in to ATPs*)
   170   | make_fixed_const c      = const_prefix ^ lookup_const c;
   171 
   172 fun make_fixed_type_const c = tconst_prefix ^ lookup_type_const c;
   173 
   174 fun make_type_class clas = class_prefix ^ ascii_of clas;
   175 
   176 
   177 (***** definitions and functions for FOL clauses, for conversion to TPTP or DFG format. *****)
   178 
   179 val keep_types = ref true;
   180 
   181 datatype kind = Axiom | Hypothesis | Conjecture;
   182 fun name_of_kind Axiom = "axiom"
   183   | name_of_kind Hypothesis = "hypothesis"
   184   | name_of_kind Conjecture = "conjecture";
   185 
   186 type clause_id = int;
   187 type axiom_name = string;
   188 
   189 
   190 type polarity = bool;
   191 
   192 (* "tag" is used for vampire specific syntax FIXME REMOVE *)
   193 type tag = bool; 
   194 
   195 
   196 (**** Isabelle FOL clauses ****)
   197 
   198 val tagged = ref false;
   199 
   200 type pred_name = string;
   201 
   202 datatype typ_var = FOLTVar of indexname | FOLTFree of string;
   203 
   204 (*FIXME: give the constructors more sensible names*)
   205 datatype fol_type = AtomV of string
   206 		  | AtomF of string
   207 		  | Comp of string * fol_type list;
   208 
   209 fun string_of_fol_type (AtomV x) = x
   210   | string_of_fol_type (AtomF x) = x
   211   | string_of_fol_type (Comp(tcon,tps)) = 
   212       tcon ^ (paren_pack (map string_of_fol_type tps));
   213       
   214 fun mk_fol_type ("Var",x,_) = AtomV(x)
   215   | mk_fol_type ("Fixed",x,_) = AtomF(x)
   216   | mk_fol_type ("Comp",con,args) = Comp(con,args)
   217 
   218 
   219 (*First string is the type class; the second is a TVar or TFfree*)
   220 datatype type_literal = LTVar of string * string | LTFree of string * string;
   221 
   222 datatype fol_term = UVar of string * fol_type
   223                  | Fun of string * fol_type list * fol_term list;
   224 datatype predicate = Predicate of pred_name * fol_type list * fol_term list;
   225 
   226 datatype literal = Literal of polarity * predicate * tag;
   227 
   228 fun mk_typ_var_sort (TFree(a,s)) = (FOLTFree a,s)
   229   | mk_typ_var_sort (TVar(v,s)) = (FOLTVar v,s);
   230 
   231 
   232 (*A clause has first-order literals and other, type-related literals*)
   233 datatype clause = 
   234 	 Clause of {clause_id: clause_id,
   235 		    axiom_name: axiom_name,
   236 		    th: thm,
   237 		    kind: kind,
   238 		    literals: literal list,
   239 		    types_sorts: (typ_var * sort) list};
   240 
   241 fun get_axiomName (Clause cls) = #axiom_name cls;
   242 
   243 exception CLAUSE of string * term;
   244 
   245 fun isFalse (Literal (pol,Predicate(pname,_,[]),_)) =
   246       (pol andalso pname = "c_False") orelse
   247       (not pol andalso pname = "c_True")
   248   | isFalse _ = false;
   249 
   250 fun isTrue (Literal (pol,Predicate(pname,_,[]),_)) =
   251       (pol andalso pname = "c_True") orelse
   252       (not pol andalso pname = "c_False")
   253   | isTrue _ = false;
   254   
   255 fun isTaut (Clause {literals,...}) = exists isTrue literals;  
   256 
   257 fun make_clause (clause_id, axiom_name, th, kind, literals, types_sorts) =
   258   if forall isFalse literals 
   259   then error "Problem too trivial for resolution (empty clause)"
   260   else
   261      Clause {clause_id = clause_id, axiom_name = axiom_name, 
   262              th = th, kind = kind, 
   263              literals = literals, types_sorts = types_sorts};
   264 
   265 
   266 (*Declarations of the current theory--to allow suppressing types.*)
   267 val const_typargs = ref (Library.K [] : (string*typ -> typ list));
   268 
   269 fun num_typargs(s,T) = if !keep_types then length (!const_typargs (s,T)) else 0;
   270 
   271 (*Initialize the type suppression mechanism with the current theory before
   272     producing any clauses!*)
   273 fun init thy = (const_typargs := Sign.const_typargs thy);
   274     
   275 
   276 (*Flatten a type to a fol_type while accumulating sort constraints on the TFrees and
   277   TVars it contains.*)    
   278 fun type_of (Type (a, Ts)) = 
   279       let val (folTyps, ts) = types_of Ts 
   280 	  val t = make_fixed_type_const a
   281       in (Comp(t,folTyps), ts) end
   282   | type_of (TFree (a,s)) = (AtomF(make_fixed_type_var a), [(FOLTFree a, s)]) 
   283   | type_of (TVar (v, s)) = (AtomV(make_schematic_type_var v), [(FOLTVar v, s)])
   284 and types_of Ts =
   285       let val (folTyps,ts) = ListPair.unzip (map type_of Ts)
   286       in (folTyps, union_all ts) end;
   287 
   288 
   289 fun const_types_of (c,T) = types_of (!const_typargs (c,T));
   290 
   291 (* Any variables created via the METAHYPS tactical should be treated as
   292    universal vars, although it is represented as "Free(...)" by Isabelle *)
   293 val isMeta = String.isPrefix "METAHYP1_"
   294 
   295 fun pred_name_type (Const(c,T)) = (make_fixed_const c, const_types_of (c,T))
   296   | pred_name_type (Free(x,T))  = 
   297       if isMeta x then raise CLAUSE("Predicate Not First Order 1", Free(x,T)) 
   298       else (make_fixed_var x, ([],[]))
   299   | pred_name_type (v as Var _) = raise CLAUSE("Predicate Not First Order 2", v)
   300   | pred_name_type t        = raise CLAUSE("Predicate input unexpected", t);
   301 
   302 
   303 (* For typed equality *)
   304 (* here "arg_typ" is the type of "="'s argument's type, not the type of the equality *)
   305 (* Find type of equality arg *)
   306 fun eq_arg_type (Type("fun",[T,_])) = 
   307     let val (folT,_) = type_of T;
   308     in  folT  end;
   309 
   310 fun fun_name_type (Const(c,T)) args = (make_fixed_const c, const_types_of (c,T))
   311   | fun_name_type (Free(x,T)) args  = 
   312        if isMeta x then raise CLAUSE("Function Not First Order", Free(x,T))
   313        else (make_fixed_var x, ([],[]))
   314   | fun_name_type f args = raise CLAUSE("Function Not First Order 1", f);
   315 
   316 (*Convert a term to a fol_term while accumulating sort constraints on the TFrees and
   317   TVars it contains.*)    
   318 fun term_of (Var(ind_nm,T)) = 
   319       let val (folType,ts) = type_of T
   320       in (UVar(make_schematic_var ind_nm, folType), ts) end
   321   | term_of (Free(x,T)) = 
   322       let val (folType, ts) = type_of T
   323       in
   324 	  if isMeta x then (UVar(make_schematic_var(x,0),folType), ts)
   325 	  else (Fun(make_fixed_var x, [folType], []), ts)
   326       end
   327   | term_of app = 
   328       let val (f,args) = strip_comb app
   329 	  val (funName,(contys,ts1)) = fun_name_type f args
   330 	  val (args',ts2) = terms_of args
   331       in
   332 	  (Fun(funName,contys,args'), union_all (ts1::ts2))
   333       end
   334 and terms_of ts = ListPair.unzip (map term_of ts)
   335 
   336 (*Create a predicate value, again accumulating sort constraints.*)    
   337 fun pred_of (Const("op =", typ), args) =
   338       let val arg_typ = eq_arg_type typ 
   339 	  val (args',ts) = terms_of args
   340 	  val equal_name = make_fixed_const "op ="
   341       in
   342 	  (Predicate(equal_name,[arg_typ],args'),
   343 	   union_all ts)
   344       end
   345   | pred_of (pred,args) = 
   346       let val (pname, (predType,ts1)) = pred_name_type pred
   347 	  val (args',ts2) = terms_of args
   348       in
   349 	  (Predicate(pname,predType,args'), union_all (ts1::ts2))
   350       end;
   351 
   352 (*Treatment of literals, possibly negated or tagged*)
   353 fun predicate_of ((Const("Not",_) $ P), polarity, tag) =
   354       predicate_of (P, not polarity, tag)
   355   | predicate_of ((Const("HOL.tag",_) $ P), polarity, tag) =
   356       predicate_of (P, polarity, true)
   357   | predicate_of (term,polarity,tag) =
   358         (pred_of (strip_comb term), polarity, tag);
   359 
   360 fun literals_of_term1 args (Const("Trueprop",_) $ P) = literals_of_term1 args P
   361   | literals_of_term1 args (Const("op |",_) $ P $ Q) = 
   362       literals_of_term1 (literals_of_term1 args P) Q
   363   | literals_of_term1 (lits, ts) P =
   364       let val ((pred, ts'), polarity, tag) = predicate_of (P,true,false)
   365 	  val lits' = Literal(polarity,pred,tag) :: lits
   366       in
   367 	  (lits', ts union ts')
   368       end;
   369 
   370 val literals_of_term = literals_of_term1 ([],[]);
   371 
   372 
   373 fun list_ord _ ([],[]) = EQUAL
   374   | list_ord _ ([],_) = LESS
   375   | list_ord _ (_,[]) = GREATER
   376   | list_ord ord (x::xs, y::ys) = 
   377       (case ord(x,y) of EQUAL => list_ord ord (xs,ys)
   378 	 	      | xy_ord => xy_ord);
   379 		     
   380 fun type_ord (AtomV(_),AtomV(_)) = EQUAL
   381   | type_ord (AtomV(_),_) = LESS
   382   | type_ord (AtomF(_),AtomV(_)) = GREATER
   383   | type_ord (AtomF(f1),AtomF(f2)) = string_ord (f1,f2)
   384   | type_ord (AtomF(_),_) = LESS
   385   | type_ord (Comp(_,_),AtomV(_)) = GREATER
   386   | type_ord (Comp(_,_),AtomF(_)) = GREATER
   387   | type_ord (Comp(con1,args1),Comp(con2,args2)) = 
   388       (case string_ord(con1,con2) of EQUAL => types_ord (args1,args2)
   389 		      | con_ord => con_ord)
   390 and
   391     types_ord ([],[]) = EQUAL
   392   | types_ord (tps1,tps2) = list_ord type_ord (tps1,tps2);
   393 
   394 
   395 fun term_ord (UVar _, UVar _) = EQUAL
   396   | term_ord (UVar _, _) = LESS
   397   | term_ord (Fun _, UVar _) = GREATER
   398   | term_ord (Fun(f1,tps1,tms1),Fun(f2,tps2,tms2)) = 
   399      (case string_ord (f1,f2) of
   400          EQUAL => 
   401 	   (case terms_ord (tms1,tms2) of EQUAL => types_ord (tps1,tps2)
   402 	      | tms_ord => tms_ord)
   403        | fn_ord => fn_ord)
   404 
   405 and
   406       terms_ord ([],[]) = EQUAL
   407     | terms_ord (tms1,tms2) = list_ord term_ord (tms1,tms2);
   408 
   409 
   410 
   411 fun predicate_ord (Predicate(pname1,ftyps1,ftms1),Predicate(pname2,ftyps2,ftms2)) = 
   412   case string_ord (pname1,pname2) of
   413        EQUAL => (case terms_ord(ftms1,ftms2) of EQUAL => types_ord(ftyps1,ftyps2)
   414 				              | ftms_ord => ftms_ord)
   415      | pname_ord => pname_ord
   416 			   
   417 
   418 fun literal_ord (Literal(false,_,_),Literal(true,_,_)) = LESS
   419   | literal_ord (Literal(true,_,_),Literal(false,_,_)) = GREATER
   420   | literal_ord (Literal(_,pred1,_),Literal(_,pred2,_)) = predicate_ord(pred1,pred2);
   421 
   422 fun sort_lits lits = sort literal_ord lits;
   423 
   424 
   425 (********** clause equivalence ******************)
   426 
   427 fun check_var_pairs (x,y) [] = 0 
   428   | check_var_pairs (x,y) ((u,v)::w) =
   429     if (x,y) = (u,v) then 1 
   430     else
   431 	if x=u orelse y=v then 2 (*conflict*)
   432 	else check_var_pairs (x,y) w;
   433 
   434 fun type_eq (AtomV(v1),AtomV(v2)) (vars,tvars) =
   435     (case check_var_pairs (v1,v2) tvars of 0 => (true,(vars,(v1,v2)::tvars))
   436 					 | 1 => (true,(vars,tvars))
   437 					 | 2 => (false,(vars,tvars)))
   438   | type_eq (AtomV(_),_) vtvars = (false,vtvars)
   439   | type_eq (AtomF(f1),AtomF(f2)) vtvars = (f1=f2,vtvars)
   440   | type_eq (AtomF(_),_) vtvars = (false,vtvars)
   441   | type_eq (Comp(con1,args1),Comp(con2,args2)) vtvars =
   442       let val (eq1,vtvars1) = 
   443 	      if con1 = con2 then types_eq (args1,args2) vtvars
   444 	      else (false,vtvars)
   445       in
   446 	  (eq1,vtvars1)
   447       end
   448   | type_eq (Comp(_,_),_) vtvars = (false,vtvars)
   449 
   450 and types_eq ([],[]) vtvars = (true,vtvars)
   451   | types_eq (tp1::tps1,tp2::tps2) vtvars =
   452       let val (eq1,vtvars1) = type_eq (tp1,tp2) vtvars
   453 	  val (eq2,vtvars2) = if eq1 then types_eq (tps1,tps2) vtvars1
   454 			      else (eq1,vtvars1)
   455       in
   456 	  (eq2,vtvars2)
   457       end;
   458 
   459 
   460 fun term_eq (UVar(v1,tp1),UVar(v2,tp2)) (vars,tvars) =
   461     (case check_var_pairs (v1,v2) vars of 0 => type_eq (tp1,tp2) (((v1,v2)::vars),tvars)
   462 					| 1 => type_eq (tp1,tp2) (vars,tvars)
   463 					| 2 => (false,(vars,tvars)))
   464   | term_eq (UVar _,_) vtvars = (false,vtvars)
   465   | term_eq (Fun(f1,tps1,tms1),Fun(f2,tps2,tms2)) vtvars =
   466       let val (eq1,vtvars1) = 
   467 	      if f1 = f2 then terms_eq (tms1,tms2) vtvars
   468 	      else (false,vtvars)
   469 	  val (eq2,vtvars2) =
   470 	      if eq1 then types_eq (tps1,tps2) vtvars1
   471 	      else (eq1,vtvars1)
   472       in
   473 	  (eq2,vtvars2)
   474       end
   475   | term_eq (Fun(_,_,_),_) vtvars = (false,vtvars)
   476 
   477 and terms_eq ([],[]) vtvars = (true,vtvars)
   478   | terms_eq (tm1::tms1,tm2::tms2) vtvars =
   479       let val (eq1,vtvars1) = term_eq (tm1,tm2) vtvars
   480 	  val (eq2,vtvars2) = if eq1 then terms_eq (tms1,tms2) vtvars1
   481 				     else (eq1,vtvars1)
   482       in
   483 	  (eq2,vtvars2)
   484       end;
   485 					     
   486 
   487 fun pred_eq (Predicate(pname1,tps1,tms1),Predicate(pname2,tps2,tms2)) vtvars =
   488     let val (eq1,vtvars1) = 
   489 	    if pname1 = pname2 then terms_eq (tms1,tms2) vtvars
   490 	    else (false,vtvars)
   491 	val (eq2,vtvars2) = 
   492 	    if eq1 then types_eq (tps1,tps2) vtvars1
   493 	    else (eq1,vtvars1)
   494     in
   495 	(eq2,vtvars2)
   496     end;
   497 					      
   498 
   499 fun lit_eq (Literal(pol1,pred1,_),Literal(pol2,pred2,_)) vtvars =
   500     if (pol1 = pol2) then pred_eq (pred1,pred2) vtvars
   501     else (false,vtvars);
   502 
   503 fun lits_eq ([],[]) vtvars = (true,vtvars)
   504   | lits_eq (l1::ls1,l2::ls2) vtvars = 
   505       let val (eq1,vtvars1) = lit_eq (l1,l2) vtvars
   506       in
   507 	  if eq1 then lits_eq (ls1,ls2) vtvars1
   508 	  else (false,vtvars1)
   509       end
   510   | lits_eq _ vtvars = (false,vtvars);
   511 
   512 (*Equality of two clauses up to variable renaming*)
   513 fun clause_eq (Clause{literals=lits1,...}, Clause{literals=lits2,...}) =
   514   #1 (lits_eq (lits1,lits2) ([],[]));
   515 
   516 
   517 (*** Hash function for clauses ***)
   518 
   519 val xor_words = List.foldl Word.xorb 0w0;
   520 
   521 fun hashw_term (UVar _, w) = w
   522   | hashw_term (Fun(f,tps,args), w) = 
   523       List.foldl hashw_term (Polyhash.hashw_string (f,w)) args;
   524   
   525 fun hashw_pred (Predicate(pn,_,args), w) = 
   526     List.foldl hashw_term (Polyhash.hashw_string (pn,w)) args;
   527     
   528 fun hash1_literal (Literal(true,pred,_)) = hashw_pred (pred, 0w0)
   529   | hash1_literal (Literal(false,pred,_)) = Word.notb (hashw_pred (pred, 0w0));
   530   
   531 fun hash_clause (Clause{literals,...}) =
   532   Word.toIntX (xor_words (map hash1_literal literals));
   533 
   534 
   535 (*Make literals for sorted type variables.  FIXME: can it use map?*) 
   536 fun sorts_on_typs (_, [])   = ([]) 
   537   | sorts_on_typs (v, "HOL.type" :: s) =
   538       sorts_on_typs (v,s)                (*IGNORE sort "type"*)
   539   | sorts_on_typs ((FOLTVar indx), s::ss) =
   540       LTVar(make_type_class s, make_schematic_type_var indx) :: 
   541       sorts_on_typs ((FOLTVar indx), ss)
   542   | sorts_on_typs ((FOLTFree x), s::ss) =
   543       LTFree(make_type_class s, make_fixed_type_var x) :: 
   544       sorts_on_typs ((FOLTFree x), ss);
   545 
   546 
   547 fun pred_of_sort (LTVar (s,ty)) = (s,1)
   548 |   pred_of_sort (LTFree (s,ty)) = (s,1)
   549 
   550 (*Given a list of sorted type variables, return two separate lists.
   551   The first is for TVars, the second for TFrees.*)
   552 fun add_typs_aux [] = ([],[])
   553   | add_typs_aux ((FOLTVar indx,s)::tss) = 
   554       let val vs = sorts_on_typs (FOLTVar indx, s)
   555 	  val (vss,fss) = add_typs_aux tss
   556       in
   557 	  (vs union vss, fss)
   558       end
   559   | add_typs_aux ((FOLTFree x,s)::tss) =
   560       let val fs = sorts_on_typs (FOLTFree x, s)
   561 	  val (vss,fss) = add_typs_aux tss
   562       in
   563 	  (vss, fs union fss)
   564       end;
   565 
   566 
   567 (** make axiom and conjecture clauses. **)
   568 
   569 fun get_tvar_strs [] = []
   570   | get_tvar_strs ((FOLTVar indx,s)::tss) = 
   571       (make_schematic_type_var indx) ins (get_tvar_strs tss)
   572   | get_tvar_strs((FOLTFree x,s)::tss) = get_tvar_strs tss
   573 
   574 (* check if a clause is first-order before making a conjecture clause*)
   575 fun make_conjecture_clause n thm =
   576     let val t = prop_of thm
   577 	val _ = check_is_fol_term t
   578 	    handle TERM("check_is_fol_term",_) => raise CLAUSE("Goal is not FOL",t)
   579 	val (lits,types_sorts) = literals_of_term t
   580     in
   581 	make_clause(n, "conjecture", thm, Conjecture, lits, types_sorts)
   582     end;
   583     
   584 fun make_conjecture_clauses_aux _ [] = []
   585   | make_conjecture_clauses_aux n (t::ts) =
   586       make_conjecture_clause n t :: make_conjecture_clauses_aux (n+1) ts
   587 
   588 val make_conjecture_clauses = make_conjecture_clauses_aux 0
   589 
   590 (** Too general means, positive equality literal with a variable X as one operand,
   591   when X does not occur properly in the other operand. This rules out clearly
   592   inconsistent clauses such as V=a|V=b, though it by no means guarantees soundness. **)
   593 
   594 fun occurs a (UVar(b,_)) = a=b
   595   | occurs a (Fun (_,_,ts)) = exists (occurs a) ts
   596 
   597 (*Is the first operand a variable that does not properly occur in the second operand?*)
   598 fun too_general_terms (UVar _, UVar _) = false
   599   | too_general_terms (Fun _, _) = false
   600   | too_general_terms (UVar (a,_), t) = not (occurs a t);
   601 
   602 fun too_general_lit (Literal (true,Predicate("equal",_,[x,y]),_)) =
   603       too_general_terms (x,y) orelse too_general_terms(y,x)
   604   | too_general_lit _ = false;
   605 
   606 (*before converting an axiom clause to "clause" format, check if it is FOL*)
   607 fun make_axiom_clause thm (ax_name,cls_id) =
   608     let val term = prop_of thm
   609 	val (lits,types_sorts) = literals_of_term term
   610     in 
   611 	if not (Meson.is_fol_term term) then
   612 	   (Output.debug ("Omitting " ^ ax_name ^ ": Axiom is not FOL"); 
   613 	    NONE)
   614 	else if forall too_general_lit lits then
   615 	   (Output.debug ("Omitting " ^ ax_name ^ ": equalities are too general"); 
   616 	    NONE)
   617 	else SOME (make_clause(cls_id, ax_name, thm, Axiom, sort_lits lits, types_sorts))
   618     end
   619     handle CLAUSE _ => NONE;
   620 
   621 
   622 fun make_axiom_clauses [] = []
   623   | make_axiom_clauses ((thm,(name,id))::thms) =
   624     case make_axiom_clause thm (name,id) of SOME cls => if isTaut cls then make_axiom_clauses thms else cls :: make_axiom_clauses thms
   625 						    | NONE => make_axiom_clauses thms;
   626 
   627 (**** Isabelle arities ****)
   628 
   629 exception ARCLAUSE of string;
   630  
   631 type class = string; 
   632 type tcons = string; 
   633 
   634 datatype arLit = TConsLit of bool * (class * tcons * string list)
   635                | TVarLit of bool * (class * string);
   636  
   637 datatype arityClause =  
   638 	 ArityClause of {clause_id: clause_id,
   639 	  	         axiom_name: axiom_name,
   640 			 kind: kind,
   641 			 conclLit: arLit,
   642 			 premLits: arLit list};
   643 
   644 
   645 fun gen_TVars 0 = []
   646   | gen_TVars n = ("T_" ^ Int.toString n) :: gen_TVars (n-1);
   647 
   648 fun pack_sort(_,[])  = []
   649   | pack_sort(tvar, "HOL.type"::srt) = pack_sort(tvar, srt)   (*IGNORE sort "type"*)
   650   | pack_sort(tvar, cls::srt) =  (make_type_class cls, tvar) :: pack_sort(tvar, srt);
   651     
   652 fun make_TVarLit (b, (cls,str)) = TVarLit(b, (cls,str));
   653 fun make_TConsLit (b, (cls,tcons,tvars)) = 
   654       TConsLit(b, (make_type_class cls, make_fixed_type_const tcons, tvars));
   655 
   656 (*Arity of type constructor tcon :: (arg1,...,argN)res*)
   657 fun make_axiom_arity_clause (tcons, n, (res,args)) =
   658    let val nargs = length args
   659        val tvars = gen_TVars nargs
   660        val tvars_srts = ListPair.zip (tvars,args)
   661        val tvars_srts' = union_all(map pack_sort tvars_srts)
   662        val false_tvars_srts' = map (pair false) tvars_srts'
   663    in
   664       ArityClause {clause_id = n, kind = Axiom, 
   665                    axiom_name = lookup_type_const tcons,
   666                    conclLit = make_TConsLit(true, (res,tcons,tvars)), 
   667                    premLits = map make_TVarLit false_tvars_srts'}
   668    end;
   669 
   670 
   671 (**** Isabelle class relations ****)
   672 
   673 datatype classrelClause = 
   674 	 ClassrelClause of {axiom_name: axiom_name,
   675 			    subclass: class,
   676 			    superclass: class};
   677 
   678 fun make_axiom_classrelClause n subclass superclass =
   679   ClassrelClause {axiom_name = clrelclause_prefix ^ ascii_of subclass ^ 
   680                                 "_" ^ Int.toString n,
   681                   subclass = make_type_class subclass, 
   682                   superclass = make_type_class superclass};
   683 
   684 fun classrelClauses_of_aux n sub [] = []
   685   | classrelClauses_of_aux n sub ("HOL.type"::sups) = (*Should be ignored*)
   686       classrelClauses_of_aux n sub sups
   687   | classrelClauses_of_aux n sub (sup::sups) =
   688       make_axiom_classrelClause n sub sup :: classrelClauses_of_aux (n+1) sub sups;
   689 
   690 fun classrelClauses_of (sub,sups) = classrelClauses_of_aux 0 sub sups;
   691 
   692 val classrel_of = #2 o #classes o Type.rep_tsig o Sign.tsig_of;
   693 
   694 fun classrel_clauses_classrel (C: Sorts.classes) =
   695   map classrelClauses_of (Graph.dest C);
   696 
   697 val classrel_clauses_thy = List.concat o classrel_clauses_classrel o classrel_of;
   698 
   699 
   700 (** Isabelle arities **)
   701 
   702 fun arity_clause _ (tcons, []) = []
   703   | arity_clause n (tcons, ("HOL.type",_)::ars) =  (*ignore*)
   704       arity_clause n (tcons,ars)
   705   | arity_clause n (tcons, ar::ars) =
   706       make_axiom_arity_clause (tcons,n,ar) :: 
   707       arity_clause (n+1) (tcons,ars);
   708 
   709 fun multi_arity_clause [] = []
   710   | multi_arity_clause ((tcons,ars) :: tc_arlists) =
   711       (*Reversal ensures that older entries always get the same axiom name*)
   712       arity_clause 0 (tcons, rev ars)  @  
   713       multi_arity_clause tc_arlists 
   714 
   715 fun arity_clause_thy thy =
   716   let val arities = #arities (Type.rep_tsig (Sign.tsig_of thy))
   717   in multi_arity_clause (rev (Symtab.dest arities)) end;
   718 
   719 
   720 (**** Find occurrences of predicates in clauses ****)
   721 
   722 (*FIXME: multiple-arity checking doesn't work, as update_new is the wrong 
   723   function (it flags repeated declarations of a function, even with the same arity)*)
   724 
   725 fun update_many (tab, keypairs) = foldl (uncurry Symtab.update) tab keypairs;
   726 
   727 fun add_predicate_preds (Predicate(pname,tys,tms), preds) = 
   728   if pname = "equal" then preds (*equality is built-in and requires no declaration*)
   729   else Symtab.update (pname, length tys + length tms) preds
   730 
   731 fun add_literal_preds (Literal(_,pred,_), preds) = add_predicate_preds (pred,preds)
   732 
   733 fun add_type_sort_preds ((FOLTVar indx,s), preds) = 
   734       update_many (preds, map pred_of_sort (sorts_on_typs (FOLTVar indx, s)))
   735   | add_type_sort_preds ((FOLTFree x,s), preds) =
   736       update_many (preds, map pred_of_sort (sorts_on_typs (FOLTFree x, s)));
   737 
   738 fun add_clause_preds (Clause {literals, types_sorts, ...}, preds) =
   739   foldl add_literal_preds (foldl add_type_sort_preds preds types_sorts) literals
   740   handle Symtab.DUP a => raise ERROR ("predicate " ^ a ^ " has multiple arities")
   741 
   742 fun add_classrelClause_preds (ClassrelClause {subclass,superclass,...}, preds) =
   743   Symtab.update (subclass,1) (Symtab.update (superclass,1) preds);
   744 
   745 fun add_arityClause_preds (ArityClause {conclLit,...}, preds) =
   746   let val TConsLit(_, (tclass, _, _)) = conclLit
   747   in  Symtab.update (tclass,1) preds  end;
   748 
   749 fun preds_of_clauses clauses clsrel_clauses arity_clauses = 
   750   Symtab.dest
   751     (foldl add_classrelClause_preds 
   752       (foldl add_arityClause_preds
   753         (foldl add_clause_preds Symtab.empty clauses)
   754         arity_clauses)
   755       clsrel_clauses)
   756 
   757 (*** Find occurrences of functions in clauses ***)
   758 
   759 fun add_foltype_funcs (AtomV _, funcs) = funcs
   760   | add_foltype_funcs (AtomF a, funcs) = Symtab.update (a,0) funcs
   761   | add_foltype_funcs (Comp(a,tys), funcs) = 
   762       foldl add_foltype_funcs (Symtab.update (a, length tys) funcs) tys;
   763 
   764 fun add_folterm_funcs (UVar _, funcs) = funcs
   765   | add_folterm_funcs (Fun(a,tys,[]), funcs) = Symtab.update (a,0) funcs
   766       (*A constant is a special case: it has no type argument even if overloaded*)
   767   | add_folterm_funcs (Fun(a,tys,tms), funcs) = 
   768       foldl add_foltype_funcs 
   769 	    (foldl add_folterm_funcs (Symtab.update (a, length tys + length tms) funcs) 
   770 	           tms) 
   771 	    tys
   772 
   773 fun add_predicate_funcs (Predicate(_,tys,tms), funcs) = 
   774     foldl add_foltype_funcs (foldl add_folterm_funcs funcs tms) tys;
   775 
   776 fun add_literal_funcs (Literal(_,pred,_), funcs) = add_predicate_funcs (pred,funcs)
   777 
   778 fun add_arityClause_funcs (ArityClause {conclLit,...}, funcs) =
   779   let val TConsLit(_, (_, tcons, tvars)) = conclLit
   780   in  Symtab.update (tcons, length tvars) funcs  end;
   781 
   782 fun add_clause_funcs (Clause {literals, ...}, funcs) =
   783   foldl add_literal_funcs funcs literals
   784   handle Symtab.DUP a => raise ERROR ("function " ^ a ^ " has multiple arities")
   785 
   786 fun funcs_of_clauses clauses arity_clauses = 
   787   Symtab.dest (foldl add_arityClause_funcs 
   788                      (foldl add_clause_funcs Symtab.empty clauses)
   789                      arity_clauses)
   790 
   791 
   792 (**** String-oriented operations ****)
   793 
   794 fun wrap_eq_type typ t = eq_typ_wrapper ^"(" ^ t ^ "," ^ typ ^ ")";
   795 
   796 (*Only need to wrap equality's arguments with "typeinfo" if the output clauses are typed 
   797  and if we specifically ask for types to be included.   *)
   798 fun string_of_equality (typ,terms) =
   799       let val [tstr1,tstr2] = map string_of_term terms
   800 	  val typ' = string_of_fol_type typ
   801       in
   802 	  if !keep_types andalso !special_equal 
   803 	  then "equal(" ^ (wrap_eq_type typ' tstr1) ^ "," ^ 
   804 		 	  (wrap_eq_type typ' tstr2) ^ ")"
   805 	  else "equal(" ^ tstr1 ^ "," ^ tstr2 ^ ")"
   806       end
   807 and string_of_term (UVar(x,_)) = x
   808   | string_of_term (Fun("equal",[typ],terms)) = string_of_equality(typ,terms)
   809   | string_of_term (Fun (name,typs,[])) = name (*Overloaded consts like 0 don't get types!*)
   810   | string_of_term (Fun (name,typs,terms)) = 
   811       let val terms_as_strings = map string_of_term terms
   812 	  val typs' = if !keep_types then map string_of_fol_type typs else []
   813       in  name ^ (paren_pack (terms_as_strings @ typs'))  end;
   814 
   815 (* before output the string of the predicate, check if the predicate corresponds to an equality or not. *)
   816 fun string_of_predicate (Predicate("equal",[typ],terms)) = string_of_equality(typ,terms)
   817   | string_of_predicate (Predicate(name,typs,terms)) = 
   818       let val terms_as_strings = map string_of_term terms
   819 	  val typs' = if !keep_types then map string_of_fol_type typs else []
   820       in  name ^ (paren_pack (terms_as_strings @ typs'))  end;
   821 
   822 fun string_of_clausename (cls_id,ax_name) = 
   823     clause_prefix ^ ascii_of ax_name ^ "_" ^ Int.toString cls_id;
   824 
   825 fun string_of_type_clsname (cls_id,ax_name,idx) = 
   826     string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx);
   827 
   828 (*Write a list of strings to a file*)
   829 fun writeln_strs os = List.app (fn s => TextIO.output (os,s));
   830 
   831     
   832 (**** Producing DFG files ****)
   833 
   834 (*Attach sign in DFG syntax: false means negate.*)
   835 fun dfg_sign true s = s
   836   | dfg_sign false s = "not(" ^ s ^ ")"  
   837 
   838 fun dfg_literal (Literal(pol,pred,tag)) = dfg_sign pol (string_of_predicate pred)
   839 
   840 fun dfg_of_typeLit (LTVar (s,ty)) = "not(" ^ s ^ "(" ^ ty ^ "))"
   841   | dfg_of_typeLit (LTFree (s,ty)) = s ^ "(" ^ ty ^ ")";
   842  
   843 (*Enclose the clause body by quantifiers, if necessary*)
   844 fun dfg_forall [] body = body  
   845   | dfg_forall vars body = "forall([" ^ commas vars ^ "],\n" ^ body ^ ")"
   846 
   847 fun gen_dfg_cls (cls_id, ax_name, knd, lits, vars) = 
   848     "clause( %(" ^ knd ^ ")\n" ^ 
   849     dfg_forall vars ("or(" ^ lits ^ ")") ^ ",\n" ^ 
   850     string_of_clausename (cls_id,ax_name) ^  ").\n\n";
   851 
   852 fun dfg_clause_aux (Clause{literals, types_sorts, ...}) = 
   853   let val lits = map dfg_literal literals
   854       val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
   855       val tvar_lits_strs = 
   856 	  if !keep_types then map dfg_of_typeLit tvar_lits else []
   857       val tfree_lits =
   858           if !keep_types then map dfg_of_typeLit tfree_lits else []
   859   in
   860       (tvar_lits_strs @ lits, tfree_lits)
   861   end; 
   862 
   863 fun dfg_folterms (Literal(pol,pred,tag)) = 
   864   let val Predicate (_, _, folterms) = pred
   865   in  folterms  end
   866 
   867 fun get_uvars (UVar(a,typ)) = [a] 
   868   | get_uvars (Fun (_,typ,tlist)) = union_all(map get_uvars tlist)
   869 
   870 fun dfg_vars (Clause {literals,...}) =
   871   union_all (map get_uvars (List.concat (map dfg_folterms literals)))
   872 
   873 fun clause2dfg (cls as Clause{axiom_name,clause_id,kind,types_sorts,...}) =
   874     let val (lits,tfree_lits) = dfg_clause_aux cls 
   875             (*"lits" includes the typing assumptions (TVars)*)
   876         val vars = dfg_vars cls
   877         val tvars = get_tvar_strs types_sorts
   878 	val knd = name_of_kind kind
   879 	val lits_str = commas lits
   880 	val cls_str = gen_dfg_cls(clause_id, axiom_name, knd, lits_str, tvars@vars) 
   881     in (cls_str, tfree_lits) end;
   882 
   883 fun string_of_arity (name, num) =  "(" ^ name ^ "," ^ Int.toString num ^ ")"
   884 
   885 fun string_of_preds [] = ""
   886   | string_of_preds preds = "predicates[" ^ commas(map string_of_arity preds) ^ "].\n";
   887 
   888 fun string_of_funcs [] = ""
   889   | string_of_funcs funcs = "functions[" ^ commas(map string_of_arity funcs) ^ "].\n" ;
   890 
   891 fun string_of_symbols predstr funcstr = 
   892   "list_of_symbols.\n" ^ predstr  ^ funcstr  ^ "end_of_list.\n\n";
   893 
   894 fun string_of_start name = "begin_problem(" ^ name ^ ").\n\n";
   895 
   896 fun string_of_descrip name = 
   897   "list_of_descriptions.\nname({*" ^ name ^ 
   898   "*}).\nauthor({*Isabelle*}).\nstatus(unknown).\ndescription({*auto-generated*}).\nend_of_list.\n\n"
   899 
   900 fun dfg_tfree_clause tfree_lit =
   901   "clause( %(conjecture)\n" ^ "or( " ^ tfree_lit ^ "),\n" ^ "tfree_tcs" ^ ").\n\n"
   902 
   903 fun string_of_arClauseID (ArityClause {clause_id,axiom_name,...}) =
   904     arclause_prefix ^ ascii_of axiom_name ^ "_" ^ Int.toString clause_id;
   905 
   906 fun dfg_of_arLit (TConsLit(pol,(c,t,args))) =
   907       dfg_sign pol (c ^ "(" ^ t ^ paren_pack args ^ ")")
   908   | dfg_of_arLit (TVarLit(pol,(c,str))) =
   909       dfg_sign pol (c ^ "(" ^ str ^ ")")
   910     
   911 fun dfg_classrelLits sub sup = 
   912     let val tvar = "(T)"
   913     in 
   914 	"not(" ^ sub ^ tvar ^ "), " ^ sup ^ tvar
   915     end;
   916 
   917 fun dfg_classrelClause (ClassrelClause {axiom_name,subclass,superclass,...}) =
   918   "clause(forall([T],\nor( " ^ dfg_classrelLits subclass superclass ^ ")),\n" ^
   919   axiom_name ^ ").\n\n";
   920       
   921 fun dfg_arity_clause (arcls as ArityClause{kind,conclLit,premLits,...}) = 
   922   let val arcls_id = string_of_arClauseID arcls
   923       val knd = name_of_kind kind
   924       val TConsLit(_, (_,_,tvars)) = conclLit
   925       val lits = map dfg_of_arLit (conclLit :: premLits)
   926   in
   927       "clause( %(" ^ knd ^ ")\n" ^ 
   928       dfg_forall tvars ("or( " ^ commas lits ^ ")") ^ ",\n" ^
   929       arcls_id ^ ").\n\n"
   930   end;
   931 
   932 (* write out a subgoal in DFG format to the file "xxxx_N"*)
   933 fun dfg_write_file ths filename (axclauses,classrel_clauses,arity_clauses) = 
   934   let 
   935     val _ = Output.debug ("Preparing to write the DFG file " ^ filename)
   936     val conjectures = make_conjecture_clauses ths
   937     val (dfg_clss, tfree_litss) = ListPair.unzip (map clause2dfg conjectures)
   938     val clss = conjectures @ axclauses
   939     val funcs = funcs_of_clauses clss arity_clauses
   940     and preds = preds_of_clauses clss classrel_clauses arity_clauses
   941     and probname = Path.pack (Path.base (Path.unpack filename))
   942     val (axstrs, _) = ListPair.unzip (map clause2dfg axclauses)
   943     val tfree_clss = map dfg_tfree_clause (union_all tfree_litss) 
   944     val out = TextIO.openOut filename
   945   in
   946     TextIO.output (out, string_of_start probname); 
   947     TextIO.output (out, string_of_descrip probname); 
   948     TextIO.output (out, string_of_symbols (string_of_funcs funcs) (string_of_preds preds)); 
   949     TextIO.output (out, "list_of_clauses(axioms,cnf).\n");
   950     writeln_strs out axstrs;
   951     List.app (curry TextIO.output out o dfg_classrelClause) classrel_clauses;
   952     List.app (curry TextIO.output out o dfg_arity_clause) arity_clauses;
   953     TextIO.output (out, "end_of_list.\n\nlist_of_clauses(conjectures,cnf).\n");
   954     writeln_strs out tfree_clss;
   955     writeln_strs out dfg_clss;
   956     TextIO.output (out, "end_of_list.\n\nend_problem.\n");
   957     TextIO.closeOut out
   958   end;
   959 
   960 
   961 (**** Produce TPTP files ****)
   962 
   963 (*Attach sign in TPTP syntax: false means negate.*)
   964 fun tptp_sign true s = "++" ^ s
   965   | tptp_sign false s = "--" ^ s
   966 
   967 fun tptp_literal (Literal(pol,pred,tag)) =  (*FIXME REMOVE TAGGING*)
   968     let val pred_string = string_of_predicate pred
   969 	val tagged_pol = 
   970 	      if (tag andalso !tagged) then (if pol then "+++" else "---")
   971 	      else (if pol then "++" else "--")
   972      in
   973 	tagged_pol ^ pred_string
   974     end;
   975 
   976 fun tptp_of_typeLit (LTVar (s,ty)) = "--" ^ s ^ "(" ^ ty ^ ")"
   977   | tptp_of_typeLit (LTFree (s,ty)) = "++" ^ s ^ "(" ^ ty ^ ")";
   978  
   979 fun gen_tptp_cls (cls_id,ax_name,knd,lits) = 
   980     "input_clause(" ^ string_of_clausename (cls_id,ax_name) ^ "," ^ 
   981     knd ^ "," ^ lits ^ ").\n";
   982 
   983 fun gen_tptp_type_cls (cls_id,ax_name,knd,tfree_lit,idx) = 
   984     "input_clause(" ^ string_of_type_clsname (cls_id,ax_name,idx) ^ "," ^ 
   985     knd ^ ",[" ^ tfree_lit ^ "]).\n";
   986 
   987 fun tptp_type_lits (Clause {literals, types_sorts, ...}) = 
   988     let val lits = map tptp_literal literals
   989 	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
   990         val tvar_lits_strs =
   991             if !keep_types then map tptp_of_typeLit tvar_lits else []
   992 	val tfree_lits =
   993 	    if !keep_types then map tptp_of_typeLit tfree_lits else []
   994     in
   995 	(tvar_lits_strs @ lits, tfree_lits)
   996     end; 
   997 
   998 fun clause2tptp (cls as Clause {clause_id, axiom_name, kind, ...}) =
   999     let val (lits,tfree_lits) = tptp_type_lits cls 
  1000             (*"lits" includes the typing assumptions (TVars)*)
  1001 	val knd = name_of_kind kind
  1002 	val cls_str = gen_tptp_cls(clause_id, axiom_name, knd, bracket_pack lits) 
  1003     in
  1004 	(cls_str,tfree_lits) 
  1005     end;
  1006 
  1007 fun tptp_tfree_clause tfree_lit =
  1008     "input_clause(" ^ "tfree_tcs," ^ "conjecture" ^ ",[" ^ tfree_lit ^ "]).\n";
  1009 
  1010 fun tptp_of_arLit (TConsLit(b,(c,t,args))) =
  1011       tptp_sign b (c ^ "(" ^ t ^ paren_pack args ^ ")")
  1012   | tptp_of_arLit (TVarLit(b,(c,str))) =
  1013       tptp_sign b (c ^ "(" ^ str ^ ")")
  1014     
  1015 fun tptp_arity_clause (arcls as ArityClause{kind,conclLit,premLits,...}) = 
  1016   let val arcls_id = string_of_arClauseID arcls
  1017       val knd = name_of_kind kind
  1018       val lits = map tptp_of_arLit (conclLit :: premLits)
  1019   in
  1020     "input_clause(" ^ arcls_id ^ "," ^ knd ^ "," ^ bracket_pack lits ^ ").\n"
  1021   end;
  1022 
  1023 fun tptp_classrelLits sub sup = 
  1024     let val tvar = "(T)"
  1025     in 
  1026 	"[--" ^ sub ^ tvar ^ ",++" ^ sup ^ tvar ^ "]"
  1027     end;
  1028 
  1029 fun tptp_classrelClause (ClassrelClause {axiom_name,subclass,superclass,...}) =
  1030   "input_clause(" ^ axiom_name ^ ",axiom," ^ tptp_classrelLits subclass superclass ^ ").\n" 
  1031 
  1032 (* write out a subgoal as tptp clauses to the file "xxxx_N"*)
  1033 fun tptp_write_file thms filename (axclauses,classrel_clauses,arity_clauses) =
  1034   let
  1035     val _ = Output.debug ("Preparing to write the TPTP file " ^ filename)
  1036     val clss = make_conjecture_clauses thms
  1037     val axclauses' = make_axiom_clauses axclauses
  1038     val (tptp_clss,tfree_litss) = ListPair.unzip (map clause2tptp clss)
  1039     val tfree_clss = map tptp_tfree_clause (foldl (op union_string) [] tfree_litss)
  1040     val out = TextIO.openOut filename
  1041   in
  1042     List.app (curry TextIO.output out o #1 o clause2tptp) axclauses';
  1043     writeln_strs out tfree_clss;
  1044     writeln_strs out tptp_clss;
  1045     List.app (curry TextIO.output out o tptp_classrelClause) classrel_clauses;
  1046     List.app (curry TextIO.output out o tptp_arity_clause) arity_clauses;
  1047     TextIO.closeOut out
  1048   end;
  1049 
  1050 
  1051 
  1052 
  1053 
  1054 end;