1 (* Author: Jia Meng, Cambridge University Computer Laboratory
4 Copyright 2004 University of Cambridge
6 ML data structure for storing/printing FOL clauses and arity clauses.
7 Typed equality is treated differently.
10 (* works for writeoutclasimp on typed *)
11 signature RES_CLAUSE =
13 val keep_types : bool ref
14 val special_equal : bool ref
17 exception ARCLAUSE of string
18 exception CLAUSE of string * term
22 val init : theory -> unit
23 val make_axiom_clause : Term.term -> string * int -> clause
24 val make_conjecture_clauses : term list -> clause list
25 val get_axiomName : clause -> string
26 val isTaut : clause -> bool
27 val num_of_clauses : clause -> int
29 val clause2dfg : clause -> string * string list
30 val clauses2dfg : clause list -> string -> clause list -> clause list ->
31 (string * int) list -> (string * int) list -> string
32 val tfree_dfg_clause : string -> string
34 val arity_clause_thy: theory -> arityClause list
35 val classrel_clauses_thy: theory -> classrelClause list
37 val tptp_arity_clause : arityClause -> string
38 val tptp_classrelClause : classrelClause -> string
39 val tptp_clause : clause -> string list
40 val clause2tptp : clause -> string * string list
41 val tfree_clause : string -> string
42 val schematic_var_prefix : string
43 val fixed_var_prefix : string
44 val tvar_prefix : string
45 val tfree_prefix : string
46 val clause_prefix : string
47 val arclause_prefix : string
48 val const_prefix : string
49 val tconst_prefix : string
50 val class_prefix : string
52 val union_all : ''a list list -> ''a list
53 val ascii_of : String.string -> String.string
54 val paren_pack : string list -> string
55 val bracket_pack : string list -> string
56 val make_schematic_var : String.string * int -> string
57 val make_fixed_var : String.string -> string
58 val make_schematic_type_var : string * int -> string
59 val make_fixed_type_var : string -> string
60 val make_fixed_const : String.string -> string
61 val make_fixed_type_const : String.string -> string
62 val make_type_class : String.string -> string
63 val isMeta : String.string -> bool
66 val mk_typ_var_sort : Term.typ -> typ_var * sort
68 val add_typs_aux2 : (typ_var * string list) list -> type_literal list * type_literal list
69 val gen_tptp_cls : int * string * string * string -> string
70 val gen_tptp_type_cls : int * string * string * string * int -> string
71 val tptp_of_typeLit : type_literal -> string
75 structure ResClause: RES_CLAUSE =
78 (* Added for typed equality *)
79 val special_equal = ref false; (* by default,equality does not carry type information *)
80 val eq_typ_wrapper = "typeinfo"; (* default string *)
83 val schematic_var_prefix = "V_";
84 val fixed_var_prefix = "v_";
86 val tvar_prefix = "T_";
87 val tfree_prefix = "t_";
89 val clause_prefix = "cls_";
90 val arclause_prefix = "clsarity_"
91 val clrelclause_prefix = "clsrel_";
93 val const_prefix = "c_";
94 val tconst_prefix = "tc_";
96 val class_prefix = "class_";
99 fun union_all xss = foldl (op union) [] xss;
102 (*Provide readable names for the more common symbolic functions*)
103 val const_trans_table =
104 Symtab.make [("op =", "equal"),
105 ("op <=", "lessequals"),
112 ("op -->", "implies"),
117 ("List.op @", "append")];
119 val type_const_trans_table =
120 Symtab.make [("*", "t_prod"),
124 (*Escaping of special characters.
125 Alphanumeric characters are left unchanged.
126 The character _ goes to __
127 Characters in the range ASCII space to / go to _A to _P, respectively.
128 Other printing characters go to _NNN where NNN is the decimal ASCII code.*)
131 val A_minus_space = Char.ord #"A" - Char.ord #" ";
134 if Char.isAlphaNum c then String.str c
135 else if c = #"_" then "__"
136 else if #" " <= c andalso c <= #"/"
137 then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
138 else if Char.isPrint c then ("_" ^ Int.toString (Char.ord c))
143 val ascii_of = String.translate ascii_of_c;
147 (* convert a list of strings into one single string; surrounded by brackets *)
148 fun paren_pack [] = "" (*empty argument list*)
149 | paren_pack strings = "(" ^ commas strings ^ ")";
151 fun bracket_pack strings = "[" ^ commas strings ^ "]";
154 (*Remove the initial ' character from a type variable, if it is present*)
155 fun trim_type_var s =
156 if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
157 else error ("trim_type: Malformed type variable encountered: " ^ s);
159 fun ascii_of_indexname (v,0) = ascii_of v
160 | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
162 fun make_schematic_var v = schematic_var_prefix ^ (ascii_of_indexname v);
163 fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x);
165 fun make_schematic_type_var (x,i) =
166 tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
167 fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
169 fun make_fixed_const c =
170 case Symtab.lookup const_trans_table c of
172 | NONE => const_prefix ^ ascii_of c;
174 fun make_fixed_type_const c =
175 case Symtab.lookup type_const_trans_table c of
177 | NONE => tconst_prefix ^ ascii_of c;
179 fun make_type_class clas = class_prefix ^ ascii_of clas;
183 (***** definitions and functions for FOL clauses, prepared for conversion into TPTP format or SPASS format. *****)
185 val keep_types = ref true;
187 datatype kind = Axiom | Hypothesis | Conjecture;
188 fun name_of_kind Axiom = "axiom"
189 | name_of_kind Hypothesis = "hypothesis"
190 | name_of_kind Conjecture = "conjecture";
192 type clause_id = int;
193 type axiom_name = string;
196 type polarity = bool;
198 type indexname = Term.indexname;
201 (* "tag" is used for vampire specific syntax *)
205 (**** Isabelle FOL clauses ****)
207 val tagged = ref false;
209 type pred_name = string;
210 type sort = Term.sort;
211 type fol_type = string;
213 datatype type_literal = LTVar of string | LTFree of string;
215 datatype folTerm = UVar of string * fol_type
216 | Fun of string * fol_type list * folTerm list;
217 datatype predicate = Predicate of pred_name * fol_type list * folTerm list;
219 datatype literal = Literal of polarity * predicate * tag;
221 datatype typ_var = FOLTVar of indexname | FOLTFree of string;
223 fun mk_typ_var_sort (TFree(a,s)) = (FOLTFree a,s)
224 | mk_typ_var_sort (TVar(v,s)) = (FOLTVar v,s);
228 (* ML datatype used to repsent one single clause: disjunction of literals. *)
230 Clause of {clause_id: clause_id,
231 axiom_name: axiom_name,
233 literals: literal list,
234 types_sorts: (typ_var * sort) list,
235 tvar_type_literals: type_literal list,
236 tfree_type_literals: type_literal list ,
238 predicates: (string*int) list,
239 functions: (string*int) list};
242 exception CLAUSE of string * term;
245 (*** make clauses ***)
247 fun isFalse (Literal (pol,Predicate(a,_,[]),_)) =
248 (pol andalso a = "c_False") orelse
249 (not pol andalso a = "c_True")
252 fun isTrue (Literal (pol,Predicate(a,_,[]),_)) =
253 (pol andalso a = "c_True") orelse
254 (not pol andalso a = "c_False")
257 fun isTaut (Clause {literals,...}) = exists isTrue literals;
259 fun make_clause (clause_id,axiom_name,kind,literals,
260 types_sorts,tvar_type_literals,
261 tfree_type_literals,tvars, predicates, functions) =
262 if forall isFalse literals
263 then error "Problem too trivial for resolution (empty clause)"
265 Clause {clause_id = clause_id, axiom_name = axiom_name, kind = kind,
266 literals = literals, types_sorts = types_sorts,
267 tvar_type_literals = tvar_type_literals,
268 tfree_type_literals = tfree_type_literals,
269 tvars = tvars, predicates = predicates,
270 functions = functions};
273 (** Some Clause destructor functions **)
275 fun string_of_kind (Clause cls) = name_of_kind (#kind cls);
277 fun get_axiomName (Clause cls) = #axiom_name cls;
279 fun get_clause_id (Clause cls) = #clause_id cls;
281 fun funcs_of_cls (Clause cls) = #functions cls;
283 fun preds_of_cls (Clause cls) = #predicates cls;
286 (*Declarations of the current theory--to allow suppressing types.*)
287 val const_typargs = ref (Library.K [] : (string*typ -> typ list));
289 fun num_typargs(s,T) = if !keep_types then length (!const_typargs (s,T)) else 0;
291 (*Initialize the type suppression mechanism with the current theory before
292 producing any clauses!*)
293 fun init thy = (const_typargs := Sign.const_typargs thy);
296 (*Flatten a type to a string while accumulating sort constraints on the TFrees and
298 fun type_of (Type (a, Ts)) =
299 let val (folTyps, (ts, funcs)) = types_of Ts
300 val t = make_fixed_type_const a
302 ((t ^ paren_pack folTyps), (ts, (t, length Ts)::funcs))
304 | type_of (TFree (a,s)) =
305 let val t = make_fixed_type_var a
306 in (t, ([((FOLTFree a),s)], [(t,0)])) end
307 | type_of (TVar (v, s)) = (make_schematic_type_var v, ([((FOLTVar v),s)], []))
309 let val foltyps_ts = map type_of Ts
310 val (folTyps,ts_funcs) = ListPair.unzip foltyps_ts
311 val (ts, funcslist) = ListPair.unzip ts_funcs
313 (folTyps, (union_all ts, union_all funcslist))
317 fun const_types_of (c,T) = types_of (!const_typargs (c,T));
319 (* Any variables created via the METAHYPS tactical should be treated as
320 universal vars, although it is represented as "Free(...)" by Isabelle *)
321 val isMeta = String.isPrefix "METAHYP1_"
323 fun pred_name_type (Const(c,T)) =
324 let val (contys,(folTyps,funcs)) = const_types_of (c,T)
325 in (make_fixed_const c, (contys,folTyps), funcs) end
326 | pred_name_type (Free(x,T)) =
327 if isMeta x then raise CLAUSE("Predicate Not First Order 1", Free(x,T))
328 else (make_fixed_var x, ([],[]), [])
329 | pred_name_type (v as Var _) = raise CLAUSE("Predicate Not First Order 2", v)
330 | pred_name_type t = raise CLAUSE("Predicate input unexpected", t);
333 (* For typed equality *)
334 (* here "arg_typ" is the type of "="'s argument's type, not the type of the equality *)
335 (* Find type of equality arg *)
336 fun eq_arg_type (Type("fun",[T,_])) =
337 let val (folT,_) = type_of T;
340 fun fun_name_type (Const("op =",T)) args = (*FIXME: Is this special treatment of = needed??*)
341 let val t = make_fixed_const "op ="
342 in (t, ([eq_arg_type T], []), [(t,2)]) end
343 | fun_name_type (Const(c,T)) args =
344 let val t = make_fixed_const c
345 val (contys, (folTyps,funcs)) = const_types_of (c,T)
346 val arity = num_typargs(c,T) + length args
348 (t, (contys,folTyps), ((t,arity)::funcs))
350 | fun_name_type (Free(x,T)) args =
351 let val t = make_fixed_var x
352 in (t, ([],[]), [(t, length args)]) end
353 | fun_name_type f args = raise CLAUSE("Function Not First Order 1", f);
356 fun term_of (Var(ind_nm,T)) =
357 let val (folType,(ts,funcs)) = type_of T
359 (UVar(make_schematic_var ind_nm, folType), (ts, funcs))
361 | term_of (Free(x,T)) =
362 let val (folType, (ts,funcs)) = type_of T
364 if isMeta x then (UVar(make_schematic_var(x,0),folType),
365 (ts, ((make_schematic_var(x,0)),0)::funcs))
367 (Fun(make_fixed_var x, [folType], []),
368 (ts, ((make_fixed_var x),0)::funcs))
370 | term_of (Const(c,T)) = (* impossible to be equality *)
371 let val (contys, (folTyps,funcs)) = const_types_of (c,T)
373 (Fun(make_fixed_const c, contys, []),
374 (folTyps, ((make_fixed_const c),0)::funcs))
377 let val (f,args) = strip_comb app
378 val _ = case f of Const(_,_) => ()
381 then raise CLAUSE("Function Not First Order 2", f)
383 | _ => raise CLAUSE("Function Not First Order 3", f);
384 val (funName,(contys,ts1),funcs) = fun_name_type f args
385 val (args',(ts2,funcs')) = terms_of args
387 (Fun(funName,contys,args'),
388 (union_all (ts1::ts2),
389 union_all(funcs::funcs')))
391 | term_of t = raise CLAUSE("Function Not First Order 4", t)
393 let val (args, ts_funcs) = ListPair.unzip (map term_of ts)
395 (args, ListPair.unzip ts_funcs)
399 fun pred_of (Const("op =", typ), args) =
400 let val arg_typ = eq_arg_type typ
401 val (args',(ts,funcs)) = terms_of args
402 val equal_name = make_fixed_const "op ="
404 (Predicate(equal_name,[arg_typ],args'),
406 [((make_fixed_var equal_name), 2)],
409 | pred_of (pred,args) =
410 let val (predName,(predType,ts1), pfuncs) = pred_name_type pred
411 val (args',(ts2,ffuncs)) = terms_of args
412 val ts3 = union_all (ts1::ts2)
413 val ffuncs' = union_all ffuncs
414 val newfuncs = pfuncs union ffuncs'
417 Const (c,T) => num_typargs(c,T) + length args
420 (Predicate(predName,predType,args'), ts3,
421 [(predName, arity)], newfuncs)
425 (*Treatment of literals, possibly negated or tagged*)
426 fun predicate_of ((Const("Not",_) $ P), polarity, tag) =
427 predicate_of (P, not polarity, tag)
428 | predicate_of ((Const("HOL.tag",_) $ P), polarity, tag) =
429 predicate_of (P, polarity, true)
430 | predicate_of (term,polarity,tag) =
431 (pred_of (strip_comb term), polarity, tag);
433 fun literals_of_term1 args (Const("Trueprop",_) $ P) = literals_of_term1 args P
434 | literals_of_term1 (args as (lits, ts, preds, funcs)) (Const("op |",_) $ P $ Q) =
435 let val (lits', ts', preds', funcs') = literals_of_term1 args P
437 literals_of_term1 (lits', ts', preds' union preds, funcs' union funcs) Q
439 | literals_of_term1 (lits, ts, preds, funcs) P =
440 let val ((pred, ts', preds', funcs'), pol, tag) = predicate_of (P,true,false)
441 val lits' = Literal(pol,pred,tag) :: lits
443 (lits', ts union ts', preds' union preds, funcs' union funcs)
447 val literals_of_term = literals_of_term1 ([],[],[],[]);
450 (* FIX: not sure what to do with these funcs *)
452 (*Make literals for sorted type variables*)
453 fun sorts_on_typs (_, []) = ([])
454 | sorts_on_typs (v, "HOL.type" :: s) =
455 sorts_on_typs (v,s) (*Ignore sort "type"*)
456 | sorts_on_typs ((FOLTVar indx), (s::ss)) =
457 LTVar((make_type_class s) ^
458 "(" ^ (make_schematic_type_var indx) ^ ")") ::
459 (sorts_on_typs ((FOLTVar indx), ss))
460 | sorts_on_typs ((FOLTFree x), (s::ss)) =
461 LTFree((make_type_class s) ^ "(" ^ (make_fixed_type_var x) ^ ")") ::
462 (sorts_on_typs ((FOLTFree x), ss));
465 (*UGLY: seems to be parsing the "show sorts" output, removing anything that
466 starts with a left parenthesis.*)
467 fun remove_type str = hd (String.fields (fn c => c = #"(") str);
469 fun pred_of_sort (LTVar x) = ((remove_type x),1)
470 | pred_of_sort (LTFree x) = ((remove_type x),1)
475 (*Given a list of sorted type variables, return two separate lists.
476 The first is for TVars, the second for TFrees.*)
477 fun add_typs_aux [] preds = ([],[], preds)
478 | add_typs_aux ((FOLTVar indx,s)::tss) preds =
479 let val vs = sorts_on_typs (FOLTVar indx, s)
480 val preds' = (map pred_of_sort vs)@preds
481 val (vss,fss, preds'') = add_typs_aux tss preds'
483 (vs union vss, fss, preds'')
485 | add_typs_aux ((FOLTFree x,s)::tss) preds =
486 let val fs = sorts_on_typs (FOLTFree x, s)
487 val preds' = (map pred_of_sort fs)@preds
488 val (vss,fss, preds'') = add_typs_aux tss preds'
490 (vss, fs union fss, preds'')
493 fun add_typs_aux2 [] = ([],[])
494 | add_typs_aux2 ((FOLTVar indx,s)::tss) =
495 let val vs = sorts_on_typs (FOLTVar indx,s)
496 val (vss,fss) = add_typs_aux2 tss
500 | add_typs_aux2 ((FOLTFree x,s)::tss) =
501 let val fs = sorts_on_typs (FOLTFree x,s)
502 val (vss,fss) = add_typs_aux2 tss
508 fun add_typs (Clause cls) preds = add_typs_aux (#types_sorts cls) preds
511 (** make axiom clauses, hypothesis clauses and conjecture clauses. **)
513 fun get_tvar_strs [] = []
514 | get_tvar_strs ((FOLTVar indx,s)::tss) =
515 let val vstr = make_schematic_type_var indx
517 vstr ins (get_tvar_strs tss)
519 | get_tvar_strs((FOLTFree x,s)::tss) = distinct (get_tvar_strs tss)
521 (* FIX add preds and funcs to add typs aux here *)
523 fun make_axiom_clause_thm thm (ax_name,cls_id) =
524 let val (lits,types_sorts, preds, funcs) = literals_of_term (prop_of thm)
525 val (tvar_lits,tfree_lits, preds) = add_typs_aux types_sorts preds
526 val tvars = get_tvar_strs types_sorts
528 make_clause(cls_id,ax_name,Axiom,
529 lits,types_sorts,tvar_lits,tfree_lits,
534 (* check if a clause is FOL first*)
535 fun make_conjecture_clause n t =
536 let val _ = check_is_fol_term t
537 handle TERM("check_is_fol_term",_) => raise CLAUSE("Goal is not FOL",t)
538 val (lits,types_sorts, preds, funcs) = literals_of_term t
539 val (tvar_lits,tfree_lits, preds) = add_typs_aux types_sorts preds
540 val tvars = get_tvar_strs types_sorts
542 make_clause(n,"conjecture",Conjecture,
543 lits,types_sorts,tvar_lits,tfree_lits,
547 fun make_conjecture_clauses_aux _ [] = []
548 | make_conjecture_clauses_aux n (t::ts) =
549 make_conjecture_clause n t :: make_conjecture_clauses_aux (n+1) ts
551 val make_conjecture_clauses = make_conjecture_clauses_aux 0
554 (*before converting an axiom clause to "clause" format, check if it is FOL*)
555 fun make_axiom_clause term (ax_name,cls_id) =
556 let val _ = check_is_fol_term term
557 handle TERM("check_is_fol_term",_) => raise CLAUSE("Axiom is not FOL", term)
558 val (lits,types_sorts, preds,funcs) = literals_of_term term
559 val (tvar_lits,tfree_lits, preds) = add_typs_aux types_sorts preds
560 val tvars = get_tvar_strs types_sorts
562 make_clause(cls_id,ax_name,Axiom,
563 lits,types_sorts,tvar_lits,tfree_lits,
570 (**** Isabelle arities ****)
572 exception ARCLAUSE of string;
579 datatype arLit = TConsLit of bool * (class * tcons * string list) | TVarLit of bool * (class * string);
581 datatype arityClause =
582 ArityClause of {clause_id: clause_id,
583 axiom_name: axiom_name,
586 premLits: arLit list};
590 | get_TVars n = ("T_" ^ (Int.toString n)) :: get_TVars (n-1);
594 fun pack_sort(_,[]) = raise ARCLAUSE("Empty Sort Found")
595 | pack_sort(tvar, [cls]) = [(make_type_class cls, tvar)]
596 | pack_sort(tvar, cls::srt) = (make_type_class cls,tvar) :: (pack_sort(tvar, srt));
599 fun make_TVarLit (b,(cls,str)) = TVarLit(b,(cls,str));
600 fun make_TConsLit (b,(cls,tcons,tvars)) = TConsLit(b,(make_type_class cls,make_fixed_type_const tcons,tvars));
602 fun make_axiom_arity_clause (tcons,n,(res,args)) =
603 let val nargs = length args
604 val tvars = get_TVars nargs
605 val tvars_srts = ListPair.zip (tvars,args)
606 val tvars_srts' = union_all(map pack_sort tvars_srts)
607 val false_tvars_srts' = map (pair false) tvars_srts'
609 ArityClause {clause_id = n, kind = Axiom,
611 conclLit = make_TConsLit(true,(res,tcons,tvars)),
612 premLits = map make_TVarLit false_tvars_srts'}
615 (*The number of clauses generated from cls, including type clauses*)
616 fun num_of_clauses (Clause cls) =
617 let val num_tfree_lits =
618 if !keep_types then length (#tfree_type_literals cls)
620 in 1 + num_tfree_lits end;
623 (**** Isabelle class relations ****)
626 datatype classrelClause =
627 ClassrelClause of {clause_id: clause_id,
629 superclass: class option};
632 fun make_axiom_classrelClause n subclass superclass =
633 ClassrelClause {clause_id = n,
634 subclass = subclass, superclass = superclass};
637 fun classrelClauses_of_aux n sub [] = []
638 | classrelClauses_of_aux n sub (sup::sups) =
639 make_axiom_classrelClause n sub (SOME sup) :: classrelClauses_of_aux (n+1) sub sups;
642 fun classrelClauses_of (sub,sups) =
643 case sups of [] => [make_axiom_classrelClause 0 sub NONE]
644 | _ => classrelClauses_of_aux 0 sub sups;
647 (***** Isabelle arities *****)
650 fun arity_clause _ (tcons, []) = []
651 | arity_clause n (tcons, ar::ars) =
652 make_axiom_arity_clause (tcons,n,ar) ::
653 arity_clause (n+1) (tcons,ars);
655 fun multi_arity_clause [] = []
656 | multi_arity_clause (tcon_ar :: tcons_ars) =
657 arity_clause 0 tcon_ar @ multi_arity_clause tcons_ars
659 fun arity_clause_thy thy =
660 let val arities = #arities (Type.rep_tsig (Sign.tsig_of thy))
661 in multi_arity_clause (Symtab.dest arities) end;
664 (* Isabelle classes *)
666 type classrelClauses = classrelClause list Symtab.table;
668 val classrel_of = #2 o #classes o Type.rep_tsig o Sign.tsig_of;
669 fun classrel_clauses_classrel (C: Sorts.classes) = map classrelClauses_of (Graph.dest C);
670 val classrel_clauses_thy = List.concat o classrel_clauses_classrel o classrel_of;
674 (****!!!! Changed for typed equality !!!!****)
676 fun wrap_eq_type typ t = eq_typ_wrapper ^"(" ^ t ^ "," ^ typ ^ ")";
678 (*Only need to wrap equality's arguments with "typeinfo" if the output clauses are typed
679 and if we specifically ask for types to be included. *)
680 fun string_of_equality (typ,terms) =
681 let val [tstr1,tstr2] = map string_of_term terms
683 if !keep_types andalso !special_equal
684 then "equal(" ^ (wrap_eq_type typ tstr1) ^ "," ^
685 (wrap_eq_type typ tstr2) ^ ")"
686 else "equal(" ^ tstr1 ^ "," ^ tstr2 ^ ")"
688 and string_of_term (UVar(x,_)) = x
689 | string_of_term (Fun("equal",[typ],terms)) = string_of_equality(typ,terms)
690 | string_of_term (Fun (name,typs,[])) = name (*Overloaded consts like 0 don't get types!*)
691 | string_of_term (Fun (name,typs,terms)) =
692 let val terms_as_strings = map string_of_term terms
693 val typs' = if !keep_types then typs else []
694 in name ^ (paren_pack (terms_as_strings @ typs')) end
695 | string_of_term _ = error "string_of_term";
697 (* before output the string of the predicate, check if the predicate corresponds to an equality or not. *)
698 fun string_of_predicate (Predicate("equal",[typ],terms)) = string_of_equality(typ,terms)
699 | string_of_predicate (Predicate(name,typs,terms)) =
700 let val terms_as_strings = map string_of_term terms
701 val typs' = if !keep_types then typs else []
702 in name ^ (paren_pack (terms_as_strings @ typs')) end
703 | string_of_predicate _ = error "string_of_predicate";
706 fun string_of_clausename (cls_id,ax_name) =
707 clause_prefix ^ ascii_of ax_name ^ "_" ^ Int.toString cls_id;
709 fun string_of_type_clsname (cls_id,ax_name,idx) =
710 string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx);
713 (********************************)
714 (* Code for producing DFG files *)
715 (********************************)
717 fun dfg_literal (Literal(pol,pred,tag)) =
718 let val pred_string = string_of_predicate pred
720 if pol then pred_string else "not(" ^pred_string ^ ")"
724 (* FIX: what does this mean? *)
725 (*fun dfg_of_typeLit (LTVar x) = "not(" ^ x ^ ")"
726 | dfg_of_typeLit (LTFree x) = "(" ^ x ^ ")";*)
728 fun dfg_of_typeLit (LTVar x) = x
729 | dfg_of_typeLit (LTFree x) = x ;
731 (*Make the string of universal quantifiers for a clause*)
732 fun forall_open ([],[]) = ""
733 | forall_open (vars,tvars) = "forall([" ^ (commas (tvars@vars))^ "],\n"
735 fun forall_close ([],[]) = ""
736 | forall_close (vars,tvars) = ")"
738 fun gen_dfg_cls (cls_id,ax_name,knd,lits,tvars,vars) =
739 "clause( %(" ^ knd ^ ")\n" ^ forall_open(vars,tvars) ^
740 "or(" ^ lits ^ ")" ^ forall_close(vars,tvars) ^ ",\n" ^
741 string_of_clausename (cls_id,ax_name) ^ ").";
743 fun gen_dfg_type_cls (cls_id,ax_name,knd,tfree_lit,idx,tvars,vars) =
744 "clause( %(" ^ knd ^ ")\n" ^ forall_open(vars,tvars) ^
745 "or( " ^ tfree_lit ^ ")" ^ forall_close(vars,tvars) ^ ",\n" ^
746 string_of_type_clsname (cls_id,ax_name,idx) ^ ").";
748 fun dfg_clause_aux (Clause cls) =
749 let val lits = map dfg_literal (#literals cls)
751 if !keep_types then map dfg_of_typeLit (#tvar_type_literals cls)
754 if !keep_types then map dfg_of_typeLit (#tfree_type_literals cls)
757 (tvar_lits_strs @ lits, tfree_lits)
761 fun dfg_folterms (Literal(pol,pred,tag)) =
762 let val Predicate (predname, _, folterms) = pred
766 fun get_uvars (UVar(a,typ)) = [a]
767 | get_uvars (Fun (_,typ,tlist)) = union_all(map get_uvars tlist)
770 fun is_uvar (UVar _) = true
771 | is_uvar (Fun _) = false;
773 fun uvar_name (UVar(a,_)) = a
774 | uvar_name (Fun (a,_,_)) = raise CLAUSE("Not a variable", Const(a,dummyT));
776 fun dfg_vars (Clause cls) =
777 let val lits = #literals cls
778 val folterms = List.concat (map dfg_folterms lits)
780 union_all(map get_uvars folterms)
784 fun dfg_tvars (Clause cls) =(#tvars cls)
788 (* make this return funcs and preds too? *)
789 fun string_of_predname (Predicate("equal",_,terms)) = "EQUALITY"
790 | string_of_predname (Predicate(name,_,terms)) = name
794 fun concat_with sep [] = ""
795 | concat_with sep [x] = "(" ^ x ^ ")"
796 | concat_with sep (x::xs) = "(" ^ x ^ ")" ^ sep ^ (concat_with sep xs);
798 fun dfg_pred (Literal(pol,pred,tag)) ax_name =
799 (string_of_predname pred) ^ " " ^ ax_name
802 let val (lits,tfree_lits) = dfg_clause_aux cls
803 (*"lits" includes the typing assumptions (TVars)*)
804 val vars = dfg_vars cls
805 val tvars = dfg_tvars cls
806 val knd = string_of_kind cls
807 val lits_str = commas lits
808 val cls_id = get_clause_id cls
809 val axname = get_axiomName cls
810 val cls_str = gen_dfg_cls(cls_id,axname,knd,lits_str,tvars, vars)
811 fun typ_clss k [] = []
812 | typ_clss k (tfree :: tfrees) =
813 (gen_dfg_type_cls(cls_id,axname,knd,tfree,k, tvars,vars)) ::
814 (typ_clss (k+1) tfrees)
816 cls_str :: (typ_clss 0 tfree_lits)
819 fun string_of_arity (name, num) = name ^ "," ^ (Int.toString num)
821 fun string_of_preds preds =
822 "predicates[" ^ (concat_with ", " (map string_of_arity preds)) ^ "].\n";
824 fun string_of_funcs funcs =
825 "functions[" ^ (concat_with ", " (map string_of_arity funcs)) ^ "].\n" ;
828 fun string_of_symbols predstr funcstr =
829 "list_of_symbols.\n" ^ predstr ^ funcstr ^ "end_of_list.\n\n";
832 fun string_of_axioms axstr =
833 "list_of_clauses(axioms,cnf).\n" ^ axstr ^ "end_of_list.\n\n";
836 fun string_of_conjectures conjstr =
837 "list_of_clauses(conjectures,cnf).\n" ^ conjstr ^ "end_of_list.\n\n";
839 fun string_of_descrip () =
840 "list_of_descriptions.\nname({*[ File : ],[ Names :]*}).\nauthor({*[ Source :]*}).\nstatus(unknown).\ndescription({*[ Refs :]*}).\nend_of_list.\n\n"
843 fun string_of_start name = "%------------------------------------------------------------------------------\nbegin_problem(" ^ name ^ ").\n\n";
846 fun string_of_end () = "end_problem.\n%------------------------------------------------------------------------------";
850 let val (lits,tfree_lits) = dfg_clause_aux cls
851 (*"lits" includes the typing assumptions (TVars)*)
852 val cls_id = get_clause_id cls
853 val ax_name = get_axiomName cls
854 val vars = dfg_vars cls
855 val tvars = dfg_tvars cls
856 val funcs = funcs_of_cls cls
857 val preds = preds_of_cls cls
858 val knd = string_of_kind cls
859 val lits_str = commas lits
860 val cls_str = gen_dfg_cls(cls_id,ax_name,knd,lits_str,tvars,vars)
867 fun tfree_dfg_clause tfree_lit =
868 "clause( %(conjecture)\n" ^ "or( " ^ tfree_lit ^ "),\n" ^ "tfree_tcs" ^ ")."
871 fun gen_dfg_file probname axioms conjectures funcs preds =
872 let val axstrs_tfrees = (map clause2dfg axioms)
873 val (axstrs, atfrees) = ListPair.unzip axstrs_tfrees
874 val axstr = (space_implode "\n" axstrs) ^ "\n\n"
875 val conjstrs_tfrees = (map clause2dfg conjectures)
876 val (conjstrs, atfrees) = ListPair.unzip conjstrs_tfrees
877 val tfree_clss = map tfree_dfg_clause (union_all atfrees)
878 val conjstr = (space_implode "\n" (tfree_clss@conjstrs)) ^ "\n\n"
879 val funcstr = string_of_funcs funcs
880 val predstr = string_of_preds preds
882 (string_of_start probname) ^ (string_of_descrip ()) ^
883 (string_of_symbols funcstr predstr) ^
884 (string_of_axioms axstr) ^
885 (string_of_conjectures conjstr) ^ (string_of_end ())
888 fun clauses2dfg [] probname axioms conjectures funcs preds =
889 let val funcs' = (union_all(map funcs_of_cls axioms)) @ funcs
890 val preds' = (union_all(map preds_of_cls axioms)) @ preds
892 gen_dfg_file probname axioms conjectures funcs' preds'
894 | clauses2dfg (cls::clss) probname axioms conjectures funcs preds =
895 let val (lits,tfree_lits) = dfg_clause_aux cls
896 (*"lits" includes the typing assumptions (TVars)*)
897 val cls_id = get_clause_id cls
898 val ax_name = get_axiomName cls
899 val vars = dfg_vars cls
900 val tvars = dfg_tvars cls
901 val funcs' = (funcs_of_cls cls) union funcs
902 val preds' = (preds_of_cls cls) union preds
903 val knd = string_of_kind cls
904 val lits_str = concat_with ", " lits
905 val axioms' = if knd = "axiom" then (cls::axioms) else axioms
907 if knd = "conjecture" then (cls::conjectures) else conjectures
909 clauses2dfg clss probname axioms' conjectures' funcs' preds'
913 fun string_of_arClauseID (ArityClause {clause_id,axiom_name,...}) =
914 arclause_prefix ^ ascii_of axiom_name ^ "_" ^ Int.toString clause_id;
916 fun string_of_arKind (ArityClause arcls) = name_of_kind(#kind arcls);
918 (*FIXME!!! currently is TPTP format!*)
919 fun dfg_of_arLit (TConsLit(b,(c,t,args))) =
920 let val pol = if b then "++" else "--"
921 val arg_strs = paren_pack args
923 pol ^ c ^ "(" ^ t ^ arg_strs ^ ")"
925 | dfg_of_arLit (TVarLit(b,(c,str))) =
926 let val pol = if b then "++" else "--"
928 pol ^ c ^ "(" ^ str ^ ")"
932 fun dfg_of_conclLit (ArityClause arcls) = dfg_of_arLit (#conclLit arcls);
935 fun dfg_of_premLits (ArityClause arcls) = map dfg_of_arLit (#premLits arcls);
939 (*FIXME: would this have variables in a forall? *)
941 fun dfg_arity_clause arcls =
942 let val arcls_id = string_of_arClauseID arcls
943 val concl_lit = dfg_of_conclLit arcls
944 val prems_lits = dfg_of_premLits arcls
945 val knd = string_of_arKind arcls
946 val all_lits = concl_lit :: prems_lits
948 "clause( %(" ^ knd ^ ")\n" ^ "or( " ^ (bracket_pack all_lits) ^ ")),\n" ^
953 (********************************)
954 (* code to produce TPTP files *)
955 (********************************)
957 fun tptp_literal (Literal(pol,pred,tag)) =
958 let val pred_string = string_of_predicate pred
960 if (tag andalso !tagged) then (if pol then "+++" else "---")
961 else (if pol then "++" else "--")
963 tagged_pol ^ pred_string
966 fun tptp_of_typeLit (LTVar x) = "--" ^ x
967 | tptp_of_typeLit (LTFree x) = "++" ^ x;
970 fun gen_tptp_cls (cls_id,ax_name,knd,lits) =
971 "input_clause(" ^ string_of_clausename (cls_id,ax_name) ^ "," ^
972 knd ^ "," ^ lits ^ ").";
974 fun gen_tptp_type_cls (cls_id,ax_name,knd,tfree_lit,idx) =
975 "input_clause(" ^ string_of_type_clsname (cls_id,ax_name,idx) ^ "," ^
976 knd ^ ",[" ^ tfree_lit ^ "]).";
978 fun tptp_type_lits (Clause cls) =
979 let val lits = map tptp_literal (#literals cls)
982 then (map tptp_of_typeLit (#tvar_type_literals cls))
986 then (map tptp_of_typeLit (#tfree_type_literals cls))
989 (tvar_lits_strs @ lits, tfree_lits)
992 fun tptp_clause cls =
993 let val (lits,tfree_lits) = tptp_type_lits cls
994 (*"lits" includes the typing assumptions (TVars)*)
995 val cls_id = get_clause_id cls
996 val ax_name = get_axiomName cls
997 val knd = string_of_kind cls
998 val lits_str = bracket_pack lits
999 val cls_str = gen_tptp_cls(cls_id,ax_name,knd,lits_str)
1000 fun typ_clss k [] = []
1001 | typ_clss k (tfree :: tfrees) =
1002 gen_tptp_type_cls(cls_id,ax_name,knd,tfree,k) ::
1003 typ_clss (k+1) tfrees
1005 cls_str :: (typ_clss 0 tfree_lits)
1008 fun clause2tptp cls =
1009 let val (lits,tfree_lits) = tptp_type_lits cls
1010 (*"lits" includes the typing assumptions (TVars)*)
1011 val cls_id = get_clause_id cls
1012 val ax_name = get_axiomName cls
1013 val knd = string_of_kind cls
1014 val lits_str = bracket_pack lits
1015 val cls_str = gen_tptp_cls(cls_id,ax_name,knd,lits_str)
1017 (cls_str,tfree_lits)
1021 fun tfree_clause tfree_lit =
1022 "input_clause(" ^ "tfree_tcs," ^ "conjecture" ^ ",[" ^ tfree_lit ^ "]).";
1025 fun tptp_of_arLit (TConsLit(b,(c,t,args))) =
1026 let val pol = if b then "++" else "--"
1027 val arg_strs = paren_pack args
1029 pol ^ c ^ "(" ^ t ^ arg_strs ^ ")"
1031 | tptp_of_arLit (TVarLit(b,(c,str))) =
1032 let val pol = if b then "++" else "--"
1034 pol ^ c ^ "(" ^ str ^ ")"
1038 fun tptp_of_conclLit (ArityClause arcls) = tptp_of_arLit (#conclLit arcls);
1040 fun tptp_of_premLits (ArityClause arcls) = map tptp_of_arLit (#premLits arcls);
1042 fun tptp_arity_clause arcls =
1043 let val arcls_id = string_of_arClauseID arcls
1044 val concl_lit = tptp_of_conclLit arcls
1045 val prems_lits = tptp_of_premLits arcls
1046 val knd = string_of_arKind arcls
1047 val all_lits = concl_lit :: prems_lits
1049 "input_clause(" ^ arcls_id ^ "," ^ knd ^ "," ^
1050 (bracket_pack all_lits) ^ ")."
1053 fun tptp_classrelLits sub sup =
1054 let val tvar = "(T)"
1056 case sup of NONE => "[++" ^ sub ^ tvar ^ "]"
1057 | (SOME supcls) => "[--" ^ sub ^ tvar ^ ",++" ^ supcls ^ tvar ^ "]"
1061 fun tptp_classrelClause (ClassrelClause {clause_id,subclass,superclass,...}) =
1062 let val relcls_id = clrelclause_prefix ^ ascii_of subclass ^ "_" ^
1063 Int.toString clause_id
1064 val lits = tptp_classrelLits (make_type_class subclass)
1065 (Option.map make_type_class superclass)
1067 "input_clause(" ^ relcls_id ^ ",axiom," ^ lits ^ ")."