src/HOL/Tools/res_clause.ML
author paulson
Tue, 22 Mar 2005 16:30:43 +0100
changeset 15615 d72b1867d09d
parent 15610 f855fd163b62
child 15774 9df37a0e935d
permissions -rw-r--r--
ensuring that "equal" is not a function
     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 ARCLAUSE of string
    12     exception CLAUSE of string
    13     type arityClause 
    14     type classrelClause
    15     val classrelClauses_of :
    16        string * string list -> classrelClause list
    17     type clause
    18     val keep_types : bool ref
    19     val make_axiom_arity_clause :
    20        string * (string * string list list) -> arityClause
    21     val make_axiom_classrelClause :
    22        string * string option -> classrelClause
    23     val make_axiom_clause : Term.term -> string * int -> clause
    24     val make_axiom_clause_thm : Thm.thm -> string * int -> clause
    25     val make_conjecture_clause : Term.term -> clause
    26     val make_conjecture_clause_thm : Thm.thm -> clause
    27     val make_hypothesis_clause : Term.term -> clause
    28     val make_hypothesis_clause_thm : Thm.thm -> clause
    29     val special_equal : bool ref
    30     val tptp_arity_clause : arityClause -> string
    31     val tptp_classrelClause : classrelClause -> string
    32     val tptp_clause : clause -> string list
    33     val tptp_clauses2str : string list -> string
    34     val typed : unit -> unit
    35     val untyped : unit -> unit
    36     val clause2tptp : clause -> string * string list
    37     val tfree_clause : string -> string
    38   end;
    39 
    40 structure ResClause : RES_CLAUSE =
    41 struct
    42 
    43 (* Added for typed equality *)
    44 val special_equal = ref false; (* by default,equality does not carry type information *)
    45 val eq_typ_wrapper = "typeinfo"; (* default string *)
    46 
    47 
    48 val schematic_var_prefix = "V_";
    49 val fixed_var_prefix = "v_";
    50 
    51 
    52 val tvar_prefix = "Typ_";
    53 val tfree_prefix = "typ_";
    54 
    55 
    56 val clause_prefix = "cls_"; 
    57 
    58 val arclause_prefix = "arcls_" 
    59 
    60 val const_prefix = "const_";
    61 val tconst_prefix = "tconst_"; 
    62 
    63 val class_prefix = "clas_"; 
    64 
    65 
    66 
    67 (**** some useful functions ****)
    68  
    69 val const_trans_table =
    70       Symtab.make [("op =", "equal"),
    71 	  	   ("op <=", "lessequals"),
    72 		   ("op <", "less"),
    73 		   ("op &", "and"),
    74 		   ("op |", "or"),
    75 		   ("op -->", "implies"),
    76 		   ("op :", "in"),
    77 		   ("op Un", "union"),
    78 		   ("op Int", "inter")];
    79 
    80 
    81 
    82 (*Escaping of special characters.
    83   Alphanumeric characters are left unchanged.
    84   The character _ goes to __
    85   Characters in the range ASCII space to / go to _A to _P, respectively.
    86   Other printing characters go to _NNN where NNN is the decimal ASCII code.*)
    87 local
    88 
    89 val A_minus_space = Char.ord #"A" - Char.ord #" ";
    90 
    91 fun ascii_of_c c =
    92   if Char.isAlphaNum c then String.str c
    93   else if c = #"_" then "__"
    94   else if #" " <= c andalso c <= #"/" 
    95        then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
    96   else if Char.isPrint c then ("_" ^ Int.toString (Char.ord c))
    97   else ""
    98 
    99 in
   100 
   101 val ascii_of = String.translate ascii_of_c;
   102 
   103 end;
   104 
   105 
   106 (* another version of above functions that remove chars that may not be allowed by Vampire *)
   107 fun make_schematic_var v = schematic_var_prefix ^ (ascii_of v);
   108 fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x);
   109 
   110 fun make_schematic_type_var v = tvar_prefix ^ (ascii_of v);
   111 fun make_fixed_type_var x = tfree_prefix ^ (ascii_of x);
   112 
   113 fun make_fixed_const c = const_prefix ^ (ascii_of c);
   114 fun make_fixed_type_const c = tconst_prefix ^ (ascii_of c);
   115 
   116 fun make_type_class clas = class_prefix ^ (ascii_of clas);
   117 
   118 
   119 
   120 
   121 fun lookup_const c =
   122     case Symtab.lookup (const_trans_table,c) of
   123         SOME c' => c'
   124       | NONE =>  make_fixed_const c;
   125 
   126 
   127 
   128 (***** definitions and functions for FOL clauses, prepared for conversion into TPTP format or SPASS format. *****)
   129 
   130 val keep_types = ref true; (* default is true *)
   131 fun untyped () = (keep_types := false);
   132 fun typed () = (keep_types := true);
   133 
   134 
   135 datatype kind = Axiom | Hypothesis | Conjecture;
   136 fun name_of_kind Axiom = "axiom"
   137   | name_of_kind Hypothesis = "hypothesis"
   138   | name_of_kind Conjecture = "conjecture";
   139 
   140 type clause_id = int;
   141 type axiom_name = string;
   142 
   143 
   144 type polarity = bool;
   145 
   146 type indexname = Term.indexname;
   147 
   148 
   149 (* "tag" is used for vampire specific syntax  *)
   150 type tag = bool; 
   151 
   152 
   153 
   154 fun string_of_indexname (name,index) = name ^ "_" ^ (string_of_int index);
   155 
   156 
   157 val id_ref = ref 0;
   158 fun generate_id () = 
   159      let val id = !id_ref
   160      in
   161 	 (id_ref:=id + 1; id)
   162      end;
   163 
   164 
   165 
   166 (**** Isabelle FOL clauses ****)
   167 
   168 (* by default it is false *)
   169 val tagged = ref false;
   170 
   171 type pred_name = string;
   172 type sort = Term.sort;
   173 type fol_type = string;
   174 
   175 
   176 datatype type_literal = LTVar of string | LTFree of string;
   177 
   178 
   179 datatype folTerm = UVar of string * fol_type| Fun of string * fol_type * folTerm list;
   180 datatype predicate = Predicate of pred_name * fol_type * folTerm list;
   181 
   182 
   183 datatype literal = Literal of polarity * predicate * tag;
   184 
   185 
   186 datatype typ_var = FOLTVar of indexname | FOLTFree of string;
   187 
   188 
   189 (* ML datatype used to repsent one single clause: disjunction of literals. *)
   190 datatype clause = 
   191 	 Clause of {clause_id: clause_id,
   192 		    axiom_name: axiom_name,
   193 		    kind: kind,
   194 		    literals: literal list,
   195 		    types_sorts: (typ_var * sort) list, 
   196                     tvar_type_literals: type_literal list, 
   197                     tfree_type_literals: type_literal list };
   198 
   199 
   200 exception CLAUSE of string;
   201 
   202 
   203 
   204 (*** make clauses ***)
   205 
   206 
   207 fun make_clause (clause_id,axiom_name,kind,literals,types_sorts,tvar_type_literals,tfree_type_literals) =
   208      Clause {clause_id = clause_id, axiom_name = axiom_name, kind = kind, literals = literals, types_sorts = types_sorts,tvar_type_literals = tvar_type_literals,tfree_type_literals = tfree_type_literals};
   209 
   210 
   211 
   212 fun type_of (Type (a, [])) = (make_fixed_type_const a,[]) 
   213   | type_of (Type (a, Ts)) = 
   214     let val foltyps_ts = map type_of Ts
   215         val (folTyps,ts) = ResLib.unzip foltyps_ts
   216         val ts' = ResLib.flat_noDup ts
   217     in    
   218         (((make_fixed_type_const a) ^ (ResLib.list_to_string folTyps)),ts') 
   219     end
   220   | type_of (TFree (a, s)) = (make_fixed_type_var a, [((FOLTFree a),s)])
   221   | type_of (TVar (v, s))  = (make_schematic_type_var (string_of_indexname v), [((FOLTVar v),s)]);
   222 
   223 (* added: checkMeta: string -> bool *)
   224 (* Any meta vars like ?x should be treated as universal vars,although it is represented as "Free(...)" by Isabelle *)
   225 fun checkMeta s =
   226     let val chars = explode s
   227     in
   228 	["M", "E", "T", "A", "H", "Y", "P", "1"] prefix chars
   229     end;
   230 
   231     
   232 fun pred_name_type (Const(c,T)) = (lookup_const c,type_of T)
   233   | pred_name_type (Free(x,T))  = 
   234     let val is_meta = checkMeta x
   235     in
   236 	if is_meta then (raise CLAUSE("Predicate Not First Order")) else
   237 	(make_fixed_var x,type_of T)
   238     end
   239   | pred_name_type (Var(_,_))   = raise CLAUSE("Predicate Not First Order")
   240   | pred_name_type _          = raise CLAUSE("Predicate input unexpected");
   241 
   242 
   243 (* For type equality *)
   244 (* here "arg_typ" is the type of "="'s argument's type, not the type of the equality *)
   245 (* Find type of equality arg *)
   246 fun eq_arg_type (Type("fun",[T,_])) = 
   247     let val (folT,_) = type_of T;
   248     in
   249 	folT
   250     end;
   251 
   252 
   253 
   254 fun fun_name_type (Const(c,T)) = (lookup_const c,type_of T)
   255   | fun_name_type (Free(x,T)) = (make_fixed_var x,type_of T)
   256   | fun_name_type _ = raise CLAUSE("Function Not First Order");
   257     
   258 
   259 
   260 fun term_of (Var(ind_nm,T)) = 
   261     let val (folType,ts) = type_of T
   262     in
   263 	(UVar(make_schematic_var(string_of_indexname ind_nm),folType),ts)
   264     end
   265   | term_of (Free(x,T)) = 
   266     let val is_meta = checkMeta x
   267 	val (folType,ts) = type_of T
   268     in
   269 	if is_meta then (UVar(make_schematic_var x,folType),ts)
   270 	else
   271             (Fun(make_fixed_var x,folType,[]),ts)
   272     end
   273   | term_of (Const(c,T)) =  (* impossible to be equality *)
   274     let val (folType,ts) = type_of T
   275      in
   276         (Fun(lookup_const c,folType,[]),ts)
   277     end    
   278   | term_of (app as (t $ a)) = 
   279     let val (f,args) = strip_comb app
   280         fun term_of_aux () = 
   281             let val (funName,(funType,ts1)) = fun_name_type f
   282                  val (args',ts2) = ResLib.unzip (map term_of args)
   283                  val ts3 = ResLib.flat_noDup (ts1::ts2)
   284             in
   285 		(Fun(funName,funType,args'),ts3)
   286             end
   287 	fun term_of_eq ((Const ("op =", typ)),args) =
   288 	    let val arg_typ = eq_arg_type typ
   289 		val (args',ts) = ResLib.unzip (map term_of args)
   290 		val equal_name = lookup_const ("op =")
   291 	    in
   292 		(Fun(equal_name,arg_typ,args'),ResLib.flat_noDup ts)
   293 	    end
   294     in
   295         case f of Const ("op =", typ) => term_of_eq (f,args)
   296 	        | Const(_,_) => term_of_aux ()
   297                 | Free(s,_)  => if (checkMeta s) then (raise CLAUSE("Function Not First Order")) else term_of_aux ()
   298                 | _          => raise CLAUSE("Function Not First Order")
   299     end
   300   | term_of _ = raise CLAUSE("Function Not First Order"); 
   301 
   302 
   303 
   304 
   305 fun pred_of_eq ((Const ("op =", typ)),args) =
   306     let val arg_typ = eq_arg_type typ 
   307 	val (args',ts) = ResLib.unzip (map term_of args)
   308 	val equal_name = lookup_const "op ="
   309     in
   310 	(Predicate(equal_name,arg_typ,args'),ResLib.flat_noDup ts)
   311     end;
   312 
   313 
   314 (* changed for non-equality predicate *)
   315 (* The input "pred" cannot be an equality *)
   316 fun pred_of_nonEq (pred,args) = 
   317     let val (predName,(predType,ts1)) = pred_name_type pred
   318 	val (args',ts2) = ResLib.unzip (map term_of args)
   319 	val ts3 = ResLib.flat_noDup (ts1::ts2)
   320     in
   321 	(Predicate(predName,predType,args'),ts3)
   322     end;
   323 
   324 
   325 (* Changed for typed equality *)
   326 (* First check if the predicate is an equality or not, then call different functions for equality and non-equalities *)
   327 fun predicate_of term =
   328     let val (pred,args) = strip_comb term
   329     in
   330 	case pred of (Const ("op =", _)) => pred_of_eq (pred,args)
   331 		   | _ => pred_of_nonEq (pred,args)
   332     end;
   333 
   334 
   335  
   336 fun literals_of_term ((Const("Trueprop",_) $ P),lits_ts) = literals_of_term(P,lits_ts)
   337   | literals_of_term ((Const("op |",_) $ P $ Q),(lits,ts)) = 
   338     let val (lits',ts') = literals_of_term(P,(lits,ts))
   339     in
   340         literals_of_term(Q,(lits',ts'))
   341     end
   342   | literals_of_term ((Const("Not",_) $ P),(lits,ts)) = 
   343     let val (pred,ts') = predicate_of P
   344         val lits' = Literal(false,pred,false) :: lits
   345         val ts'' = ResLib.no_rep_app ts ts'
   346     in
   347         (lits',ts'')
   348     end
   349   | literals_of_term (P,(lits,ts)) = 
   350     let val (pred,ts') = predicate_of P
   351         val lits' = Literal(true,pred,false) :: lits
   352         val ts'' = ResLib.no_rep_app ts ts' 
   353     in
   354         (lits',ts'')
   355     end
   356   | literals_of_term _ = raise CLAUSE("Unexpected clause format");
   357      
   358 
   359 fun literals_of_thm thm = 
   360     let val term_of_thm = prop_of thm
   361 			  
   362     in
   363 	literals_of_term (term_of_thm,([],[]))
   364     end;
   365     
   366  
   367 fun sorts_on_typs (_, []) = []
   368   | sorts_on_typs ((FOLTVar(indx)), [s]) = [LTVar((make_type_class s) ^ "(" ^ (make_schematic_type_var(string_of_indexname indx)) ^ ")")]
   369   | sorts_on_typs ((FOLTVar(indx)), (s::ss))=  LTVar((make_type_class s) ^ "(" ^ (make_schematic_type_var(string_of_indexname indx)) ^ ")") :: (sorts_on_typs ((FOLTVar(indx)), ss))
   370   | sorts_on_typs ((FOLTFree(x)), [s]) = [LTFree((make_type_class s) ^ "(" ^ (make_fixed_type_var(x)) ^ ")")]
   371   | sorts_on_typs ((FOLTFree(x)), (s::ss)) = LTFree((make_type_class s) ^ "(" ^ (make_fixed_type_var(x)) ^ ")") :: (sorts_on_typs ((FOLTFree(x)), ss));
   372 
   373 
   374 fun add_typs_aux [] = ([],[])
   375   | add_typs_aux ((FOLTVar(indx),s)::tss) = 
   376     let val vs = sorts_on_typs (FOLTVar(indx),s)
   377         val (vss,fss) = add_typs_aux tss
   378     in
   379         (ResLib.no_rep_app vs vss,fss)
   380     end
   381   | add_typs_aux ((FOLTFree(x),s)::tss) =
   382     let val fs = sorts_on_typs (FOLTFree(x),s)
   383         val (vss,fss) = add_typs_aux tss
   384     in
   385         (vss,ResLib.no_rep_app fs fss)
   386     end;
   387 
   388 
   389 fun add_typs (Clause cls) = 
   390     let val ts = #types_sorts cls
   391     in
   392 	add_typs_aux ts
   393     end;    
   394 
   395 
   396 
   397 
   398 (** make axiom clauses, hypothesis clauses and conjecture clauses. **)
   399 
   400 local 
   401     fun replace_dot "." = "_"
   402       | replace_dot c = c;
   403 
   404 in
   405 
   406 fun proper_ax_name ax_name = 
   407     let val chars = explode ax_name
   408     in
   409 	implode (map replace_dot chars)
   410     end;
   411 end;
   412 
   413 fun make_axiom_clause_thm thm (name,number)=
   414     let val (lits,types_sorts) = literals_of_thm thm
   415 	val cls_id = number
   416 	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
   417 	val ax_name = proper_ax_name name
   418     in 
   419 	make_clause(cls_id,ax_name,Axiom,lits,types_sorts,tvar_lits,tfree_lits)
   420     end;
   421    
   422 fun make_hypothesis_clause_thm thm =
   423     let val (lits,types_sorts) = literals_of_thm thm
   424 	val cls_id = generate_id()
   425 	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
   426     in
   427 	make_clause(cls_id,"",Hypothesis,lits,types_sorts,tvar_lits,tfree_lits)
   428     end;
   429 
   430 
   431 fun make_conjecture_clause_thm thm =
   432     let val (lits,types_sorts) = literals_of_thm thm
   433 	val cls_id = generate_id()
   434 	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
   435     in
   436 	make_clause(cls_id,"",Conjecture,lits,types_sorts,tvar_lits,tfree_lits)
   437     end;
   438 
   439 
   440 fun make_axiom_clause term (name,number)=
   441     let val (lits,types_sorts) = literals_of_term (term,([],[]))
   442 	val cls_id = number
   443 	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
   444 	val ax_name = proper_ax_name name
   445     in 
   446 	make_clause(cls_id,ax_name,Axiom,lits,types_sorts,tvar_lits,tfree_lits)
   447     end;
   448 
   449 
   450 fun make_hypothesis_clause term =
   451     let val (lits,types_sorts) = literals_of_term (term,([],[]))
   452 	val cls_id = generate_id()
   453 	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
   454     in
   455 	make_clause(cls_id,"",Hypothesis,lits,types_sorts,tvar_lits,tfree_lits)
   456     end;
   457     
   458  
   459 fun make_conjecture_clause term =
   460     let val (lits,types_sorts) = literals_of_term (term,([],[]))
   461 	val cls_id = generate_id()
   462 	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
   463     in
   464 	make_clause(cls_id,"",Conjecture,lits,types_sorts,tvar_lits,tfree_lits)
   465     end;
   466  
   467 
   468  
   469 (**** Isabelle arities ****)
   470 
   471 exception ARCLAUSE of string;
   472  
   473 
   474 type class = string; 
   475 type tcons = string; 
   476 
   477 
   478 datatype arLit = TConsLit of bool * (class * tcons * string list) | TVarLit of bool * (class * string);
   479  
   480 datatype arityClause =  
   481 	 ArityClause of {clause_id: clause_id,
   482 			 kind: kind,
   483 			 conclLit: arLit,
   484 			 premLits: arLit list};
   485 
   486 
   487 fun get_TVars 0 = []
   488   | get_TVars n = ("T_" ^ (string_of_int n)) :: get_TVars (n-1);
   489 
   490 
   491 
   492 fun pack_sort(_,[])  = raise ARCLAUSE("Empty Sort Found") 
   493   | pack_sort(tvar, [cls]) = [(make_type_class cls, tvar)] 
   494   | pack_sort(tvar, cls::srt) =  (make_type_class cls,tvar) :: (pack_sort(tvar, srt));
   495     
   496     
   497 fun make_TVarLit (b,(cls,str)) = TVarLit(b,(cls,str));
   498 fun make_TConsLit (b,(cls,tcons,tvars)) = TConsLit(b,(make_type_class cls,make_fixed_type_const tcons,tvars));
   499 
   500 
   501 fun make_arity_clause (clause_id,kind,conclLit,premLits) =
   502     ArityClause {clause_id = clause_id, kind = kind, conclLit = conclLit, premLits = premLits};
   503 
   504 
   505 fun make_axiom_arity_clause (tcons,(res,args)) =
   506      let val cls_id = generate_id()
   507 	 val nargs = length args
   508 	 val tvars = get_TVars nargs
   509 	 val conclLit = make_TConsLit(true,(res,tcons,tvars))
   510          val tvars_srts = ResLib.zip tvars args
   511 	 val tvars_srts' = ResLib.flat_noDup(map pack_sort tvars_srts)
   512          val false_tvars_srts' = ResLib.pair_ins false tvars_srts'
   513 	 val premLits = map make_TVarLit false_tvars_srts'
   514      in
   515 	 make_arity_clause (cls_id,Axiom,conclLit,premLits)
   516      end;
   517     
   518 
   519 
   520 (**** Isabelle class relations ****)
   521 
   522 
   523 datatype classrelClause = 
   524 	 ClassrelClause of {clause_id: clause_id,
   525 			    subclass: class,
   526 			    superclass: class option};
   527 
   528 fun make_classrelClause (clause_id,subclass,superclass) =
   529     ClassrelClause {clause_id = clause_id,subclass = subclass, superclass = superclass};
   530 
   531 
   532 fun make_axiom_classrelClause (subclass,superclass) =
   533     let val cls_id = generate_id()
   534 	val sub = make_type_class subclass
   535 	val sup = case superclass of NONE => NONE 
   536 				   | SOME s => SOME (make_type_class s)
   537     in
   538 	make_classrelClause(cls_id,sub,sup)
   539     end;
   540 
   541 
   542 
   543 fun classrelClauses_of_aux (sub,[]) = []
   544   | classrelClauses_of_aux (sub,(sup::sups)) = make_axiom_classrelClause(sub,SOME sup) :: (classrelClauses_of_aux (sub,sups));
   545 
   546 
   547 fun classrelClauses_of (sub,sups) = 
   548     case sups of [] => [make_axiom_classrelClause (sub,NONE)]
   549 	       | _ => classrelClauses_of_aux (sub, sups);
   550 
   551 
   552 
   553 (***** convert clauses to tptp format *****)
   554 
   555 
   556 fun string_of_clauseID (Clause cls) = clause_prefix ^ (string_of_int (#clause_id cls));
   557 
   558 
   559 fun string_of_kind (Clause cls) = name_of_kind (#kind cls);
   560 
   561 fun string_of_axiomName (Clause cls) = #axiom_name cls;
   562 
   563 (****!!!! Changed for typed equality !!!!****)
   564 fun wrap_eq_type typ t = eq_typ_wrapper ^"(" ^ t ^ "," ^ typ ^ ")";
   565 
   566 
   567 (****!!!! Changed for typed equality !!!!****)
   568 (* Only need to wrap equality's arguments with "typeinfo" if the output clauses are typed && if we specifically ask for types to be included.   *)
   569 fun string_of_equality (typ,terms) =
   570     let val [tstr1,tstr2] = map string_of_term terms
   571     in
   572 	if ((!keep_types) andalso (!special_equal)) then 
   573 	    "equal(" ^ (wrap_eq_type typ tstr1) ^ "," ^ (wrap_eq_type typ tstr2) ^ ")"
   574 	else
   575 	    "equal(" ^ tstr1 ^ "," ^ tstr2 ^ ")"
   576     end
   577 
   578 
   579 and
   580     string_of_term (UVar(x,_)) = x
   581   | string_of_term (Fun("equal",typ,terms)) = string_of_equality(typ,terms)
   582   | string_of_term (Fun (name,typ,[])) = name
   583   | string_of_term (Fun (name,typ,terms)) = 
   584     let val terms' = map string_of_term terms
   585     in
   586         if (!keep_types) then name ^ (ResLib.list_to_string (typ :: terms'))
   587         else name ^ (ResLib.list_to_string terms')
   588     end;
   589 
   590 
   591 
   592 (* Changed for typed equality *)
   593 (* before output the string of the predicate, check if the predicate corresponds to an equality or not. *)
   594 fun string_of_predicate (Predicate("equal",typ,terms)) = string_of_equality(typ,terms)
   595   | string_of_predicate (Predicate(name,_,[])) = name 
   596   | string_of_predicate (Predicate(name,typ,terms)) = 
   597     let val terms_as_strings = map string_of_term terms
   598     in
   599         if (!keep_types) then name ^ (ResLib.list_to_string  (typ :: terms_as_strings))
   600         else name ^ (ResLib.list_to_string terms_as_strings) 
   601     end;
   602 
   603     
   604 
   605 
   606 fun tptp_literal (Literal(pol,pred,tag)) =
   607     let val pred_string = string_of_predicate pred
   608 	val tagged_pol = if (tag andalso !tagged) then (if pol then "+++" else "---")
   609                          else (if pol then "++" else "--")
   610      in
   611 	tagged_pol ^ pred_string
   612     end;
   613 
   614 
   615 
   616 fun tptp_of_typeLit (LTVar x) = "--" ^ x
   617   | tptp_of_typeLit (LTFree x) = "++" ^ x;
   618  
   619 
   620 fun gen_tptp_cls (cls_id,ax_name,knd,lits) = 
   621     let val ax_str = (if ax_name = "" then "" else ("_" ^ ax_name))
   622     in
   623 	"input_clause(" ^ cls_id ^ ax_str ^ "," ^ knd ^ "," ^ lits ^ ")."
   624     end;
   625 
   626 
   627 fun gen_tptp_type_cls (cls_id,knd,tfree_lit,idx) = "input_clause(" ^ cls_id ^ "_tcs" ^ (string_of_int idx) ^ "," ^ knd ^ ",[" ^ tfree_lit ^ "]).";
   628 
   629 
   630 fun tptp_clause_aux (Clause cls) = 
   631     let val lits = map tptp_literal (#literals cls)
   632 	val tvar_lits_strs = if (!keep_types) then (map tptp_of_typeLit (#tvar_type_literals cls)) else []
   633 	val tfree_lits = if (!keep_types) then (map tptp_of_typeLit (#tfree_type_literals cls)) else []
   634     in
   635 	(tvar_lits_strs @ lits,tfree_lits)
   636     end; 
   637 
   638 
   639 fun tptp_clause cls =
   640     let val (lits,tfree_lits) = tptp_clause_aux cls (*"lits" includes the typing assumptions (TVars)*)
   641 	val cls_id = string_of_clauseID cls
   642 	val ax_name = string_of_axiomName cls
   643 	val knd = string_of_kind cls
   644 	val lits_str = ResLib.list_to_string' lits
   645 	val cls_str = gen_tptp_cls(cls_id,ax_name,knd,lits_str) 			fun typ_clss k [] = []
   646           | typ_clss k (tfree :: tfrees) = 
   647             (gen_tptp_type_cls(cls_id,knd,tfree,k)) ::  (typ_clss (k+1) tfrees)
   648     in 
   649 	cls_str :: (typ_clss 0 tfree_lits)
   650     end;
   651 
   652 
   653 fun clause2tptp cls =
   654     let val (lits,tfree_lits) = tptp_clause_aux cls (*"lits" includes the typing assumptions (TVars)*)
   655 	val cls_id = string_of_clauseID cls
   656 	val ax_name = string_of_axiomName cls
   657 	val knd = string_of_kind cls
   658 	val lits_str = ResLib.list_to_string' lits
   659 	val cls_str = gen_tptp_cls(cls_id,ax_name,knd,lits_str) 
   660     in
   661 	(cls_str,tfree_lits) 
   662     end;
   663 
   664 
   665 fun tfree_clause tfree_lit = "input_clause(" ^ "tfree_tcs," ^ "conjecture" ^ ",[" ^ tfree_lit ^ "]).";
   666 
   667 val delim = "\n";
   668 val tptp_clauses2str = ResLib.list2str_sep delim; 
   669      
   670 
   671 fun string_of_arClauseID (ArityClause arcls) = arclause_prefix ^ string_of_int(#clause_id arcls);
   672 
   673 
   674 fun string_of_arLit (TConsLit(b,(c,t,args))) =
   675     let val pol = if b then "++" else "--"
   676 	val  arg_strs = (case args of [] => "" | _ => ResLib.list_to_string args)
   677     in 
   678 	pol ^ c ^ "(" ^ t ^ arg_strs ^ ")"
   679     end
   680   | string_of_arLit (TVarLit(b,(c,str))) =
   681     let val pol = if b then "++" else "--"
   682     in
   683 	pol ^ c ^ "(" ^ str ^ ")"
   684     end;
   685     
   686 
   687 fun string_of_conclLit (ArityClause arcls) = string_of_arLit (#conclLit arcls);
   688      
   689 
   690 fun strings_of_premLits (ArityClause arcls) = map string_of_arLit (#premLits arcls);
   691 		
   692 
   693 fun string_of_arKind (ArityClause arcls) = name_of_kind(#kind arcls);
   694 
   695 fun tptp_arity_clause arcls = 
   696     let val arcls_id = string_of_arClauseID arcls
   697 	val concl_lit = string_of_conclLit arcls
   698 	val prems_lits = strings_of_premLits arcls
   699 	val knd = string_of_arKind arcls
   700 	val all_lits = concl_lit :: prems_lits
   701     in
   702 	"input_clause(" ^ arcls_id ^ "," ^ knd ^ "," ^ (ResLib.list_to_string' all_lits) ^ ")."
   703 	
   704     end;
   705 
   706 
   707 val clrelclause_prefix = "relcls_";
   708 
   709 
   710 fun tptp_classrelLits sub sup = 
   711     let val tvar = "(T)"
   712     in 
   713 	case sup of NONE => "[++" ^ sub ^ tvar ^ "]"
   714 		  | (SOME supcls) =>  "[--" ^ sub ^ tvar ^ ",++" ^ supcls ^ tvar ^ "]"
   715     end;
   716 
   717 
   718 fun tptp_classrelClause (ClassrelClause cls) =
   719     let val relcls_id = clrelclause_prefix ^ string_of_int(#clause_id cls)
   720 	val sub = #subclass cls
   721 	val sup = #superclass cls
   722 	val lits = tptp_classrelLits sub sup
   723     in
   724 	"input_clause(" ^ relcls_id ^ ",axiom," ^ lits ^ ")."
   725     end; 
   726     
   727 end;