give more sensible names to "fol_type" constructors, as requested by a FIXME comment
1 (* Title: HOL/Tools/Sledgehammer/sledgehammer_fol_clause.ML
2 Author: Jia Meng, Cambridge University Computer Laboratory
4 Storing/printing FOL clauses and arity clauses. Typed equality is
7 FIXME: combine with res_hol_clause!
10 signature SLEDGEHAMMER_FOL_CLAUSE =
12 val schematic_var_prefix: string
13 val fixed_var_prefix: string
14 val tvar_prefix: string
15 val tfree_prefix: string
16 val clause_prefix: string
17 val const_prefix: string
18 val tconst_prefix: string
19 val class_prefix: string
20 val union_all: ''a list list -> ''a list
21 val const_trans_table: string Symtab.table
22 val type_const_trans_table: string Symtab.table
23 val ascii_of: string -> string
24 val undo_ascii_of: string -> string
25 val paren_pack : string list -> string
26 val make_schematic_var : string * int -> string
27 val make_fixed_var : string -> string
28 val make_schematic_type_var : string * int -> string
29 val make_fixed_type_var : string -> string
30 val make_fixed_const : bool -> string -> string
31 val make_fixed_type_const : bool -> string -> string
32 val make_type_class : string -> string
33 datatype kind = Axiom | Conjecture
34 type axiom_name = string
38 TyConstr of string * fol_type list
39 val string_of_fol_type : fol_type -> string
40 datatype type_literal = LTVar of string * string | LTFree of string * string
41 exception CLAUSE of string * term
42 val add_typs : typ list -> type_literal list
43 val get_tvar_strs: typ list -> string list
45 TConsLit of class * string * string list
46 | TVarLit of class * string
47 datatype arity_clause = ArityClause of
48 {axiom_name: axiom_name, conclLit: arLit, premLits: arLit list}
49 datatype classrel_clause = ClassrelClause of
50 {axiom_name: axiom_name, subclass: class, superclass: class}
51 val make_classrel_clauses: theory -> class list -> class list -> classrel_clause list
52 val make_arity_clauses_dfg: bool -> theory -> string list -> class list -> class list * arity_clause list
53 val make_arity_clauses: theory -> string list -> class list -> class list * arity_clause list
54 val add_type_sort_preds: typ * int Symtab.table -> int Symtab.table
55 val add_classrel_clause_preds : classrel_clause * int Symtab.table -> int Symtab.table
56 val class_of_arityLit: arLit -> class
57 val add_arity_clause_preds: arity_clause * int Symtab.table -> int Symtab.table
58 val add_foltype_funcs: fol_type * int Symtab.table -> int Symtab.table
59 val add_arity_clause_funcs: arity_clause * int Symtab.table -> int Symtab.table
60 val init_functab: int Symtab.table
61 val dfg_sign: bool -> string -> string
62 val dfg_of_typeLit: bool -> type_literal -> string
63 val gen_dfg_cls: int * string * kind * string list * string list * string list -> string
64 val string_of_preds: (string * Int.int) list -> string
65 val string_of_funcs: (string * int) list -> string
66 val string_of_symbols: string -> string -> string
67 val string_of_start: string -> string
68 val string_of_descrip : string -> string
69 val dfg_tfree_clause : string -> string
70 val dfg_classrel_clause: classrel_clause -> string
71 val dfg_arity_clause: arity_clause -> string
72 val tptp_sign: bool -> string -> string
73 val tptp_of_typeLit : bool -> type_literal -> string
74 val gen_tptp_cls : int * string * kind * string list * string list -> string
75 val tptp_tfree_clause : string -> string
76 val tptp_arity_clause : arity_clause -> string
77 val tptp_classrel_clause : classrel_clause -> string
80 structure Sledgehammer_FOL_Clause : SLEDGEHAMMER_FOL_CLAUSE =
83 open Sledgehammer_Util
85 val schematic_var_prefix = "V_";
86 val fixed_var_prefix = "v_";
88 val tvar_prefix = "T_";
89 val tfree_prefix = "t_";
91 val clause_prefix = "cls_";
92 val arclause_prefix = "clsarity_"
93 val clrelclause_prefix = "clsrel_";
95 val const_prefix = "c_";
96 val tconst_prefix = "tc_";
97 val class_prefix = "class_";
99 fun union_all xss = List.foldl (uncurry (union (op =))) [] xss;
101 (* Provide readable names for the more common symbolic functions *)
102 val const_trans_table =
103 Symtab.make [(@{const_name "op ="}, "equal"),
104 (@{const_name Orderings.less_eq}, "lessequals"),
105 (@{const_name "op &"}, "and"),
106 (@{const_name "op |"}, "or"),
107 (@{const_name "op -->"}, "implies"),
108 (@{const_name "op :"}, "in"),
109 (@{const_name fequal}, "fequal"),
110 (@{const_name COMBI}, "COMBI"),
111 (@{const_name COMBK}, "COMBK"),
112 (@{const_name COMBB}, "COMBB"),
113 (@{const_name COMBC}, "COMBC"),
114 (@{const_name COMBS}, "COMBS")];
116 val type_const_trans_table =
117 Symtab.make [("*", "prod"), ("+", "sum"), ("~=>", "map")];
119 (*Escaping of special characters.
120 Alphanumeric characters are left unchanged.
121 The character _ goes to __
122 Characters in the range ASCII space to / go to _A to _P, respectively.
123 Other printing characters go to _nnn where nnn is the decimal ASCII code.*)
124 val A_minus_space = Char.ord #"A" - Char.ord #" ";
126 fun stringN_of_int 0 _ = ""
127 | stringN_of_int k n = stringN_of_int (k-1) (n div 10) ^ Int.toString (n mod 10);
130 if Char.isAlphaNum c then String.str c
131 else if c = #"_" then "__"
132 else if #" " <= c andalso c <= #"/"
133 then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
134 else if Char.isPrint c
135 then ("_" ^ stringN_of_int 3 (Char.ord c)) (*fixed width, in case more digits follow*)
138 val ascii_of = String.translate ascii_of_c;
140 (** Remove ASCII armouring from names in proof files **)
142 (*We don't raise error exceptions because this code can run inside the watcher.
143 Also, the errors are "impossible" (hah!)*)
144 fun undo_ascii_aux rcs [] = String.implode(rev rcs)
145 | undo_ascii_aux rcs [#"_"] = undo_ascii_aux (#"_"::rcs) [] (*ERROR*)
146 (*Three types of _ escapes: __, _A to _P, _nnn*)
147 | undo_ascii_aux rcs (#"_" :: #"_" :: cs) = undo_ascii_aux (#"_"::rcs) cs
148 | undo_ascii_aux rcs (#"_" :: c :: cs) =
149 if #"A" <= c andalso c<= #"P" (*translation of #" " to #"/"*)
150 then undo_ascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs
152 let val digits = List.take (c::cs, 3) handle Subscript => []
154 case Int.fromString (String.implode digits) of
155 NONE => undo_ascii_aux (c:: #"_"::rcs) cs (*ERROR*)
156 | SOME n => undo_ascii_aux (Char.chr n :: rcs) (List.drop (cs, 2))
158 | undo_ascii_aux rcs (c::cs) = undo_ascii_aux (c::rcs) cs;
160 val undo_ascii_of = undo_ascii_aux [] o String.explode;
162 (* convert a list of strings into one single string; surrounded by brackets *)
163 fun paren_pack [] = "" (*empty argument list*)
164 | paren_pack strings = "(" ^ commas strings ^ ")";
166 (*TSTP format uses (...) rather than the old [...]*)
167 fun tptp_pack strings = "(" ^ space_implode " | " strings ^ ")";
170 (*Remove the initial ' character from a type variable, if it is present*)
171 fun trim_type_var s =
172 if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
173 else error ("trim_type: Malformed type variable encountered: " ^ s);
175 fun ascii_of_indexname (v,0) = ascii_of v
176 | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
178 fun make_schematic_var v = schematic_var_prefix ^ (ascii_of_indexname v);
179 fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x);
181 fun make_schematic_type_var (x,i) =
182 tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
183 fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
185 (* HACK because SPASS 3.0 truncates identifiers to 63 characters. (This is
186 solved in 3.7 and perhaps in earlier versions too.) *)
187 (* 32-bit hash, so we expect no collisions. *)
188 fun controlled_length dfg s =
189 if dfg andalso size s > 60 then Word.toString (hashw_string (s, 0w0)) else s;
191 fun lookup_const dfg c =
192 case Symtab.lookup const_trans_table c of
194 | NONE => controlled_length dfg (ascii_of c);
196 fun lookup_type_const dfg c =
197 case Symtab.lookup type_const_trans_table c of
199 | NONE => controlled_length dfg (ascii_of c);
201 (* "op =" MUST BE "equal" because it's built into ATPs. *)
202 fun make_fixed_const _ (@{const_name "op ="}) = "equal"
203 | make_fixed_const dfg c = const_prefix ^ lookup_const dfg c;
205 fun make_fixed_type_const dfg c = tconst_prefix ^ lookup_type_const dfg c;
207 fun make_type_class clas = class_prefix ^ ascii_of clas;
210 (***** definitions and functions for FOL clauses, for conversion to TPTP or DFG format. *****)
212 datatype kind = Axiom | Conjecture;
214 type axiom_name = string;
216 (**** Isabelle FOL clauses ****)
221 TyConstr of string * fol_type list
223 fun string_of_fol_type (TyVar s) = s
224 | string_of_fol_type (TyFree s) = s
225 | string_of_fol_type (TyConstr (tcon, tps)) =
226 tcon ^ (paren_pack (map string_of_fol_type tps));
228 (*First string is the type class; the second is a TVar or TFfree*)
229 datatype type_literal = LTVar of string * string | LTFree of string * string;
231 exception CLAUSE of string * term;
233 fun atomic_type (TFree (a,_)) = TyFree(make_fixed_type_var a)
234 | atomic_type (TVar (v,_)) = TyVar(make_schematic_type_var v);
236 (*Flatten a type to a fol_type while accumulating sort constraints on the TFrees and
238 fun type_of dfg (Type (a, Ts)) =
239 let val (folTyps, ts) = types_of dfg Ts
240 val t = make_fixed_type_const dfg a
241 in (TyConstr (t, folTyps), ts) end
242 | type_of dfg T = (atomic_type T, [T])
243 and types_of dfg Ts =
244 let val (folTyps,ts) = ListPair.unzip (map (type_of dfg) Ts)
245 in (folTyps, union_all ts) end;
247 (*Make literals for sorted type variables*)
248 fun sorts_on_typs_aux (_, []) = []
249 | sorts_on_typs_aux ((x,i), s::ss) =
250 let val sorts = sorts_on_typs_aux ((x,i), ss)
252 if s = "HOL.type" then sorts
253 else if i = ~1 then LTFree(make_type_class s, make_fixed_type_var x) :: sorts
254 else LTVar(make_type_class s, make_schematic_type_var (x,i)) :: sorts
257 fun sorts_on_typs (TFree (a,s)) = sorts_on_typs_aux ((a,~1),s)
258 | sorts_on_typs (TVar (v,s)) = sorts_on_typs_aux (v,s);
260 fun pred_of_sort (LTVar (s,ty)) = (s,1)
261 | pred_of_sort (LTFree (s,ty)) = (s,1)
263 (*Given a list of sorted type variables, return a list of type literals.*)
264 fun add_typs Ts = List.foldl (uncurry (union (op =))) [] (map sorts_on_typs Ts);
266 (*The correct treatment of TFrees like 'a in lemmas (axiom clauses) is not clear.
267 * Ignoring them leads to unsound proofs, since we do nothing to ensure that 'a
268 in a lemma has the same sort as 'a in the conjecture.
269 * Deleting such clauses will lead to problems with locales in other use of local results
270 where 'a is fixed. Probably we should delete clauses unless the sorts agree.
271 * Currently we include a class constraint in the clause, exactly as with TVars.
274 (** make axiom and conjecture clauses. **)
276 fun get_tvar_strs [] = []
277 | get_tvar_strs ((TVar (indx,s))::Ts) =
278 insert (op =) (make_schematic_type_var indx) (get_tvar_strs Ts)
279 | get_tvar_strs((TFree _)::Ts) = get_tvar_strs Ts
283 (**** Isabelle arities ****)
285 exception ARCLAUSE of string;
287 datatype arLit = TConsLit of class * string * string list
288 | TVarLit of class * string;
290 datatype arity_clause =
291 ArityClause of {axiom_name: axiom_name,
293 premLits: arLit list};
297 | gen_TVars n = ("T_" ^ Int.toString n) :: gen_TVars (n-1);
299 fun pack_sort(_,[]) = []
300 | pack_sort(tvar, "HOL.type"::srt) = pack_sort(tvar, srt) (*IGNORE sort "type"*)
301 | pack_sort(tvar, cls::srt) = (cls, tvar) :: pack_sort(tvar, srt);
303 (*Arity of type constructor tcon :: (arg1,...,argN)res*)
304 fun make_axiom_arity_clause dfg (tcons, axiom_name, (cls,args)) =
305 let val tvars = gen_TVars (length args)
306 val tvars_srts = ListPair.zip (tvars,args)
308 ArityClause {axiom_name = axiom_name,
309 conclLit = TConsLit (cls, make_fixed_type_const dfg tcons, tvars),
310 premLits = map TVarLit (union_all(map pack_sort tvars_srts))}
314 (**** Isabelle class relations ****)
316 datatype classrel_clause =
317 ClassrelClause of {axiom_name: axiom_name,
321 (*Generate all pairs (sub,super) such that sub is a proper subclass of super in theory thy.*)
322 fun class_pairs thy [] supers = []
323 | class_pairs thy subs supers =
324 let val class_less = Sorts.class_less(Sign.classes_of thy)
325 fun add_super sub (super,pairs) =
326 if class_less (sub,super) then (sub,super)::pairs else pairs
327 fun add_supers (sub,pairs) = List.foldl (add_super sub) pairs supers
328 in List.foldl add_supers [] subs end;
330 fun make_classrel_clause (sub,super) =
331 ClassrelClause {axiom_name = clrelclause_prefix ^ ascii_of sub ^ "_" ^ ascii_of super,
332 subclass = make_type_class sub,
333 superclass = make_type_class super};
335 fun make_classrel_clauses thy subs supers =
336 map make_classrel_clause (class_pairs thy subs supers);
339 (** Isabelle arities **)
341 fun arity_clause dfg _ _ (tcons, []) = []
342 | arity_clause dfg seen n (tcons, ("HOL.type",_)::ars) = (*ignore*)
343 arity_clause dfg seen n (tcons,ars)
344 | arity_clause dfg seen n (tcons, (ar as (class,_)) :: ars) =
345 if class mem_string seen then (*multiple arities for the same tycon, class pair*)
346 make_axiom_arity_clause dfg (tcons, lookup_type_const dfg tcons ^ "_" ^ class ^ "_" ^ Int.toString n, ar) ::
347 arity_clause dfg seen (n+1) (tcons,ars)
349 make_axiom_arity_clause dfg (tcons, lookup_type_const dfg tcons ^ "_" ^ class, ar) ::
350 arity_clause dfg (class::seen) n (tcons,ars)
352 fun multi_arity_clause dfg [] = []
353 | multi_arity_clause dfg ((tcons,ars) :: tc_arlists) =
354 arity_clause dfg [] 1 (tcons, ars) @ multi_arity_clause dfg tc_arlists
356 (*Generate all pairs (tycon,class,sorts) such that tycon belongs to class in theory thy
357 provided its arguments have the corresponding sorts.*)
358 fun type_class_pairs thy tycons classes =
359 let val alg = Sign.classes_of thy
360 fun domain_sorts (tycon,class) = Sorts.mg_domain alg tycon [class]
361 fun add_class tycon (class,pairs) =
362 (class, domain_sorts(tycon,class))::pairs
363 handle Sorts.CLASS_ERROR _ => pairs
364 fun try_classes tycon = (tycon, List.foldl (add_class tycon) [] classes)
365 in map try_classes tycons end;
367 (*Proving one (tycon, class) membership may require proving others, so iterate.*)
368 fun iter_type_class_pairs thy tycons [] = ([], [])
369 | iter_type_class_pairs thy tycons classes =
370 let val cpairs = type_class_pairs thy tycons classes
371 val newclasses = union_all (union_all (union_all (map (map #2 o #2) cpairs)))
372 |> subtract (op =) classes |> subtract (op =) HOLogic.typeS
373 val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
374 in (union (op =) classes' classes, union (op =) cpairs' cpairs) end;
376 fun make_arity_clauses_dfg dfg thy tycons classes =
377 let val (classes', cpairs) = iter_type_class_pairs thy tycons classes
378 in (classes', multi_arity_clause dfg cpairs) end;
379 val make_arity_clauses = make_arity_clauses_dfg false;
381 (**** Find occurrences of predicates in clauses ****)
383 (*FIXME: multiple-arity checking doesn't work, as update_new is the wrong
384 function (it flags repeated declarations of a function, even with the same arity)*)
386 fun update_many (tab, keypairs) = List.foldl (uncurry Symtab.update) tab keypairs;
388 fun add_type_sort_preds (T, preds) =
389 update_many (preds, map pred_of_sort (sorts_on_typs T));
391 fun add_classrel_clause_preds (ClassrelClause {subclass,superclass,...}, preds) =
392 Symtab.update (subclass,1) (Symtab.update (superclass,1) preds);
394 fun class_of_arityLit (TConsLit (tclass, _, _)) = tclass
395 | class_of_arityLit (TVarLit (tclass, _)) = tclass;
397 fun add_arity_clause_preds (ArityClause {conclLit,premLits,...}, preds) =
398 let val classes = map (make_type_class o class_of_arityLit) (conclLit::premLits)
399 fun upd (class,preds) = Symtab.update (class,1) preds
400 in List.foldl upd preds classes end;
402 (*** Find occurrences of functions in clauses ***)
404 fun add_foltype_funcs (TyVar _, funcs) = funcs
405 | add_foltype_funcs (TyFree a, funcs) = Symtab.update (a,0) funcs
406 | add_foltype_funcs (TyConstr (a, tys), funcs) =
407 List.foldl add_foltype_funcs (Symtab.update (a, length tys) funcs) tys;
409 (*TFrees are recorded as constants*)
410 fun add_type_sort_funcs (TVar _, funcs) = funcs
411 | add_type_sort_funcs (TFree (a, _), funcs) =
412 Symtab.update (make_fixed_type_var a, 0) funcs
414 fun add_arity_clause_funcs (ArityClause {conclLit,...}, funcs) =
415 let val TConsLit (_, tcons, tvars) = conclLit
416 in Symtab.update (tcons, length tvars) funcs end;
418 (*This type can be overlooked because it is built-in...*)
419 val init_functab = Symtab.update ("tc_itself", 1) Symtab.empty;
422 (**** String-oriented operations ****)
424 fun string_of_clausename (cls_id,ax_name) =
425 clause_prefix ^ ascii_of ax_name ^ "_" ^ Int.toString cls_id;
427 fun string_of_type_clsname (cls_id,ax_name,idx) =
428 string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx);
431 (**** Producing DFG files ****)
433 (*Attach sign in DFG syntax: false means negate.*)
434 fun dfg_sign true s = s
435 | dfg_sign false s = "not(" ^ s ^ ")"
437 fun dfg_of_typeLit pos (LTVar (s,ty)) = dfg_sign pos (s ^ "(" ^ ty ^ ")")
438 | dfg_of_typeLit pos (LTFree (s,ty)) = dfg_sign pos (s ^ "(" ^ ty ^ ")");
440 (*Enclose the clause body by quantifiers, if necessary*)
441 fun dfg_forall [] body = body
442 | dfg_forall vars body = "forall([" ^ commas vars ^ "],\n" ^ body ^ ")"
444 fun gen_dfg_cls (cls_id, ax_name, Axiom, lits, tylits, vars) =
445 "clause( %(axiom)\n" ^
446 dfg_forall vars ("or(" ^ commas (tylits@lits) ^ ")") ^ ",\n" ^
447 string_of_clausename (cls_id,ax_name) ^ ").\n\n"
448 | gen_dfg_cls (cls_id, ax_name, Conjecture, lits, _, vars) =
449 "clause( %(negated_conjecture)\n" ^
450 dfg_forall vars ("or(" ^ commas lits ^ ")") ^ ",\n" ^
451 string_of_clausename (cls_id,ax_name) ^ ").\n\n";
453 fun string_of_arity (name, num) = "(" ^ name ^ "," ^ Int.toString num ^ ")"
455 fun string_of_preds [] = ""
456 | string_of_preds preds = "predicates[" ^ commas(map string_of_arity preds) ^ "].\n";
458 fun string_of_funcs [] = ""
459 | string_of_funcs funcs = "functions[" ^ commas(map string_of_arity funcs) ^ "].\n" ;
461 fun string_of_symbols predstr funcstr =
462 "list_of_symbols.\n" ^ predstr ^ funcstr ^ "end_of_list.\n\n";
464 fun string_of_start name = "begin_problem(" ^ name ^ ").\n\n";
466 fun string_of_descrip name =
467 "list_of_descriptions.\nname({*" ^ name ^
468 "*}).\nauthor({*Isabelle*}).\nstatus(unknown).\ndescription({*auto-generated*}).\nend_of_list.\n\n"
470 fun dfg_tfree_clause tfree_lit =
471 "clause( %(negated_conjecture)\n" ^ "or( " ^ tfree_lit ^ "),\n" ^ "tfree_tcs" ^ ").\n\n"
473 fun dfg_of_arLit (TConsLit (c,t,args)) =
474 dfg_sign true (make_type_class c ^ "(" ^ t ^ paren_pack args ^ ")")
475 | dfg_of_arLit (TVarLit (c,str)) =
476 dfg_sign false (make_type_class c ^ "(" ^ str ^ ")")
478 fun dfg_classrelLits sub sup = "not(" ^ sub ^ "(T)), " ^ sup ^ "(T)";
480 fun dfg_classrel_clause (ClassrelClause {axiom_name,subclass,superclass,...}) =
481 "clause(forall([T],\nor( " ^ dfg_classrelLits subclass superclass ^ ")),\n" ^
482 axiom_name ^ ").\n\n";
484 fun string_of_ar axiom_name = arclause_prefix ^ ascii_of axiom_name;
486 fun dfg_arity_clause (ArityClause{axiom_name,conclLit,premLits,...}) =
487 let val TConsLit (_,_,tvars) = conclLit
488 val lits = map dfg_of_arLit (conclLit :: premLits)
490 "clause( %(axiom)\n" ^
491 dfg_forall tvars ("or( " ^ commas lits ^ ")") ^ ",\n" ^
492 string_of_ar axiom_name ^ ").\n\n"
496 (**** Produce TPTP files ****)
498 fun tptp_sign true s = s
499 | tptp_sign false s = "~ " ^ s
501 fun tptp_of_typeLit pos (LTVar (s, ty)) = tptp_sign pos (s ^ "(" ^ ty ^ ")")
502 | tptp_of_typeLit pos (LTFree (s, ty)) = tptp_sign pos (s ^ "(" ^ ty ^ ")")
504 fun tptp_cnf name kind formula =
505 "cnf(" ^ name ^ ", " ^ kind ^ ",\n " ^ formula ^ ").\n"
507 fun gen_tptp_cls (cls_id, ax_name, Axiom, lits, tylits) =
508 tptp_cnf (string_of_clausename (cls_id, ax_name)) "axiom"
509 (tptp_pack (tylits @ lits))
510 | gen_tptp_cls (cls_id, ax_name, Conjecture, lits, _) =
511 tptp_cnf (string_of_clausename (cls_id, ax_name)) "negated_conjecture"
514 fun tptp_tfree_clause tfree_lit =
515 tptp_cnf "tfree_tcs" "negated_conjecture" (tptp_pack [tfree_lit])
517 fun tptp_of_arLit (TConsLit (c,t,args)) =
518 tptp_sign true (make_type_class c ^ "(" ^ t ^ paren_pack args ^ ")")
519 | tptp_of_arLit (TVarLit (c,str)) =
520 tptp_sign false (make_type_class c ^ "(" ^ str ^ ")")
522 fun tptp_arity_clause (ArityClause{axiom_name,conclLit,premLits,...}) =
523 tptp_cnf (string_of_ar axiom_name) "axiom"
524 (tptp_pack (map tptp_of_arLit (conclLit :: premLits)))
526 fun tptp_classrelLits sub sup =
528 in tptp_pack [tptp_sign false (sub^tvar), tptp_sign true (sup^tvar)] end;
530 fun tptp_classrel_clause (ClassrelClause {axiom_name,subclass,superclass,...}) =
531 tptp_cnf axiom_name "axiom" (tptp_classrelLits subclass superclass)