src/HOL/Tools/res_hol_clause.ML
author mengj
Tue, 18 Apr 2006 05:38:18 +0200
changeset 19444 085568445283
parent 19354 aebf9dddccd7
child 19452 ef0ed2fe7bf2
permissions -rw-r--r--
Take conjectures and axioms as thms when convert them to ResHolClause.clause format.
     1 (* ID: $Id$ 
     2    Author: Jia Meng, NICTA
     3 
     4 FOL clauses translated from HOL formulae.  Combinators are used to represent lambda terms.
     5 
     6 *)
     7 
     8 structure ResHolClause =
     9 
    10 struct
    11 
    12 
    13 val include_combS = ref false;
    14 val include_min_comb = ref false;
    15 
    16 val const_typargs = ref (Library.K [] : (string*typ -> typ list));
    17 
    18 fun init thy = (include_combS:=false;include_min_comb:=false;const_typargs := Sign.const_typargs thy);
    19 
    20 
    21 
    22 (**********************************************************************)
    23 (* convert a Term.term with lambdas into a Term.term with combinators *) 
    24 (**********************************************************************)
    25 
    26 fun is_free (Bound(a)) n = (a = n)
    27   | is_free (Abs(x,_,b)) n = (is_free b (n+1))
    28   | is_free (P $ Q) n = ((is_free P n) orelse (is_free Q n))
    29   | is_free _ _ = false;
    30 
    31 
    32 exception LAM2COMB of term;
    33 
    34 exception BND of term;
    35 
    36 fun decre_bndVar (Bound n) = Bound (n-1)
    37   | decre_bndVar (P $ Q) = (decre_bndVar P) $ (decre_bndVar Q)
    38   | decre_bndVar t =
    39     case t of Const(_,_) => t
    40 	    | Free(_,_) => t
    41 	    | Var(_,_) => t
    42 	    | Abs(_,_,_) => raise BND(t); (*should not occur*)
    43 
    44 
    45 (*******************************************)
    46 fun lam2comb (Abs(x,tp,Bound 0)) _ = 
    47     let val tpI = Type("fun",[tp,tp])
    48     in 
    49 	include_min_comb:=true;
    50 	Const("COMBI",tpI) 
    51     end
    52   | lam2comb (Abs(x,tp,Bound n)) Bnds = 
    53     let val (Bound n') = decre_bndVar (Bound n)
    54 	val tb = List.nth(Bnds,n')
    55 	val tK = Type("fun",[tb,Type("fun",[tp,tb])])
    56     in
    57 	include_min_comb:=true;
    58 	Const("COMBK",tK) $ (Bound n')
    59     end
    60   | lam2comb (Abs(x,t1,Const(c,t2))) _ = 
    61     let val tK = Type("fun",[t2,Type("fun",[t1,t2])])
    62     in 
    63 	include_min_comb:=true;
    64 	Const("COMBK",tK) $ Const(c,t2) 
    65     end
    66   | lam2comb (Abs(x,t1,Free(v,t2))) _ =
    67     let val tK = Type("fun",[t2,Type("fun",[t1,t2])])
    68     in
    69 	include_min_comb:=true;
    70 	Const("COMBK",tK) $ Free(v,t2)
    71     end
    72   | lam2comb (Abs(x,t1,Var(ind,t2))) _=
    73     let val tK = Type("fun",[t2,Type("fun",[t1,t2])])
    74     in
    75 	include_min_comb:=true;
    76 	Const("COMBK",tK) $ Var(ind,t2)
    77     end
    78   | lam2comb (t as (Abs(x,t1,P$(Bound 0)))) Bnds =
    79     let val nfreeP = not(is_free P 0)
    80 	val tr = Term.type_of1(t1::Bnds,P)
    81     in
    82 	if nfreeP then (decre_bndVar P)
    83 	else (
    84 	      let val tI = Type("fun",[t1,t1])
    85 		  val P' = lam2comb (Abs(x,t1,P)) Bnds
    86 		  val tp' = Term.type_of1(Bnds,P')
    87 		  val tS = Type("fun",[tp',Type("fun",[tI,tr])])
    88 	      in
    89 		  include_min_comb:=true;
    90 		  include_combS:=true;
    91 		  Const("COMBS",tS) $ P' $ Const("COMBI",tI)
    92 	      end)
    93     end
    94 	    
    95   | lam2comb (t as (Abs(x,t1,P$Q))) Bnds =
    96     let val (nfreeP,nfreeQ) = (not(is_free P 0),not(is_free Q 0))
    97 	val tpq = Term.type_of1(t1::Bnds, P$Q) 
    98     in
    99 	if(nfreeP andalso nfreeQ) then (
   100 	    let val tK = Type("fun",[tpq,Type("fun",[t1,tpq])])
   101 		val PQ' = decre_bndVar(P $ Q)
   102 	    in 
   103 		include_min_comb:=true;
   104 		Const("COMBK",tK) $ PQ'
   105 	    end)
   106 	else (
   107 	      if nfreeP then (
   108 			       let val Q' = lam2comb (Abs(x,t1,Q)) Bnds
   109 				   val P' = decre_bndVar P
   110 				   val tp = Term.type_of1(Bnds,P')
   111 				   val tq' = Term.type_of1(Bnds, Q')
   112 				   val tB = Type("fun",[tp,Type("fun",[tq',Type("fun",[t1,tpq])])])
   113 			       in
   114 				   include_min_comb:=true;
   115 				   Const("COMBB",tB) $ P' $ Q' 
   116 			       end)
   117 	      else (
   118 		    if nfreeQ then (
   119 				    let val P' = lam2comb (Abs(x,t1,P)) Bnds
   120 					val Q' = decre_bndVar Q
   121 					val tq = Term.type_of1(Bnds,Q')
   122 					val tp' = Term.type_of1(Bnds, P')
   123 					val tC = Type("fun",[tp',Type("fun",[tq,Type("fun",[t1,tpq])])])
   124 				    in
   125 					include_min_comb:=true;
   126 					Const("COMBC",tC) $ P' $ Q'
   127 				    end)
   128 		    else(
   129 			 let val P' = lam2comb (Abs(x,t1,P)) Bnds
   130 			     val Q' = lam2comb (Abs(x,t1,Q)) Bnds
   131 			     val tp' = Term.type_of1(Bnds,P')
   132 			     val tq' = Term.type_of1(Bnds,Q')
   133 			     val tS = Type("fun",[tp',Type("fun",[tq',Type("fun",[t1,tpq])])])
   134 			 in
   135 			     include_min_comb:=true;
   136 			     include_combS:=true;
   137 			     Const("COMBS",tS) $ P' $ Q'
   138 			 end)))
   139     end
   140   | lam2comb (t as (Abs(x,t1,_))) _ = raise LAM2COMB (t);
   141 
   142 	     
   143 
   144 (*********************)
   145 
   146 fun to_comb (Abs(x,tp,b)) Bnds =
   147     let val b' = to_comb b (tp::Bnds)
   148     in lam2comb (Abs(x,tp,b')) Bnds end
   149   | to_comb (P $ Q) Bnds = ((to_comb P Bnds) $ (to_comb Q Bnds))
   150   | to_comb t _ = t;
   151  
   152     
   153 fun comb_of t = to_comb t [];
   154 
   155 
   156 (* print a term containing combinators, used for debugging *)
   157 exception TERM_COMB of term;
   158 
   159 fun string_of_term (Const(c,t)) = c
   160   | string_of_term (Free(v,t)) = v
   161   | string_of_term (Var((x,n),t)) =
   162     let val xn = x ^ "_" ^ (string_of_int n)
   163     in xn end
   164   | string_of_term (P $ Q) =
   165     let val P' = string_of_term P
   166 	val Q' = string_of_term Q
   167     in
   168 	"(" ^ P' ^ " " ^ Q' ^ ")" end
   169   | string_of_term t =  raise TERM_COMB (t);
   170 
   171 
   172 
   173 (******************************************************)
   174 (* data types for typed combinator expressions        *)
   175 (******************************************************)
   176 
   177 type axiom_name = string;
   178 datatype kind = Axiom | Conjecture;
   179 fun name_of_kind Axiom = "axiom"
   180   | name_of_kind Conjecture = "conjecture";
   181 
   182 type polarity = bool;
   183 type indexname = Term.indexname;
   184 type clause_id = int;
   185 type csort = Term.sort;
   186 type ctyp = ResClause.fol_type;
   187 
   188 val string_of_ctyp = ResClause.string_of_fol_type;
   189 
   190 type ctyp_var = ResClause.typ_var;
   191 
   192 type ctype_literal = ResClause.type_literal;
   193 
   194 
   195 datatype combterm = CombConst of string * ctyp * ctyp list
   196 		  | CombFree of string * ctyp
   197 		  | CombVar of string * ctyp
   198 		  | CombApp of combterm * combterm * ctyp
   199 		  | Bool of combterm
   200 		  | Equal of combterm * combterm;
   201 datatype literal = Literal of polarity * combterm;
   202 
   203 
   204 
   205 datatype clause = 
   206 	 Clause of {clause_id: clause_id,
   207 		    axiom_name: axiom_name,
   208 		    kind: kind,
   209 		    literals: literal list,
   210 		    ctypes_sorts: (ctyp_var * csort) list, 
   211                     ctvar_type_literals: ctype_literal list, 
   212                     ctfree_type_literals: ctype_literal list};
   213 
   214 
   215 
   216 fun string_of_kind (Clause cls) = name_of_kind (#kind cls);
   217 fun get_axiomName (Clause cls) = #axiom_name cls;
   218 fun get_clause_id (Clause cls) = #clause_id cls;
   219 
   220 fun get_literals (c as Clause(cls)) = #literals cls;
   221 
   222 
   223 
   224 exception TERM_ORD of string
   225 
   226 fun term_ord (CombVar(_,_),CombVar(_,_)) = EQUAL
   227   | term_ord (CombVar(_,_),_) = LESS
   228   | term_ord (CombFree(_,_),CombVar(_,_)) = GREATER
   229   | term_ord (CombFree(f1,tp1),CombFree(f2,tp2)) = 
   230     let val ord1 = string_ord(f1,f2)
   231     in
   232 	case ord1 of EQUAL => ResClause.types_ord ([tp1],[tp2])
   233 		   | _ => ord1
   234     end
   235   | term_ord (CombFree(_,_),_) = LESS
   236   | term_ord (CombConst(_,_,_),CombVar(_,_)) = GREATER
   237   | term_ord (CombConst(_,_,_),CombFree(_,_)) = GREATER
   238   | term_ord (CombConst(c1,tp1,_),CombConst(c2,tp2,_)) = 
   239     let val ord1 = string_ord (c1,c2)
   240     in
   241 	case ord1 of EQUAL => ResClause.types_ord ([tp1],[tp2])
   242 		   | _ => ord1
   243     end
   244   | term_ord (CombConst(_,_,_),_) = LESS
   245   | term_ord (CombApp(_,_,_),Bool(_)) = raise TERM_ORD("bool")
   246   | term_ord (CombApp(_,_,_),Equal(_,_)) = LESS
   247   | term_ord (CombApp(f1,arg1,tp1),CombApp(f2,arg2,tp2)) =
   248     let val ord1 = term_ord (f1,f2)
   249 	val ord2 = case ord1 of EQUAL => term_ord (arg1,arg2)
   250 			      | _ => ord1
   251     in
   252 	case ord2 of EQUAL => ResClause.types_ord ([tp1],[tp2])
   253 		   | _ => ord2
   254     end
   255   | term_ord (CombApp(_,_,_),_) = GREATER
   256   | term_ord (Bool(_),_) = raise TERM_ORD("bool")
   257   | term_ord (Equal(t1,t2),Equal(t3,t4)) = ResClause.list_ord term_ord ([t1,t2],[t3,t4])
   258   | term_ord (Equal(_,_),_) = GREATER;
   259 
   260 fun predicate_ord (Equal(_,_),Bool(_)) = LESS
   261   | predicate_ord (Equal(t1,t2),Equal(t3,t4)) = 
   262     ResClause.list_ord term_ord ([t1,t2],[t3,t4])
   263   | predicate_ord (Bool(_),Equal(_,_)) = GREATER
   264   | predicate_ord (Bool(t1),Bool(t2)) = term_ord (t1,t2)
   265 
   266 
   267 fun literal_ord (Literal(false,_),Literal(true,_)) = LESS
   268   | literal_ord (Literal(true,_),Literal(false,_)) = GREATER
   269   | literal_ord (Literal(_,pred1),Literal(_,pred2)) = predicate_ord(pred1,pred2);
   270 
   271 fun sort_lits lits = sort literal_ord lits;
   272 
   273 (*********************************************************************)
   274 (* convert a clause with type Term.term to a clause with type clause *)
   275 (*********************************************************************)
   276 
   277 fun isFalse (Literal(pol,Bool(CombConst(c,_,_)))) =
   278     (pol andalso c = "c_False") orelse
   279     (not pol andalso c = "c_True")
   280   | isFalse _ = false;
   281 
   282 
   283 fun isTrue (Literal (pol,Bool(CombConst(c,_,_)))) =
   284       (pol andalso c = "c_True") orelse
   285       (not pol andalso c = "c_False")
   286   | isTrue _ = false;
   287   
   288 fun isTaut (Clause {literals,...}) = exists isTrue literals;  
   289 
   290 
   291 
   292 fun make_clause(clause_id,axiom_name,kind,literals,ctypes_sorts,ctvar_type_literals,ctfree_type_literals) =
   293     if forall isFalse literals
   294     then error "Problem too trivial for resolution (empty clause)"
   295     else
   296 	Clause {clause_id = clause_id, axiom_name = axiom_name, kind = kind,
   297 		literals = literals, ctypes_sorts = ctypes_sorts, 
   298 		ctvar_type_literals = ctvar_type_literals,
   299 		ctfree_type_literals = ctfree_type_literals};
   300 
   301 fun type_of (Type (a, Ts)) =
   302     let val (folTypes,ts) = types_of Ts
   303 	val t = ResClause.make_fixed_type_const a
   304     in
   305 	(ResClause.mk_fol_type("Comp",t,folTypes),ts)
   306     end
   307   | type_of (tp as (TFree(a,s))) =
   308     let val t = ResClause.make_fixed_type_var a
   309     in
   310 	(ResClause.mk_fol_type("Fixed",t,[]),[ResClause.mk_typ_var_sort tp])
   311     end
   312   | type_of (tp as (TVar(v,s))) =
   313     let val t = ResClause.make_schematic_type_var v
   314     in
   315 	(ResClause.mk_fol_type("Var",t,[]),[ResClause.mk_typ_var_sort tp])
   316     end
   317 
   318 and types_of Ts =
   319     let val foltyps_ts = map type_of Ts
   320 	val (folTyps,ts) = ListPair.unzip foltyps_ts
   321     in
   322 	(folTyps,ResClause.union_all ts)
   323     end;
   324 
   325 (* same as above, but no gathering of sort information *)
   326 fun simp_type_of (Type (a, Ts)) = 
   327     let val typs = map simp_type_of Ts
   328 	val t = ResClause.make_fixed_type_const a
   329     in
   330 	ResClause.mk_fol_type("Comp",t,typs)
   331     end
   332   | simp_type_of (TFree (a,s)) = ResClause.mk_fol_type("Fixed",ResClause.make_fixed_type_var a,[])
   333   | simp_type_of (TVar (v,s)) = ResClause.mk_fol_type("Var",ResClause.make_schematic_type_var v,[]);
   334 
   335 fun comb_typ ("COMBI",t) = 
   336     let val t' = domain_type t
   337     in
   338 	[simp_type_of t']
   339     end
   340   | comb_typ ("COMBK",t) = 
   341     let val a = domain_type t
   342 	val b = domain_type (range_type t)
   343     in
   344 	map simp_type_of [a,b]
   345     end
   346   | comb_typ ("COMBS",t) = 
   347     let val t' = domain_type t
   348 	val a = domain_type t'
   349 	val b = domain_type (range_type t')
   350 	val c = range_type (range_type t')
   351     in 
   352 	map simp_type_of [a,b,c]
   353     end
   354   | comb_typ ("COMBB",t) = 
   355     let val ab = domain_type t
   356 	val ca = domain_type (range_type t)
   357 	val a = domain_type ab
   358 	val b = range_type ab
   359 	val c = domain_type ca
   360     in
   361 	map simp_type_of [a,b,c]
   362     end
   363   | comb_typ ("COMBC",t) =
   364     let val t1 = domain_type t
   365 	val a = domain_type t1
   366 	val b = domain_type (range_type t1)
   367 	val c = range_type (range_type t1)
   368     in
   369 	map simp_type_of [a,b,c]
   370     end;
   371 
   372 fun const_type_of ("COMBI",t) = 
   373     let val (tp,ts) = type_of t
   374 	val I_var = comb_typ ("COMBI",t)
   375     in
   376 	(tp,ts,I_var)
   377     end
   378   | const_type_of ("COMBK",t) =
   379     let val (tp,ts) = type_of t
   380 	val K_var = comb_typ ("COMBK",t)
   381     in
   382 	(tp,ts,K_var)
   383     end
   384   | const_type_of ("COMBS",t) =
   385     let val (tp,ts) = type_of t
   386 	val S_var = comb_typ ("COMBS",t)
   387     in
   388 	(tp,ts,S_var)
   389     end
   390   | const_type_of ("COMBB",t) =
   391     let val (tp,ts) = type_of t
   392 	val B_var = comb_typ ("COMBB",t)
   393     in
   394 	(tp,ts,B_var)
   395     end
   396   | const_type_of ("COMBC",t) =
   397     let val (tp,ts) = type_of t
   398 	val C_var = comb_typ ("COMBC",t)
   399     in
   400 	(tp,ts,C_var)
   401     end
   402   | const_type_of (c,t) =
   403     let val (tp,ts) = type_of t
   404 	val tvars = !const_typargs(c,t)
   405 	val tvars' = map simp_type_of tvars
   406     in
   407 	(tp,ts,tvars')
   408     end;
   409 
   410 fun is_bool_type (Type("bool",[])) = true
   411   | is_bool_type _ = false;
   412 
   413 
   414 (* convert a Term.term (with combinators) into a combterm, also accummulate sort info *)
   415 fun combterm_of (Const(c,t)) =
   416     let val (tp,ts,tvar_list) = const_type_of (c,t)
   417 	val is_bool = is_bool_type t
   418 	val c' = CombConst(ResClause.make_fixed_const c,tp,tvar_list)
   419 	val c'' = if is_bool then Bool(c') else c'
   420     in
   421 	(c'',ts)
   422     end
   423   | combterm_of (Free(v,t)) =
   424     let val (tp,ts) = type_of t
   425 	val is_bool = is_bool_type t
   426 	val v' = if ResClause.isMeta v then CombVar(ResClause.make_schematic_var(v,0),tp)
   427 		 else CombFree(ResClause.make_fixed_var v,tp)
   428 	val v'' = if is_bool then Bool(v') else v'
   429     in
   430 	(v'',ts)
   431     end
   432   | combterm_of (Var(v,t)) =
   433     let val (tp,ts) = type_of t
   434 	val is_bool = is_bool_type t
   435 	val v' = CombVar(ResClause.make_schematic_var v,tp)
   436 	val v'' = if is_bool then Bool(v') else v'
   437     in
   438 	(v'',ts)
   439     end
   440   | combterm_of (Const("op =",T) $ P $ Q) = (*FIXME: allow equal between bools?*)
   441     let val (P',tsP) = combterm_of P        
   442 	val (Q',tsQ) = combterm_of Q
   443     in
   444 	(Equal(P',Q'),tsP union tsQ)
   445     end
   446   | combterm_of (t as (P $ Q)) =
   447     let val (P',tsP) = combterm_of P
   448 	val (Q',tsQ) = combterm_of Q
   449 	val tp = Term.type_of t
   450 	val tp' = simp_type_of tp
   451 	val is_bool = is_bool_type tp
   452 	val t' = CombApp(P',Q',tp')
   453 	val t'' = if is_bool then Bool(t') else t'
   454     in
   455 	(t'',tsP union tsQ)
   456     end;
   457 
   458 fun predicate_of ((Const("Not",_) $ P), polarity) =
   459     predicate_of (P, not polarity)
   460   | predicate_of (term,polarity) = (combterm_of term,polarity);
   461 
   462 
   463 fun literals_of_term1 args (Const("Trueprop",_) $ P) = literals_of_term1 args P
   464   | literals_of_term1 args (Const("op |",_) $ P $ Q) = 
   465     let val args' = literals_of_term1 args P
   466     in
   467 	literals_of_term1 args' Q
   468     end
   469   | literals_of_term1 (lits,ts) P =
   470     let val ((pred,ts'),pol) = predicate_of (P,true)
   471 	val lits' = Literal(pol,pred)::lits
   472     in
   473 	(lits',ts union ts')
   474     end;
   475 
   476 
   477 fun literals_of_term P = literals_of_term1 ([],[]) P;
   478 
   479 
   480 (* making axiom and conjecture clauses *)
   481 fun make_axiom_clause thm (ax_name,cls_id) =
   482     let val term = prop_of thm
   483 	val term' = comb_of term
   484 	val (lits,ctypes_sorts) = literals_of_term term'
   485 	val lits' = sort_lits lits
   486 	val (ctvar_lits,ctfree_lits) = ResClause.add_typs_aux ctypes_sorts
   487     in
   488 	make_clause(cls_id,ax_name,Axiom,
   489 		    lits',ctypes_sorts,ctvar_lits,ctfree_lits)
   490     end;
   491 
   492 fun make_axiom_clauses [] = []
   493   | make_axiom_clauses ((thm,(name,id))::thms) =
   494     let val cls = make_axiom_clause thm (name,id)
   495 	val clss = make_axiom_clauses thms
   496     in
   497 	if isTaut cls then clss else (cls::clss)
   498     end;
   499 
   500 
   501 fun make_conjecture_clause n thm =
   502     let val t = prop_of thm
   503 	val t' = comb_of t
   504 	val (lits,ctypes_sorts) = literals_of_term t'
   505 	val (ctvar_lits,ctfree_lits) = ResClause.add_typs_aux ctypes_sorts
   506     in
   507 	make_clause(n,"conjecture",Conjecture,lits,ctypes_sorts,ctvar_lits,ctfree_lits)
   508     end;
   509 
   510 
   511 
   512 fun make_conjecture_clauses_aux _ [] = []
   513   | make_conjecture_clauses_aux n (t::ts) =
   514     make_conjecture_clause n t :: make_conjecture_clauses_aux (n+1) ts;
   515 
   516 val make_conjecture_clauses = make_conjecture_clauses_aux 0;
   517 
   518 
   519 (**********************************************************************)
   520 (* convert clause into ATP specific formats:                          *)
   521 (* TPTP used by Vampire and E                                         *)
   522 (**********************************************************************)
   523 
   524 val type_wrapper = "typeinfo";
   525 
   526 datatype type_level = T_FULL | T_PARTIAL | T_CONST | T_NONE;
   527 
   528 val typ_level = ref T_FULL;
   529 
   530 fun full_types () = (typ_level:=T_FULL);
   531 fun partial_types () = (typ_level:=T_PARTIAL);
   532 fun const_types_only () = (typ_level:=T_CONST);
   533 fun no_types () = (typ_level:=T_NONE);
   534 
   535 
   536 fun find_typ_level () = !typ_level;
   537 
   538 fun wrap_type (c,t) = 
   539     case !typ_level of T_FULL => type_wrapper ^ (ResClause.paren_pack [c,t])
   540 		     | _ => c;
   541     
   542 
   543 val bool_tp = ResClause.make_fixed_type_const "bool";
   544 
   545 val app_str = "hAPP";
   546 
   547 val bool_str = "hBOOL";
   548 
   549 exception STRING_OF_COMBTERM of int;
   550 
   551 (* convert literals of clauses into strings *)
   552 fun string_of_combterm1_aux _ (CombConst(c,tp,_)) = 
   553     let val tp' = string_of_ctyp tp
   554     in
   555 	(wrap_type (c,tp'),tp')
   556     end
   557   | string_of_combterm1_aux _ (CombFree(v,tp)) = 
   558     let val tp' = string_of_ctyp tp
   559     in
   560 	(wrap_type (v,tp'),tp')
   561     end
   562   | string_of_combterm1_aux _ (CombVar(v,tp)) = 
   563     let val tp' = string_of_ctyp tp
   564     in
   565 	(wrap_type (v,tp'),tp')
   566     end
   567   | string_of_combterm1_aux is_pred (CombApp(t1,t2,tp)) =
   568     let val (s1,tp1) = string_of_combterm1_aux is_pred t1
   569 	val (s2,tp2) = string_of_combterm1_aux is_pred t2
   570 	val tp' = ResClause.string_of_fol_type tp
   571 	val r =	case !typ_level of T_FULL => type_wrapper ^  (ResClause.paren_pack [(app_str ^ (ResClause.paren_pack [s1,s2])),tp'])
   572 				 | T_PARTIAL => app_str ^ (ResClause.paren_pack [s1,s2,tp1])
   573 				 | T_NONE => app_str ^ (ResClause.paren_pack [s1,s2])
   574 				 | T_CONST => raise STRING_OF_COMBTERM (1) (*should not happen, if happened may be a bug*)
   575     in	(r,tp')
   576 
   577     end
   578   | string_of_combterm1_aux is_pred (Bool(t)) = 
   579     let val (t',_) = string_of_combterm1_aux false t
   580 	val r = if is_pred then bool_str ^ (ResClause.paren_pack [t'])
   581 		else t'
   582     in
   583 	(r,bool_tp)
   584     end
   585   | string_of_combterm1_aux _ (Equal(t1,t2)) =
   586     let val (s1,_) = string_of_combterm1_aux false t1
   587 	val (s2,_) = string_of_combterm1_aux false t2
   588     in
   589 	("equal" ^ (ResClause.paren_pack [s1,s2]),bool_tp) 
   590     end;
   591 
   592 fun string_of_combterm1 is_pred term = fst (string_of_combterm1_aux is_pred term);
   593 
   594 fun string_of_combterm2 _ (CombConst(c,tp,tvars)) = 
   595     let val tvars' = map string_of_ctyp tvars
   596     in
   597 	c ^ (ResClause.paren_pack tvars')
   598     end
   599   | string_of_combterm2 _ (CombFree(v,tp)) = v
   600   | string_of_combterm2 _ (CombVar(v,tp)) = v
   601   | string_of_combterm2 is_pred (CombApp(t1,t2,tp)) =
   602     let val s1 = string_of_combterm2 is_pred t1
   603 	val s2 = string_of_combterm2 is_pred t2
   604     in
   605 	app_str ^ (ResClause.paren_pack [s1,s2])
   606     end
   607   | string_of_combterm2 is_pred (Bool(t)) = 
   608     let val t' = string_of_combterm2 false t
   609     in
   610 	if is_pred then bool_str ^ (ResClause.paren_pack [t'])
   611 	else t'
   612     end
   613   | string_of_combterm2 _ (Equal(t1,t2)) =
   614     let val s1 = string_of_combterm2 false t1
   615 	val s2 = string_of_combterm2 false t2
   616     in
   617 	("equal" ^ (ResClause.paren_pack [s1,s2])) 
   618     end;
   619 
   620 
   621 
   622 fun string_of_combterm is_pred term = 
   623     case !typ_level of T_CONST => string_of_combterm2 is_pred term
   624 		     | _ => string_of_combterm1 is_pred term;
   625 
   626 
   627 fun string_of_clausename (cls_id,ax_name) = 
   628     ResClause.clause_prefix ^ ResClause.ascii_of ax_name ^ "_" ^ Int.toString cls_id;
   629 
   630 fun string_of_type_clsname (cls_id,ax_name,idx) = 
   631     string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx);
   632 
   633 
   634 fun tptp_literal (Literal(pol,pred)) =
   635     let val pred_string = string_of_combterm true pred
   636 	val pol_str = if pol then "++" else "--"
   637     in
   638 	pol_str ^ pred_string
   639     end;
   640 
   641  
   642 fun tptp_type_lits (Clause cls) = 
   643     let val lits = map tptp_literal (#literals cls)
   644 	val ctvar_lits_strs =
   645 	    case !typ_level of T_NONE => []
   646 			     | _ => (map ResClause.tptp_of_typeLit (#ctvar_type_literals cls)) 
   647 	val ctfree_lits = 
   648 	    case !typ_level of T_NONE => []
   649 			     | _ => (map ResClause.tptp_of_typeLit (#ctfree_type_literals cls)) 
   650     in
   651 	(ctvar_lits_strs @ lits, ctfree_lits)
   652     end; 
   653     
   654     
   655 fun clause2tptp cls =
   656     let val (lits,ctfree_lits) = tptp_type_lits cls
   657 	val cls_id = get_clause_id cls
   658 	val ax_name = get_axiomName cls
   659 	val knd = string_of_kind cls
   660 	val lits_str = ResClause.bracket_pack lits
   661 	val cls_str = ResClause.gen_tptp_cls(cls_id,ax_name,knd,lits_str)
   662     in
   663 	(cls_str,ctfree_lits)
   664     end;
   665 
   666 
   667 
   668 (**********************************************************************)
   669 (* clause equalities and hashing functions                            *)
   670 (**********************************************************************)
   671 
   672 
   673 fun combterm_eq (CombConst(c1,tp1,tps1),CombConst(c2,tp2,tps2)) vtvars =
   674     let val (eq1,vtvars1) = if c1 = c2 then ResClause.types_eq (tps1,tps2) vtvars
   675 			    else (false,vtvars)
   676     in
   677 	(eq1,vtvars1)
   678     end
   679   | combterm_eq (CombConst(_,_,_),_) vtvars = (false,vtvars)
   680   | combterm_eq (CombFree(a1,tp1),CombFree(a2,tp2)) vtvars = 
   681     if a1 = a2 then ResClause.types_eq ([tp1],[tp2]) vtvars
   682     else (false,vtvars)
   683   | combterm_eq (CombFree(_,_),_) vtvars = (false,vtvars)
   684   | combterm_eq (CombVar(v1,tp1),CombVar(v2,tp2)) (vars,tvars) = 
   685     (case ResClause.check_var_pairs(v1,v2) vars of 0 => ResClause.types_eq ([tp1],[tp2]) ((v1,v2)::vars,tvars)
   686 						 | 1 => ResClause.types_eq ([tp1],[tp2]) (vars,tvars)
   687 						 | 2 => (false,(vars,tvars)))
   688   | combterm_eq (CombVar(_,_),_) vtvars = (false,vtvars)
   689   | combterm_eq (CombApp(f1,arg1,tp1),CombApp(f2,arg2,tp2)) vtvars =
   690     let val (eq1,vtvars1) = combterm_eq (f1,f2) vtvars
   691 	val (eq2,vtvars2) = if eq1 then combterm_eq (arg1,arg2) vtvars1
   692 			    else (eq1,vtvars1)
   693     in
   694 	if eq2 then ResClause.types_eq ([tp1],[tp2]) vtvars2
   695 	else (eq2,vtvars2)
   696     end
   697   | combterm_eq (CombApp(_,_,_),_) vtvars = (false,vtvars)
   698   | combterm_eq (Bool(t1),Bool(t2)) vtvars = combterm_eq (t1,t2) vtvars
   699   | combterm_eq (Bool(_),_) vtvars = (false,vtvars)
   700   | combterm_eq (Equal(t1,t2),Equal(t3,t4)) vtvars =
   701     let val (eq1,vtvars1) = combterm_eq (t1,t3) vtvars
   702     in
   703 	if eq1 then combterm_eq (t2,t4) vtvars1
   704 	else (eq1,vtvars1)
   705     end
   706   | combterm_eq (Equal(t1,t2),_) vtvars = (false,vtvars);
   707 
   708 fun lit_eq (Literal(pol1,pred1),Literal(pol2,pred2)) vtvars =
   709     if (pol1 = pol2) then combterm_eq (pred1,pred2) vtvars
   710     else (false,vtvars);
   711 
   712 fun lits_eq ([],[]) vtvars = (true,vtvars)
   713   | lits_eq (l1::ls1,l2::ls2) vtvars = 
   714     let val (eq1,vtvars1) = lit_eq (l1,l2) vtvars
   715     in
   716 	if eq1 then lits_eq (ls1,ls2) vtvars1
   717 	else (false,vtvars1)
   718     end;
   719 
   720 fun clause_eq (cls1,cls2) =
   721     let val lits1 = get_literals cls1
   722 	val lits2 = get_literals cls2
   723     in
   724 	length lits1 = length lits2 andalso #1 (lits_eq (lits1,lits2) ([],[]))
   725     end;
   726 
   727 val xor_words = List.foldl Word.xorb 0w0;
   728 
   729 fun hash_combterm (CombVar(_,_),w) = w
   730   | hash_combterm (CombFree(f,_),w) = Polyhash.hashw_string(f,w)
   731   | hash_combterm (CombConst(c,tp,tps),w) = Polyhash.hashw_string(c,w)
   732   | hash_combterm (CombApp(f,arg,tp),w) = hash_combterm (arg, hash_combterm (f,w))
   733   | hash_combterm (Bool(t),w) = hash_combterm (t,w)
   734   | hash_combterm (Equal(t1,t2),w) = 
   735     List.foldl hash_combterm (Polyhash.hashw_string ("equal",w)) [t1,t2]
   736 
   737 fun hash_literal (Literal(true,pred)) = hash_combterm(pred,0w0)
   738   | hash_literal (Literal(false,pred)) = Word.notb(hash_combterm(pred,0w0));
   739 
   740 fun hash_clause clause = xor_words (map hash_literal (get_literals clause));
   741 
   742 (**********************************************************************)
   743 (* write clauses to files                                             *)
   744 (**********************************************************************)
   745 
   746 fun read_in [] = []
   747   | read_in (f1::fs) =
   748     let val lines = read_in fs
   749 	val input = TextIO.openIn f1
   750 	fun reading () =
   751 	    let val nextline = TextIO.inputLine input
   752 	    in
   753 		if nextline = "" then (TextIO.closeIn input; [])
   754 		else nextline::(reading ())
   755 	    end
   756     in
   757 	((reading ()) @ lines)
   758     end;
   759 
   760  
   761 fun get_helper_clauses (full,partial,const,untyped) =
   762     let val (helper1,noS,inclS) = case !typ_level of T_FULL => (warning "Fully-typed HOL"; full)
   763 						   | T_PARTIAL => (warning "Partially-typed HOL"; partial)
   764 						   | T_CONST => (warning "Const-only-typed HOL"; const)
   765 						   | T_NONE => (warning "Untyped HOL"; untyped)
   766 	val needed_helpers = if !include_combS then (warning "Include combinator S"; [helper1,inclS]) else
   767 			     if !include_min_comb then (warning "Include min combinators"; [helper1,noS])
   768 			     else (warning "No combinator is used"; [helper1])
   769     in
   770 	read_in needed_helpers
   771     end;
   772 
   773 
   774 (* write TPTP format to a single file *)
   775 (* when "get_helper_clauses" is called, "include_combS" and "include_min_comb" should have correct values already *)
   776 fun tptp_write_file thms filename (axclauses,classrel_clauses,arity_clauses) helpers =
   777     let val clss = make_conjecture_clauses thms
   778 	val axclauses' = make_axiom_clauses axclauses
   779 	val (tptp_clss,tfree_litss) = ListPair.unzip (map clause2tptp clss)
   780 	val tfree_clss = map ResClause.tptp_tfree_clause (foldl (op union_string) [] tfree_litss)
   781 	val out = TextIO.openOut filename
   782 	val helper_clauses = get_helper_clauses helpers
   783     in
   784 	List.app (curry TextIO.output out o #1 o clause2tptp) axclauses';
   785 	ResClause.writeln_strs out tfree_clss;
   786 	ResClause.writeln_strs out tptp_clss;
   787 	List.app (curry TextIO.output out o ResClause.tptp_classrelClause) classrel_clauses;
   788 	List.app (curry TextIO.output out o ResClause.tptp_arity_clause) arity_clauses;
   789 	List.app (curry TextIO.output out) helper_clauses;
   790 	TextIO.closeOut out
   791     end;
   792 
   793 end