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