-- before converting axiom and conjecture clauses into ResClause.clause format, perform "check_is_fol_term" first.
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
74 structure ResClause: RES_CLAUSE =
77 (* Added for typed equality *)
78 val special_equal = ref false; (* by default,equality does not carry type information *)
79 val eq_typ_wrapper = "typeinfo"; (* default string *)
82 val schematic_var_prefix = "V_";
83 val fixed_var_prefix = "v_";
85 val tvar_prefix = "T_";
86 val tfree_prefix = "t_";
88 val clause_prefix = "cls_";
89 val arclause_prefix = "clsarity_"
90 val clrelclause_prefix = "clsrel_";
92 val const_prefix = "c_";
93 val tconst_prefix = "tc_";
95 val class_prefix = "class_";
98 fun union_all xss = foldl (op union) [] xss;
101 (*Provide readable names for the more common symbolic functions*)
102 val const_trans_table =
103 Symtab.make [("op =", "equal"),
104 ("op <=", "lessequals"),
111 ("op -->", "implies"),
115 ("op Int", "inter")];
117 val type_const_trans_table =
118 Symtab.make [("*", "t_prod"),
122 (*Escaping of special characters.
123 Alphanumeric characters are left unchanged.
124 The character _ goes to __
125 Characters in the range ASCII space to / go to _A to _P, respectively.
126 Other printing characters go to _NNN where NNN is the decimal ASCII code.*)
129 val A_minus_space = Char.ord #"A" - Char.ord #" ";
132 if Char.isAlphaNum c then String.str c
133 else if c = #"_" then "__"
134 else if #" " <= c andalso c <= #"/"
135 then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
136 else if Char.isPrint c then ("_" ^ Int.toString (Char.ord c))
141 val ascii_of = String.translate ascii_of_c;
145 (* convert a list of strings into one single string; surrounded by brackets *)
146 fun paren_pack strings = "(" ^ commas strings ^ ")";
148 fun bracket_pack strings = "[" ^ commas strings ^ "]";
151 (*Remove the initial ' character from a type variable, if it is present*)
152 fun trim_type_var s =
153 if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
154 else error ("trim_type: Malformed type variable encountered: " ^ s);
156 fun ascii_of_indexname (v,0) = ascii_of v
157 | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
159 fun make_schematic_var v = schematic_var_prefix ^ (ascii_of_indexname v);
160 fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x);
162 (*Type variables contain _H because the character ' translates to that.*)
163 fun make_schematic_type_var (x,i) =
164 tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
165 fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
167 fun make_fixed_const c =
168 case Symtab.lookup const_trans_table c of
170 | NONE => const_prefix ^ ascii_of c;
172 fun make_fixed_type_const c =
173 case Symtab.lookup type_const_trans_table c of
175 | NONE => tconst_prefix ^ ascii_of c;
177 fun make_type_class clas = class_prefix ^ ascii_of clas;
181 (***** definitions and functions for FOL clauses, prepared for conversion into TPTP format or SPASS format. *****)
183 val keep_types = ref true;
185 datatype kind = Axiom | Hypothesis | Conjecture;
186 fun name_of_kind Axiom = "axiom"
187 | name_of_kind Hypothesis = "hypothesis"
188 | name_of_kind Conjecture = "conjecture";
190 type clause_id = int;
191 type axiom_name = string;
194 type polarity = bool;
196 type indexname = Term.indexname;
199 (* "tag" is used for vampire specific syntax *)
203 (**** Isabelle FOL clauses ****)
205 val tagged = ref false;
207 type pred_name = string;
208 type sort = Term.sort;
209 type fol_type = string;
212 datatype type_literal = LTVar of string | LTFree of string;
215 datatype folTerm = UVar of string * fol_type
216 | Fun of string * fol_type * folTerm list;
217 datatype predicate = Predicate of pred_name * fol_type * 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;
287 (*Declarations of the current theory--to allow suppressing types.*)
288 val monomorphic = ref (fn (_: string) => false);
289 fun no_types_needed s = ! monomorphic s;
291 (*Initialize the type suppression mechanism with the current theory before
292 producing any clauses!*)
293 fun init thy = (monomorphic := Sign.const_monomorphic thy);
296 (*Flatten a type to a string while accumulating sort constraints on the TFress and
298 fun type_of (Type (a, [])) =
299 let val t = make_fixed_type_const a
300 in (t,([],[(t,0)])) end
301 | type_of (Type (a, Ts)) =
302 let val foltyps_ts = map type_of Ts
303 val (folTyps,ts_funcs) = ListPair.unzip foltyps_ts
304 val (ts, funcslist) = ListPair.unzip ts_funcs
305 val ts' = union_all ts
306 val funcs' = union_all funcslist
307 val t = make_fixed_type_const a
309 ((t ^ paren_pack folTyps), (ts', (t, length Ts)::funcs'))
311 | type_of (TFree (a, s)) =
312 let val t = make_fixed_type_var a
313 in (t, ([((FOLTFree a),s)],[(t,0)])) end
314 | type_of (TVar (v, s)) = (make_schematic_type_var v, ([((FOLTVar v),s)], []))
317 fun maybe_type_of c T =
318 if no_types_needed c then ("",([],[])) else type_of T;
320 (* Any variables created via the METAHYPS tactical should be treated as
321 universal vars, although it is represented as "Free(...)" by Isabelle *)
322 val isMeta = String.isPrefix "METAHYP1_"
324 fun pred_name_type (Const(c,T)) =
325 let val (typof,(folTyps,funcs)) = maybe_type_of c T
326 in (make_fixed_const c, (typof,folTyps), funcs) end
327 | pred_name_type (Free(x,T)) =
328 if isMeta x then raise CLAUSE("Predicate Not First Order 1", Free(x,T))
329 else (make_fixed_var x, ("",[]), [])
330 | pred_name_type (v as Var _) = raise CLAUSE("Predicate Not First Order 2", v)
331 | pred_name_type t = raise CLAUSE("Predicate input unexpected", t);
334 (* For type equality *)
335 (* here "arg_typ" is the type of "="'s argument's type, not the type of the equality *)
336 (* Find type of equality arg *)
337 fun eq_arg_type (Type("fun",[T,_])) =
338 let val (folT,_) = type_of T;
341 fun fun_name_type (Const(c,T)) args =
342 let val t = make_fixed_const c
343 val (typof, (folTyps,funcs)) = maybe_type_of c T
344 val arity = if !keep_types andalso not (no_types_needed c)
348 (t, (typof,folTyps), ((t,arity)::funcs))
350 | fun_name_type (Free(x,T)) args =
351 let val t = make_fixed_var x
353 (t, ("",[]), [(t, length args)])
355 | fun_name_type f args = raise CLAUSE("Function Not First Order 1", f);
358 fun term_of (Var(ind_nm,T)) =
359 let val (folType,(ts,funcs)) = type_of T
361 (UVar(make_schematic_var ind_nm, folType), (ts, funcs))
363 | term_of (Free(x,T)) =
364 let val (folType, (ts,funcs)) = type_of T
366 if isMeta x then (UVar(make_schematic_var(x,0),folType),
367 (ts, ((make_schematic_var(x,0)),0)::funcs))
369 (Fun(make_fixed_var x, folType, []),
370 (ts, ((make_fixed_var x),0)::funcs))
372 | term_of (Const(c,T)) = (* impossible to be equality *)
373 let val (folType,(ts,funcs)) = type_of T
375 (Fun(make_fixed_const c, folType, []),
376 (ts, ((make_fixed_const c),0)::funcs))
378 | term_of (app as (t $ a)) =
379 let val (f,args) = strip_comb app
381 let val (funName,(funType,ts1),funcs) = fun_name_type f args
382 val (args',ts_funcs) = ListPair.unzip (map term_of args)
383 val (ts2,funcs') = ListPair.unzip ts_funcs
384 val ts3 = union_all (ts1::ts2)
385 val funcs'' = union_all(funcs::funcs')
387 (Fun(funName,funType,args'), (ts3,funcs''))
389 fun term_of_eq ((Const ("op =", typ)),args) =
390 let val arg_typ = eq_arg_type typ
391 val (args',ts_funcs) = ListPair.unzip (map term_of args)
392 val (ts,funcs) = ListPair.unzip ts_funcs
393 val equal_name = make_fixed_const ("op =")
395 (Fun(equal_name,arg_typ,args'),
397 (make_fixed_var equal_name, 2):: union_all funcs))
400 case f of Const ("op =", typ) => term_of_eq (f,args)
401 | Const(_,_) => term_of_aux ()
404 then raise CLAUSE("Function Not First Order 2", f)
406 | _ => raise CLAUSE("Function Not First Order 3", f)
408 | term_of t = raise CLAUSE("Function Not First Order 4", t);
411 fun pred_of (Const("op =", typ), args) =
412 let val arg_typ = eq_arg_type typ
413 val (args',ts_funcs) = ListPair.unzip (map term_of args)
414 val (ts,funcs) = ListPair.unzip ts_funcs
415 val equal_name = make_fixed_const "op ="
417 (Predicate(equal_name,arg_typ,args'),
419 [((make_fixed_var equal_name), 2)],
422 | pred_of (pred,args) =
423 let val (predName,(predType,ts1), pfuncs) = pred_name_type pred
424 val (args',ts_funcs) = ListPair.unzip (map term_of args)
425 val (ts2,ffuncs) = ListPair.unzip ts_funcs
426 val ts3 = union_all (ts1::ts2)
427 val ffuncs' = union_all ffuncs
428 val newfuncs = pfuncs union ffuncs'
432 if !keep_types andalso not (no_types_needed c)
437 (Predicate(predName,predType,args'), ts3,
438 [(predName, arity)], newfuncs)
442 (*Treatment of literals, possibly negated or tagged*)
443 fun predicate_of ((Const("Not",_) $ P), polarity, tag) =
444 predicate_of (P, not polarity, tag)
445 | predicate_of ((Const("HOL.tag",_) $ P), polarity, tag) =
446 predicate_of (P, polarity, true)
447 | predicate_of (term,polarity,tag) =
448 (pred_of (strip_comb term), polarity, tag);
450 fun literals_of_term1 args (Const("Trueprop",_) $ P) = literals_of_term1 args P
451 | literals_of_term1 (args as (lits, ts, preds, funcs)) (Const("op |",_) $ P $ Q) =
452 let val (lits', ts', preds', funcs') = literals_of_term1 args P
454 literals_of_term1 (lits', ts', preds' union preds, funcs' union funcs) Q
456 | literals_of_term1 (lits, ts, preds, funcs) P =
457 let val ((pred, ts', preds', funcs'), pol, tag) = predicate_of (P,true,false)
458 val lits' = Literal(pol,pred,tag) :: lits
460 (lits', ts union ts', preds' union preds, funcs' union funcs)
464 val literals_of_term = literals_of_term1 ([],[],[],[]);
467 (* FIX: not sure what to do with these funcs *)
469 (*Make literals for sorted type variables*)
470 fun sorts_on_typs (_, []) = ([])
471 | sorts_on_typs (v, "HOL.type" :: s) =
472 sorts_on_typs (v,s) (*Ignore sort "type"*)
473 | sorts_on_typs ((FOLTVar indx), (s::ss)) =
474 LTVar((make_type_class s) ^
475 "(" ^ (make_schematic_type_var indx) ^ ")") ::
476 (sorts_on_typs ((FOLTVar indx), ss))
477 | sorts_on_typs ((FOLTFree x), (s::ss)) =
478 LTFree((make_type_class s) ^ "(" ^ (make_fixed_type_var x) ^ ")") ::
479 (sorts_on_typs ((FOLTFree x), ss));
482 (*UGLY: seems to be parsing the "show sorts" output, removing anything that
483 starts with a left parenthesis.*)
484 fun remove_type str = hd (String.fields (fn c => c = #"(") str);
486 fun pred_of_sort (LTVar x) = ((remove_type x),1)
487 | pred_of_sort (LTFree x) = ((remove_type x),1)
492 (*Given a list of sorted type variables, return two separate lists.
493 The first is for TVars, the second for TFrees.*)
494 fun add_typs_aux [] preds = ([],[], preds)
495 | add_typs_aux ((FOLTVar indx,s)::tss) preds =
496 let val vs = sorts_on_typs (FOLTVar indx, s)
497 val preds' = (map pred_of_sort vs)@preds
498 val (vss,fss, preds'') = add_typs_aux tss preds'
500 (vs union vss, fss, preds'')
502 | add_typs_aux ((FOLTFree x,s)::tss) preds =
503 let val fs = sorts_on_typs (FOLTFree x, s)
504 val preds' = (map pred_of_sort fs)@preds
505 val (vss,fss, preds'') = add_typs_aux tss preds'
507 (vss, fs union fss, preds'')
510 fun add_typs_aux2 [] = ([],[])
511 | add_typs_aux2 ((FOLTVar indx,s)::tss) =
512 let val vs = sorts_on_typs (FOLTVar indx,s)
513 val (vss,fss) = add_typs_aux2 tss
517 | add_typs_aux2 ((FOLTFree x,s)::tss) =
518 let val fs = sorts_on_typs (FOLTFree x,s)
519 val (vss,fss) = add_typs_aux2 tss
525 fun add_typs (Clause cls) preds = add_typs_aux (#types_sorts cls) preds
528 (** make axiom clauses, hypothesis clauses and conjecture clauses. **)
530 fun get_tvar_strs [] = []
531 | get_tvar_strs ((FOLTVar indx,s)::tss) =
532 let val vstr = make_schematic_type_var indx
534 vstr ins (get_tvar_strs tss)
536 | get_tvar_strs((FOLTFree x,s)::tss) = distinct (get_tvar_strs tss)
538 (* FIX add preds and funcs to add typs aux here *)
540 fun make_axiom_clause_thm thm (ax_name,cls_id) =
541 let val (lits,types_sorts, preds, funcs) = literals_of_term (prop_of thm)
542 val (tvar_lits,tfree_lits, preds) = add_typs_aux types_sorts preds
543 val tvars = get_tvar_strs types_sorts
545 make_clause(cls_id,ax_name,Axiom,
546 lits,types_sorts,tvar_lits,tfree_lits,
551 (* check if a clause is FOL first*)
552 fun make_conjecture_clause n t =
553 let val _ = check_is_fol_term t
554 handle TERM("check_is_fol_term",_) => raise CLAUSE("Goal is not FOL",t)
555 val (lits,types_sorts, preds, funcs) = literals_of_term t
556 val (tvar_lits,tfree_lits, preds) = add_typs_aux types_sorts preds
557 val tvars = get_tvar_strs types_sorts
559 make_clause(n,"conjecture",Conjecture,
560 lits,types_sorts,tvar_lits,tfree_lits,
564 fun make_conjecture_clauses_aux _ [] = []
565 | make_conjecture_clauses_aux n (t::ts) =
566 make_conjecture_clause n t :: make_conjecture_clauses_aux (n+1) ts
568 val make_conjecture_clauses = make_conjecture_clauses_aux 0
571 (*before converting an axiom clause to "clause" format, check if it is FOL*)
572 fun make_axiom_clause term (ax_name,cls_id) =
573 let val _ = check_is_fol_term term
574 handle TERM("check_is_fol_term",_) => raise CLAUSE("Axiom is not FOL", term)
575 val (lits,types_sorts, preds,funcs) = literals_of_term term
576 val (tvar_lits,tfree_lits, preds) = add_typs_aux types_sorts preds
577 val tvars = get_tvar_strs types_sorts
579 make_clause(cls_id,ax_name,Axiom,
580 lits,types_sorts,tvar_lits,tfree_lits,
587 (**** Isabelle arities ****)
589 exception ARCLAUSE of string;
596 datatype arLit = TConsLit of bool * (class * tcons * string list) | TVarLit of bool * (class * string);
598 datatype arityClause =
599 ArityClause of {clause_id: clause_id,
600 axiom_name: axiom_name,
603 premLits: arLit list};
607 | get_TVars n = ("T_" ^ (Int.toString n)) :: get_TVars (n-1);
611 fun pack_sort(_,[]) = raise ARCLAUSE("Empty Sort Found")
612 | pack_sort(tvar, [cls]) = [(make_type_class cls, tvar)]
613 | pack_sort(tvar, cls::srt) = (make_type_class cls,tvar) :: (pack_sort(tvar, srt));
616 fun make_TVarLit (b,(cls,str)) = TVarLit(b,(cls,str));
617 fun make_TConsLit (b,(cls,tcons,tvars)) = TConsLit(b,(make_type_class cls,make_fixed_type_const tcons,tvars));
619 fun make_axiom_arity_clause (tcons,n,(res,args)) =
620 let val nargs = length args
621 val tvars = get_TVars nargs
622 val tvars_srts = ListPair.zip (tvars,args)
623 val tvars_srts' = union_all(map pack_sort tvars_srts)
624 val false_tvars_srts' = map (pair false) tvars_srts'
626 ArityClause {clause_id = n, kind = Axiom,
628 conclLit = make_TConsLit(true,(res,tcons,tvars)),
629 premLits = map make_TVarLit false_tvars_srts'}
632 (*The number of clauses generated from cls, including type clauses*)
633 fun num_of_clauses (Clause cls) =
634 let val num_tfree_lits =
635 if !keep_types then length (#tfree_type_literals cls)
637 in 1 + num_tfree_lits end;
640 (**** Isabelle class relations ****)
643 datatype classrelClause =
644 ClassrelClause of {clause_id: clause_id,
646 superclass: class option};
649 fun make_axiom_classrelClause n subclass superclass =
650 ClassrelClause {clause_id = n,
651 subclass = subclass, superclass = superclass};
654 fun classrelClauses_of_aux n sub [] = []
655 | classrelClauses_of_aux n sub (sup::sups) =
656 make_axiom_classrelClause n sub (SOME sup) :: classrelClauses_of_aux (n+1) sub sups;
659 fun classrelClauses_of (sub,sups) =
660 case sups of [] => [make_axiom_classrelClause 0 sub NONE]
661 | _ => classrelClauses_of_aux 0 sub sups;
664 (***** Isabelle arities *****)
667 fun arity_clause _ (tcons, []) = []
668 | arity_clause n (tcons, ar::ars) =
669 make_axiom_arity_clause (tcons,n,ar) ::
670 arity_clause (n+1) (tcons,ars);
672 fun multi_arity_clause [] = []
673 | multi_arity_clause (tcon_ar :: tcons_ars) =
674 arity_clause 0 tcon_ar @ multi_arity_clause tcons_ars
676 fun arity_clause_thy thy =
677 let val arities = #arities (Type.rep_tsig (Sign.tsig_of thy))
678 in multi_arity_clause (Symtab.dest arities) end;
681 (* Isabelle classes *)
683 type classrelClauses = classrelClause list Symtab.table;
685 val classrel_of = #2 o #classes o Type.rep_tsig o Sign.tsig_of;
686 fun classrel_clauses_classrel (C: Sorts.classes) = map classrelClauses_of (Graph.dest C);
687 val classrel_clauses_thy = List.concat o classrel_clauses_classrel o classrel_of;
691 (****!!!! Changed for typed equality !!!!****)
693 fun wrap_eq_type typ t = eq_typ_wrapper ^"(" ^ t ^ "," ^ typ ^ ")";
695 (* Only need to wrap equality's arguments with "typeinfo" if the output clauses are typed && if we specifically ask for types to be included. *)
696 fun string_of_equality (typ,terms) =
697 let val [tstr1,tstr2] = map string_of_term terms
699 if !keep_types andalso !special_equal
700 then "equal(" ^ (wrap_eq_type typ tstr1) ^ "," ^
701 (wrap_eq_type typ tstr2) ^ ")"
702 else "equal(" ^ tstr1 ^ "," ^ tstr2 ^ ")"
704 and string_of_term (UVar(x,_)) = x
705 | string_of_term (Fun("equal",typ,terms)) = string_of_equality(typ,terms)
706 | string_of_term (Fun (name,typ,[])) = name
707 | string_of_term (Fun (name,typ,terms)) =
708 let val terms' = map string_of_term terms
710 if !keep_types andalso typ<>""
711 then name ^ (paren_pack (terms' @ [typ]))
712 else name ^ (paren_pack terms')
715 (* before output the string of the predicate, check if the predicate corresponds to an equality or not. *)
716 fun string_of_predicate (Predicate("equal",typ,terms)) =
717 string_of_equality(typ,terms)
718 | string_of_predicate (Predicate(name,_,[])) = name
719 | string_of_predicate (Predicate(name,typ,terms)) =
720 let val terms_as_strings = map string_of_term terms
722 if !keep_types andalso typ<>""
723 then name ^ (paren_pack (terms_as_strings @ [typ]))
724 else name ^ (paren_pack terms_as_strings)
728 fun string_of_clausename (cls_id,ax_name) =
729 clause_prefix ^ ascii_of ax_name ^ "_" ^ Int.toString cls_id;
731 fun string_of_type_clsname (cls_id,ax_name,idx) =
732 string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx);
735 (********************************)
736 (* Code for producing DFG files *)
737 (********************************)
739 fun dfg_literal (Literal(pol,pred,tag)) =
740 let val pred_string = string_of_predicate pred
742 if pol then pred_string else "not(" ^pred_string ^ ")"
746 (* FIX: what does this mean? *)
747 (*fun dfg_of_typeLit (LTVar x) = "not(" ^ x ^ ")"
748 | dfg_of_typeLit (LTFree x) = "(" ^ x ^ ")";*)
750 fun dfg_of_typeLit (LTVar x) = x
751 | dfg_of_typeLit (LTFree x) = x ;
753 (*Make the string of universal quantifiers for a clause*)
754 fun forall_open ([],[]) = ""
755 | forall_open (vars,tvars) = "forall([" ^ (commas (tvars@vars))^ "],\n"
757 fun forall_close ([],[]) = ""
758 | forall_close (vars,tvars) = ")"
760 fun gen_dfg_cls (cls_id,ax_name,knd,lits,tvars,vars) =
761 "clause( %(" ^ knd ^ ")\n" ^ forall_open(vars,tvars) ^
762 "or(" ^ lits ^ ")" ^ forall_close(vars,tvars) ^ ",\n" ^
763 string_of_clausename (cls_id,ax_name) ^ ").";
765 fun gen_dfg_type_cls (cls_id,ax_name,knd,tfree_lit,idx,tvars,vars) =
766 "clause( %(" ^ knd ^ ")\n" ^ forall_open(vars,tvars) ^
767 "or( " ^ tfree_lit ^ ")" ^ forall_close(vars,tvars) ^ ",\n" ^
768 string_of_type_clsname (cls_id,ax_name,idx) ^ ").";
770 fun dfg_clause_aux (Clause cls) =
771 let val lits = map dfg_literal (#literals cls)
773 if !keep_types then map dfg_of_typeLit (#tvar_type_literals cls)
776 if !keep_types then map dfg_of_typeLit (#tfree_type_literals cls)
779 (tvar_lits_strs @ lits, tfree_lits)
783 fun dfg_folterms (Literal(pol,pred,tag)) =
784 let val Predicate (predname, foltype, folterms) = pred
790 fun get_uvars (UVar(a,typ)) = [a]
791 | get_uvars (Fun (_,typ,tlist)) = union_all(map get_uvars tlist)
794 fun is_uvar (UVar _) = true
795 | is_uvar (Fun _) = false;
797 fun uvar_name (UVar(a,_)) = a
798 | uvar_name (Fun (a,_,_)) = raise CLAUSE("Not a variable", Const(a,dummyT));
800 fun mergelist [] = []
801 | mergelist (x::xs) = x @ mergelist xs
803 fun dfg_vars (Clause cls) =
804 let val lits = #literals cls
805 val folterms = mergelist(map dfg_folterms lits)
807 union_all(map get_uvars folterms)
811 fun dfg_tvars (Clause cls) =(#tvars cls)
815 (* make this return funcs and preds too? *)
816 fun string_of_predname (Predicate("equal",typ,terms)) = "EQUALITY"
817 | string_of_predname (Predicate(name,_,[])) = name
818 | string_of_predname (Predicate(name,typ,terms)) = name
821 (* make this return funcs and preds too? *)
823 fun string_of_predicate (Predicate("equal",typ,terms)) =
824 string_of_equality(typ,terms)
825 | string_of_predicate (Predicate(name,_,[])) = name
826 | string_of_predicate (Predicate(name,typ,terms)) =
827 let val terms_as_strings = map string_of_term terms
829 if !keep_types andalso typ<>""
830 then name ^ (paren_pack (terms_as_strings @ [typ]))
831 else name ^ (paren_pack terms_as_strings)
835 fun concat_with sep [] = ""
836 | concat_with sep [x] = "(" ^ x ^ ")"
837 | concat_with sep (x::xs) = "(" ^ x ^ ")" ^ sep ^ (concat_with sep xs);
839 fun dfg_pred (Literal(pol,pred,tag)) ax_name =
840 (string_of_predname pred) ^ " " ^ ax_name
843 let val (lits,tfree_lits) = dfg_clause_aux cls
844 (*"lits" includes the typing assumptions (TVars)*)
845 val vars = dfg_vars cls
846 val tvars = dfg_tvars cls
847 val knd = string_of_kind cls
848 val lits_str = commas lits
849 val cls_id = get_clause_id cls
850 val axname = get_axiomName cls
851 val cls_str = gen_dfg_cls(cls_id,axname,knd,lits_str,tvars, vars)
852 fun typ_clss k [] = []
853 | typ_clss k (tfree :: tfrees) =
854 (gen_dfg_type_cls(cls_id,axname,knd,tfree,k, tvars,vars)) ::
855 (typ_clss (k+1) tfrees)
857 cls_str :: (typ_clss 0 tfree_lits)
860 fun string_of_arity (name, num) = name ^ "," ^ (Int.toString num)
862 fun string_of_preds preds =
863 "predicates[" ^ (concat_with ", " (map string_of_arity preds)) ^ "].\n";
865 fun string_of_funcs funcs =
866 "functions[" ^ (concat_with ", " (map string_of_arity funcs)) ^ "].\n" ;
869 fun string_of_symbols predstr funcstr =
870 "list_of_symbols.\n" ^ predstr ^ funcstr ^ "end_of_list.\n\n";
873 fun string_of_axioms axstr =
874 "list_of_clauses(axioms,cnf).\n" ^ axstr ^ "end_of_list.\n\n";
877 fun string_of_conjectures conjstr =
878 "list_of_clauses(conjectures,cnf).\n" ^ conjstr ^ "end_of_list.\n\n";
880 fun string_of_descrip () =
881 "list_of_descriptions.\nname({*[ File : ],[ Names :]*}).\nauthor({*[ Source :]*}).\nstatus(unknown).\ndescription({*[ Refs :]*}).\nend_of_list.\n\n"
884 fun string_of_start name = "%------------------------------------------------------------------------------\nbegin_problem(" ^ name ^ ").\n\n";
887 fun string_of_end () = "end_problem.\n%------------------------------------------------------------------------------";
891 let val (lits,tfree_lits) = dfg_clause_aux cls
892 (*"lits" includes the typing assumptions (TVars)*)
893 val cls_id = get_clause_id cls
894 val ax_name = get_axiomName cls
895 val vars = dfg_vars cls
896 val tvars = dfg_tvars cls
897 val funcs = funcs_of_cls cls
898 val preds = preds_of_cls cls
899 val knd = string_of_kind cls
900 val lits_str = commas lits
901 val cls_str = gen_dfg_cls(cls_id,ax_name,knd,lits_str,tvars,vars)
908 fun tfree_dfg_clause tfree_lit =
909 "clause( %(conjecture)\n" ^ "or( " ^ tfree_lit ^ "),\n" ^ "tfree_tcs" ^ ")."
912 fun gen_dfg_file probname axioms conjectures funcs preds =
913 let val axstrs_tfrees = (map clause2dfg axioms)
914 val (axstrs, atfrees) = ListPair.unzip axstrs_tfrees
915 val axstr = (space_implode "\n" axstrs) ^ "\n\n"
916 val conjstrs_tfrees = (map clause2dfg conjectures)
917 val (conjstrs, atfrees) = ListPair.unzip conjstrs_tfrees
918 val tfree_clss = map tfree_dfg_clause (union_all atfrees)
919 val conjstr = (space_implode "\n" (tfree_clss@conjstrs)) ^ "\n\n"
920 val funcstr = string_of_funcs funcs
921 val predstr = string_of_preds preds
923 (string_of_start probname) ^ (string_of_descrip ()) ^
924 (string_of_symbols funcstr predstr) ^
925 (string_of_axioms axstr) ^
926 (string_of_conjectures conjstr) ^ (string_of_end ())
929 fun clauses2dfg [] probname axioms conjectures funcs preds =
930 let val funcs' = (union_all(map funcs_of_cls axioms)) @ funcs
931 val preds' = (union_all(map preds_of_cls axioms)) @ preds
933 gen_dfg_file probname axioms conjectures funcs' preds'
935 | clauses2dfg (cls::clss) probname axioms conjectures funcs preds =
936 let val (lits,tfree_lits) = dfg_clause_aux cls
937 (*"lits" includes the typing assumptions (TVars)*)
938 val cls_id = get_clause_id cls
939 val ax_name = get_axiomName cls
940 val vars = dfg_vars cls
941 val tvars = dfg_tvars cls
942 val funcs' = (funcs_of_cls cls) union funcs
943 val preds' = (preds_of_cls cls) union preds
944 val knd = string_of_kind cls
945 val lits_str = concat_with ", " lits
946 val axioms' = if knd = "axiom" then (cls::axioms) else axioms
948 if knd = "conjecture" then (cls::conjectures) else conjectures
950 clauses2dfg clss probname axioms' conjectures' funcs' preds'
954 fun string_of_arClauseID (ArityClause {clause_id,axiom_name,...}) =
955 arclause_prefix ^ ascii_of axiom_name ^ "_" ^ Int.toString clause_id;
957 fun string_of_arKind (ArityClause arcls) = name_of_kind(#kind arcls);
959 (*FIXME!!! currently is TPTP format!*)
960 fun dfg_of_arLit (TConsLit(b,(c,t,args))) =
961 let val pol = if b then "++" else "--"
962 val arg_strs = (case args of [] => "" | _ => paren_pack args)
964 pol ^ c ^ "(" ^ t ^ arg_strs ^ ")"
966 | dfg_of_arLit (TVarLit(b,(c,str))) =
967 let val pol = if b then "++" else "--"
969 pol ^ c ^ "(" ^ str ^ ")"
973 fun dfg_of_conclLit (ArityClause arcls) = dfg_of_arLit (#conclLit arcls);
976 fun dfg_of_premLits (ArityClause arcls) = map dfg_of_arLit (#premLits arcls);
980 (*FIXME: would this have variables in a forall? *)
982 fun dfg_arity_clause arcls =
983 let val arcls_id = string_of_arClauseID arcls
984 val concl_lit = dfg_of_conclLit arcls
985 val prems_lits = dfg_of_premLits arcls
986 val knd = string_of_arKind arcls
987 val all_lits = concl_lit :: prems_lits
989 "clause( %(" ^ knd ^ ")\n" ^ "or( " ^ (bracket_pack all_lits) ^ ")),\n" ^
994 (********************************)
995 (* code to produce TPTP files *)
996 (********************************)
998 fun tptp_literal (Literal(pol,pred,tag)) =
999 let val pred_string = string_of_predicate pred
1001 if (tag andalso !tagged) then (if pol then "+++" else "---")
1002 else (if pol then "++" else "--")
1004 tagged_pol ^ pred_string
1009 fun tptp_of_typeLit (LTVar x) = "--" ^ x
1010 | tptp_of_typeLit (LTFree x) = "++" ^ x;
1013 fun gen_tptp_cls (cls_id,ax_name,knd,lits) =
1014 "input_clause(" ^ string_of_clausename (cls_id,ax_name) ^ "," ^
1015 knd ^ "," ^ lits ^ ").";
1017 fun gen_tptp_type_cls (cls_id,ax_name,knd,tfree_lit,idx) =
1018 "input_clause(" ^ string_of_type_clsname (cls_id,ax_name,idx) ^ "," ^
1019 knd ^ ",[" ^ tfree_lit ^ "]).";
1021 fun tptp_type_lits (Clause cls) =
1022 let val lits = map tptp_literal (#literals cls)
1023 val tvar_lits_strs =
1025 then (map tptp_of_typeLit (#tvar_type_literals cls))
1029 then (map tptp_of_typeLit (#tfree_type_literals cls))
1032 (tvar_lits_strs @ lits, tfree_lits)
1035 fun tptp_clause cls =
1036 let val (lits,tfree_lits) = tptp_type_lits cls
1037 (*"lits" includes the typing assumptions (TVars)*)
1038 val cls_id = get_clause_id cls
1039 val ax_name = get_axiomName cls
1040 val knd = string_of_kind cls
1041 val lits_str = bracket_pack lits
1042 val cls_str = gen_tptp_cls(cls_id,ax_name,knd,lits_str)
1043 fun typ_clss k [] = []
1044 | typ_clss k (tfree :: tfrees) =
1045 gen_tptp_type_cls(cls_id,ax_name,knd,tfree,k) ::
1046 typ_clss (k+1) tfrees
1048 cls_str :: (typ_clss 0 tfree_lits)
1051 fun clause2tptp cls =
1052 let val (lits,tfree_lits) = tptp_type_lits cls
1053 (*"lits" includes the typing assumptions (TVars)*)
1054 val cls_id = get_clause_id cls
1055 val ax_name = get_axiomName cls
1056 val knd = string_of_kind cls
1057 val lits_str = bracket_pack lits
1058 val cls_str = gen_tptp_cls(cls_id,ax_name,knd,lits_str)
1060 (cls_str,tfree_lits)
1064 fun tfree_clause tfree_lit =
1065 "input_clause(" ^ "tfree_tcs," ^ "conjecture" ^ ",[" ^ tfree_lit ^ "]).";
1068 fun tptp_of_arLit (TConsLit(b,(c,t,args))) =
1069 let val pol = if b then "++" else "--"
1070 val arg_strs = (case args of [] => "" | _ => paren_pack args)
1072 pol ^ c ^ "(" ^ t ^ arg_strs ^ ")"
1074 | tptp_of_arLit (TVarLit(b,(c,str))) =
1075 let val pol = if b then "++" else "--"
1077 pol ^ c ^ "(" ^ str ^ ")"
1081 fun tptp_of_conclLit (ArityClause arcls) = tptp_of_arLit (#conclLit arcls);
1083 fun tptp_of_premLits (ArityClause arcls) = map tptp_of_arLit (#premLits arcls);
1085 fun tptp_arity_clause arcls =
1086 let val arcls_id = string_of_arClauseID arcls
1087 val concl_lit = tptp_of_conclLit arcls
1088 val prems_lits = tptp_of_premLits arcls
1089 val knd = string_of_arKind arcls
1090 val all_lits = concl_lit :: prems_lits
1092 "input_clause(" ^ arcls_id ^ "," ^ knd ^ "," ^
1093 (bracket_pack all_lits) ^ ")."
1096 fun tptp_classrelLits sub sup =
1097 let val tvar = "(T)"
1099 case sup of NONE => "[++" ^ sub ^ tvar ^ "]"
1100 | (SOME supcls) => "[--" ^ sub ^ tvar ^ ",++" ^ supcls ^ tvar ^ "]"
1104 fun tptp_classrelClause (ClassrelClause {clause_id,subclass,superclass,...}) =
1105 let val relcls_id = clrelclause_prefix ^ ascii_of subclass ^ "_" ^
1106 Int.toString clause_id
1107 val lits = tptp_classrelLits (make_type_class subclass)
1108 (Option.map make_type_class superclass)
1110 "input_clause(" ^ relcls_id ^ ",axiom," ^ lits ^ ")."