src/HOL/Tools/res_clause.ML
author mengj
Wed, 19 Oct 2005 13:59:33 +0200
changeset 17908 ac97527724ba
parent 17888 116a8d1c7a67
child 17993 e6e5b28740ec
permissions -rw-r--r--
More functions are added to the signature of ResClause
     1 (*  Author: Jia Meng, Cambridge University Computer Laboratory
     2 
     3     ID: $Id$
     4     Copyright 2004 University of Cambridge
     5 
     6 ML data structure for storing/printing FOL clauses and arity clauses.
     7 Typed equality is treated differently.
     8 *)
     9 
    10 (* works for writeoutclasimp on typed *)
    11 signature RES_CLAUSE =
    12   sig
    13   val keep_types : bool ref
    14   val special_equal : bool ref
    15   val tagged : bool ref
    16 
    17   exception ARCLAUSE of string
    18   exception CLAUSE of string * term
    19   type arityClause 
    20   type classrelClause
    21   type clause
    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
    28 
    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
    33 
    34   val arity_clause_thy: theory -> arityClause list 
    35   val classrel_clauses_thy: theory -> classrelClause list 
    36 
    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 
    51 
    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 
    64   end;
    65 
    66 structure ResClause: RES_CLAUSE =
    67 struct
    68 
    69 (* Added for typed equality *)
    70 val special_equal = ref false; (* by default,equality does not carry type information *)
    71 val eq_typ_wrapper = "typeinfo"; (* default string *)
    72 
    73 
    74 val schematic_var_prefix = "V_";
    75 val fixed_var_prefix = "v_";
    76 
    77 val tvar_prefix = "T_";
    78 val tfree_prefix = "t_";
    79 
    80 val clause_prefix = "cls_"; 
    81 val arclause_prefix = "clsarity_" 
    82 val clrelclause_prefix = "clsrel_";
    83 
    84 val const_prefix = "c_";
    85 val tconst_prefix = "tc_"; 
    86 
    87 val class_prefix = "class_"; 
    88 
    89 
    90 fun union_all xss = foldl (op union) [] xss;
    91 
    92  
    93 (*Provide readable names for the more common symbolic functions*)
    94 val const_trans_table =
    95       Symtab.make [("op =", "equal"),
    96 	  	   ("op <=", "lessequals"),
    97 		   ("op <", "less"),
    98 		   ("op &", "and"),
    99 		   ("op |", "or"),
   100 		   ("op +", "plus"),
   101 		   ("op -", "minus"),
   102 		   ("op *", "times"),
   103 		   ("op -->", "implies"),
   104 		   ("{}", "emptyset"),
   105 		   ("op :", "in"),
   106 		   ("op Un", "union"),
   107 		   ("op Int", "inter")];
   108 
   109 val type_const_trans_table =
   110       Symtab.make [("*", "t_prod"),
   111 	  	   ("+", "t_sum"),
   112 		   ("~=>", "t_map")];
   113 
   114 (*Escaping of special characters.
   115   Alphanumeric characters are left unchanged.
   116   The character _ goes to __
   117   Characters in the range ASCII space to / go to _A to _P, respectively.
   118   Other printing characters go to _NNN where NNN is the decimal ASCII code.*)
   119 local
   120 
   121 val A_minus_space = Char.ord #"A" - Char.ord #" ";
   122 
   123 fun ascii_of_c c =
   124   if Char.isAlphaNum c then String.str c
   125   else if c = #"_" then "__"
   126   else if #" " <= c andalso c <= #"/" 
   127        then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
   128   else if Char.isPrint c then ("_" ^ Int.toString (Char.ord c))
   129   else ""
   130 
   131 in
   132 
   133 val ascii_of = String.translate ascii_of_c;
   134 
   135 end;
   136 
   137 (* convert a list of strings into one single string; surrounded by brackets *)
   138 fun paren_pack strings = "(" ^ commas strings ^ ")";
   139 
   140 fun bracket_pack strings = "[" ^ commas strings ^ "]";
   141 
   142 
   143 (*Remove the initial ' character from a type variable, if it is present*)
   144 fun trim_type_var s =
   145   if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
   146   else error ("trim_type: Malformed type variable encountered: " ^ s);
   147 
   148 fun ascii_of_indexname (v,0) = ascii_of v
   149   | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
   150 
   151 fun make_schematic_var v = schematic_var_prefix ^ (ascii_of_indexname v);
   152 fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x);
   153 
   154 (*Type variables contain _H because the character ' translates to that.*)
   155 fun make_schematic_type_var (x,i) = 
   156       tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
   157 fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
   158 
   159 fun make_fixed_const c =
   160     case Symtab.lookup const_trans_table c of
   161         SOME c' => c'
   162       | NONE =>  const_prefix ^ ascii_of c;
   163 
   164 fun make_fixed_type_const c = 
   165     case Symtab.lookup type_const_trans_table c of
   166         SOME c' => c'
   167       | NONE =>  tconst_prefix ^ ascii_of c;
   168 
   169 fun make_type_class clas = class_prefix ^ ascii_of clas;
   170 
   171 
   172 
   173 (***** definitions and functions for FOL clauses, prepared for conversion into TPTP format or SPASS format. *****)
   174 
   175 val keep_types = ref true;
   176 
   177 datatype kind = Axiom | Hypothesis | Conjecture;
   178 fun name_of_kind Axiom = "axiom"
   179   | name_of_kind Hypothesis = "hypothesis"
   180   | name_of_kind Conjecture = "conjecture";
   181 
   182 type clause_id = int;
   183 type axiom_name = string;
   184 
   185 
   186 type polarity = bool;
   187 
   188 type indexname = Term.indexname;
   189 
   190 
   191 (* "tag" is used for vampire specific syntax  *)
   192 type tag = bool; 
   193 
   194 
   195 (**** Isabelle FOL clauses ****)
   196 
   197 val tagged = ref false;
   198 
   199 type pred_name = string;
   200 type sort = Term.sort;
   201 type fol_type = string;
   202 
   203 
   204 datatype type_literal = LTVar of string | LTFree of string;
   205 
   206 
   207 datatype folTerm = UVar of string * fol_type
   208                  | Fun of string * fol_type * folTerm list;
   209 datatype predicate = Predicate of pred_name * fol_type * folTerm list;
   210 
   211 datatype literal = Literal of polarity * predicate * tag;
   212 
   213 datatype typ_var = FOLTVar of indexname | FOLTFree of string;
   214 
   215 
   216 (* ML datatype used to repsent one single clause: disjunction of literals. *)
   217 datatype clause = 
   218 	 Clause of {clause_id: clause_id,
   219 		    axiom_name: axiom_name,
   220 		    kind: kind,
   221 		    literals: literal list,
   222 		    types_sorts: (typ_var * sort) list, 
   223                     tvar_type_literals: type_literal list, 
   224                     tfree_type_literals: type_literal list ,
   225                     tvars: string list,
   226                     predicates: (string*int) list,
   227                     functions: (string*int) list};
   228 
   229 
   230 exception CLAUSE of string * term;
   231 
   232 
   233 (*** make clauses ***)
   234 
   235 fun isFalse (Literal (pol,Predicate(a,_,[]),_)) =
   236       (pol andalso a = "c_False") orelse
   237       (not pol andalso a = "c_True")
   238   | isFalse _ = false;
   239 
   240 fun isTrue (Literal (pol,Predicate(a,_,[]),_)) =
   241       (pol andalso a = "c_True") orelse
   242       (not pol andalso a = "c_False")
   243   | isTrue _ = false;
   244   
   245 fun isTaut (Clause {literals,...}) = exists isTrue literals;  
   246 
   247 fun make_clause (clause_id,axiom_name,kind,literals,
   248                  types_sorts,tvar_type_literals,
   249                  tfree_type_literals,tvars, predicates, functions) =
   250   if forall isFalse literals 
   251   then error "Problem too trivial for resolution (empty clause)"
   252   else
   253      Clause {clause_id = clause_id, axiom_name = axiom_name, kind = kind, 
   254              literals = literals, types_sorts = types_sorts,
   255              tvar_type_literals = tvar_type_literals,
   256              tfree_type_literals = tfree_type_literals,
   257              tvars = tvars, predicates = predicates, 
   258              functions = functions};
   259 
   260 
   261 (** Some Clause destructor functions **)
   262 
   263 fun string_of_kind (Clause cls) = name_of_kind (#kind cls);
   264 
   265 fun get_axiomName (Clause cls) = #axiom_name cls;
   266 
   267 fun get_clause_id (Clause cls) = #clause_id cls;
   268 
   269 fun funcs_of_cls (Clause cls) = #functions cls;
   270 
   271 fun preds_of_cls (Clause cls) = #predicates cls;
   272 
   273 
   274 
   275 (*Definitions of the current theory--to allow suppressing types.*)
   276 val curr_defs = ref Defs.empty;
   277 
   278 (*Initialize the type suppression mechanism with the current theory before
   279     producing any clauses!*)
   280 fun init thy = (curr_defs := Theory.defs_of thy);
   281 
   282 fun no_types_needed s = Defs.monomorphic (!curr_defs) s;
   283     
   284 
   285 (*Flatten a type to a string while accumulating sort constraints on the TFress and
   286   TVars it contains.*)    
   287 fun type_of (Type (a, [])) = 
   288       let val t = make_fixed_type_const a
   289       in (t,([],[(t,0)]))  end
   290   | type_of (Type (a, Ts)) = 
   291       let val foltyps_ts = map type_of Ts 
   292 	  val (folTyps,ts_funcs) = ListPair.unzip foltyps_ts
   293 	  val (ts, funcslist) = ListPair.unzip ts_funcs
   294 	  val ts' = union_all ts
   295 	  val funcs' = union_all funcslist
   296 	  val t = make_fixed_type_const a
   297       in    
   298 	  ((t ^ paren_pack folTyps), (ts', (t, length Ts)::funcs'))
   299       end
   300   | type_of (TFree (a, s)) = 
   301       let val t = make_fixed_type_var a
   302       in (t, ([((FOLTFree a),s)],[(t,0)])) end
   303   | type_of (TVar (v, s)) = (make_schematic_type_var v, ([((FOLTVar v),s)], []))
   304 
   305 
   306 fun maybe_type_of c T =
   307  if no_types_needed c then ("",([],[])) else type_of T;
   308 
   309 (* Any variables created via the METAHYPS tactical should be treated as
   310    universal vars, although it is represented as "Free(...)" by Isabelle *)
   311 val isMeta = String.isPrefix "METAHYP1_"
   312 
   313 fun pred_name_type (Const(c,T)) = 
   314       let val (typof,(folTyps,funcs)) = maybe_type_of c T
   315       in (make_fixed_const c, (typof,folTyps), funcs) end
   316   | pred_name_type (Free(x,T))  = 
   317       if isMeta x then raise CLAUSE("Predicate Not First Order 1", Free(x,T)) 
   318       else (make_fixed_var x, ("",[]), [])
   319   | pred_name_type (v as Var _) = raise CLAUSE("Predicate Not First Order 2", v)
   320   | pred_name_type t        = raise CLAUSE("Predicate input unexpected", t);
   321 
   322 
   323 (* For type equality *)
   324 (* here "arg_typ" is the type of "="'s argument's type, not the type of the equality *)
   325 (* Find type of equality arg *)
   326 fun eq_arg_type (Type("fun",[T,_])) = 
   327     let val (folT,_) = type_of T;
   328     in  folT  end;
   329 
   330 fun fun_name_type (Const(c,T)) args = 
   331       let val t = make_fixed_const c
   332 	val (typof, (folTyps,funcs)) = maybe_type_of c T
   333 	val arity = if !keep_types andalso not (no_types_needed c)
   334 	            then 1 + length args
   335 	            else length args
   336       in
   337 	  (t, (typof,folTyps), ((t,arity)::funcs))
   338       end
   339  | fun_name_type (Free(x,T)) args  = 
   340       let val t = make_fixed_var x
   341       in
   342 	    (t, ("",[]), [(t, length args)])
   343       end
   344   | fun_name_type f args = raise CLAUSE("Function Not First Order 1", f);
   345 
   346 
   347 fun term_of (Var(ind_nm,T)) = 
   348       let val (folType,(ts,funcs)) = type_of T
   349       in
   350 	  (UVar(make_schematic_var ind_nm, folType), (ts, funcs))
   351       end
   352   | term_of (Free(x,T)) = 
   353       let val (folType, (ts,funcs)) = type_of T
   354       in
   355 	  if isMeta x then (UVar(make_schematic_var(x,0),folType),
   356 			    (ts, ((make_schematic_var(x,0)),0)::funcs))
   357 	  else
   358 	      (Fun(make_fixed_var x, folType, []), 
   359 	       (ts, ((make_fixed_var x),0)::funcs))
   360       end
   361   | term_of (Const(c,T)) =  (* impossible to be equality *)
   362       let val (folType,(ts,funcs)) = type_of T
   363       in
   364 	  (Fun(make_fixed_const c, folType, []),
   365 	   (ts, ((make_fixed_const c),0)::funcs))
   366       end    
   367   | term_of (app as (t $ a)) = 
   368       let val (f,args) = strip_comb app
   369 	  fun term_of_aux () = 
   370 	      let val (funName,(funType,ts1),funcs) = fun_name_type f args
   371 		  val (args',ts_funcs) = ListPair.unzip (map term_of args)
   372 		  val (ts2,funcs') = ListPair.unzip ts_funcs
   373 		  val ts3 = union_all (ts1::ts2)
   374 		  val funcs'' = union_all(funcs::funcs')
   375 	      in
   376 		  (Fun(funName,funType,args'), (ts3,funcs''))
   377 	      end
   378 	  fun term_of_eq ((Const ("op =", typ)),args) =
   379 	      let val arg_typ = eq_arg_type typ
   380 		  val (args',ts_funcs) = ListPair.unzip (map term_of args)
   381 		  val (ts,funcs) = ListPair.unzip ts_funcs
   382 		  val equal_name = make_fixed_const ("op =")
   383 	      in
   384 		  (Fun(equal_name,arg_typ,args'),
   385 		   (union_all ts, 
   386 		    (make_fixed_var equal_name, 2):: union_all funcs))
   387 	      end
   388       in
   389 	 case f of Const ("op =", typ) => term_of_eq (f,args)
   390 		 | Const(_,_) => term_of_aux ()
   391 		 | Free(s,_)  => 
   392 		     if isMeta s 
   393 		     then raise CLAUSE("Function Not First Order 2", f)
   394 		     else term_of_aux()
   395 		 | _ => raise CLAUSE("Function Not First Order 3", f)
   396       end
   397   | term_of t = raise CLAUSE("Function Not First Order 4", t); 
   398 
   399 
   400 fun pred_of (Const("op =", typ), args) =
   401       let val arg_typ = eq_arg_type typ 
   402 	  val (args',ts_funcs) = ListPair.unzip (map term_of args)
   403 	  val (ts,funcs) = ListPair.unzip ts_funcs
   404 	  val equal_name = make_fixed_const "op ="
   405       in
   406 	  (Predicate(equal_name,arg_typ,args'),
   407 	   union_all ts, 
   408 	   [((make_fixed_var equal_name), 2)], 
   409 	   union_all funcs)
   410       end
   411   | pred_of (pred,args) = 
   412       let val (predName,(predType,ts1), pfuncs) = pred_name_type pred
   413 	  val (args',ts_funcs) = ListPair.unzip (map term_of args)
   414 	  val (ts2,ffuncs) = ListPair.unzip ts_funcs
   415 	  val ts3 = union_all (ts1::ts2)
   416 	  val ffuncs' = union_all ffuncs
   417 	  val newfuncs = pfuncs union ffuncs'
   418 	  val arity = 
   419 	    case pred of
   420 		Const (c,_) => 
   421 		      if !keep_types andalso not (no_types_needed c)
   422 		      then 1 + length args
   423 		      else length args
   424 	      | _ => length args
   425       in
   426 	  (Predicate(predName,predType,args'), ts3, 
   427 	   [(predName, arity)], newfuncs)
   428       end;
   429 
   430 
   431 (*Treatment of literals, possibly negated or tagged*)
   432 fun predicate_of ((Const("Not",_) $ P), polarity, tag) =
   433       predicate_of (P, not polarity, tag)
   434   | predicate_of ((Const("HOL.tag",_) $ P), polarity, tag) =
   435       predicate_of (P, polarity, true)
   436   | predicate_of (term,polarity,tag) =
   437         (pred_of (strip_comb term), polarity, tag);
   438 
   439 fun literals_of_term1 args (Const("Trueprop",_) $ P) = literals_of_term1 args P
   440   | literals_of_term1 (args as (lits, ts, preds, funcs)) (Const("op |",_) $ P $ Q) = 
   441       let val (lits', ts', preds', funcs') = literals_of_term1 args P
   442       in
   443 	  literals_of_term1 (lits', ts', preds' union preds, funcs' union funcs) Q
   444       end
   445   | literals_of_term1 (lits, ts, preds, funcs) P =
   446       let val ((pred, ts', preds', funcs'), pol, tag) = predicate_of (P,true,false)
   447 	  val lits' = Literal(pol,pred,tag) :: lits
   448       in
   449 	  (lits', ts union ts', preds' union preds, funcs' union funcs)
   450       end;
   451 
   452 
   453 val literals_of_term = literals_of_term1 ([],[],[],[]);
   454 
   455 
   456 (* FIX: not sure what to do with these funcs *)
   457 
   458 (*Make literals for sorted type variables*) 
   459 fun sorts_on_typs (_, [])   = ([]) 
   460   | sorts_on_typs (v, "HOL.type" :: s) =
   461       sorts_on_typs (v,s)   (*Ignore sort "type"*)
   462   | sorts_on_typs ((FOLTVar indx), (s::ss)) =
   463       LTVar((make_type_class s) ^ 
   464         "(" ^ (make_schematic_type_var indx) ^ ")") :: 
   465       (sorts_on_typs ((FOLTVar indx), ss))
   466   | sorts_on_typs ((FOLTFree x), (s::ss)) =
   467       LTFree((make_type_class s) ^ "(" ^ (make_fixed_type_var x) ^ ")") :: 
   468       (sorts_on_typs ((FOLTFree x), ss));
   469 
   470 
   471 (*UGLY: seems to be parsing the "show sorts" output, removing anything that
   472   starts with a left parenthesis.*)
   473 fun remove_type str = hd (String.fields (fn c => c = #"(") str);
   474 
   475 fun pred_of_sort (LTVar x) = ((remove_type x),1)
   476 |   pred_of_sort (LTFree x) = ((remove_type x),1)
   477 
   478 
   479 
   480 
   481 (*Given a list of sorted type variables, return two separate lists.
   482   The first is for TVars, the second for TFrees.*)
   483 fun add_typs_aux [] preds  = ([],[], preds)
   484   | add_typs_aux ((FOLTVar indx,s)::tss) preds = 
   485       let val vs = sorts_on_typs (FOLTVar indx, s)
   486           val preds' = (map pred_of_sort vs)@preds
   487 	  val (vss,fss, preds'') = add_typs_aux tss preds'
   488       in
   489 	  (vs union vss, fss, preds'')
   490       end
   491   | add_typs_aux ((FOLTFree x,s)::tss) preds  =
   492       let val fs = sorts_on_typs (FOLTFree x, s)
   493           val preds' = (map pred_of_sort fs)@preds
   494 	  val (vss,fss, preds'') = add_typs_aux tss preds'
   495       in
   496 	  (vss, fs union fss, preds'')
   497       end;
   498 
   499 fun add_typs (Clause cls) preds  = add_typs_aux (#types_sorts cls) preds 
   500 
   501 
   502 (** make axiom clauses, hypothesis clauses and conjecture clauses. **)
   503 
   504 fun get_tvar_strs [] = []
   505   | get_tvar_strs ((FOLTVar indx,s)::tss) = 
   506       let val vstr = make_schematic_type_var indx
   507       in
   508 	  vstr ins (get_tvar_strs tss)
   509       end
   510   | get_tvar_strs((FOLTFree x,s)::tss) = distinct (get_tvar_strs tss)
   511 
   512 (* FIX add preds and funcs to add typs aux here *)
   513 
   514 fun make_axiom_clause_thm thm (ax_name,cls_id) =
   515     let val (lits,types_sorts, preds, funcs) = literals_of_term (prop_of thm)
   516 	val (tvar_lits,tfree_lits, preds) = add_typs_aux types_sorts preds 
   517         val tvars = get_tvar_strs types_sorts
   518     in 
   519 	make_clause(cls_id,ax_name,Axiom,
   520 	            lits,types_sorts,tvar_lits,tfree_lits,
   521 	            tvars, preds, funcs)
   522     end;
   523 
   524 
   525 
   526 fun make_conjecture_clause n t =
   527     let val (lits,types_sorts, preds, funcs) = literals_of_term t
   528 	val (tvar_lits,tfree_lits, preds) = add_typs_aux types_sorts preds 
   529         val tvars = get_tvar_strs types_sorts
   530     in
   531 	make_clause(n,"conjecture",Conjecture,
   532 	            lits,types_sorts,tvar_lits,tfree_lits,
   533 	            tvars, preds, funcs)
   534     end;
   535     
   536 fun make_conjecture_clauses_aux _ [] = []
   537   | make_conjecture_clauses_aux n (t::ts) =
   538       make_conjecture_clause n t :: make_conjecture_clauses_aux (n+1) ts
   539 
   540 val make_conjecture_clauses = make_conjecture_clauses_aux 0
   541 
   542 
   543 fun make_axiom_clause term (ax_name,cls_id) =
   544     let val (lits,types_sorts, preds,funcs) = literals_of_term term
   545 	val (tvar_lits,tfree_lits, preds) = add_typs_aux types_sorts preds
   546         val tvars = get_tvar_strs types_sorts	
   547     in 
   548 	make_clause(cls_id,ax_name,Axiom,
   549 	            lits,types_sorts,tvar_lits,tfree_lits,
   550 	            tvars, preds,funcs)
   551     end;
   552 
   553 
   554 
   555  
   556 (**** Isabelle arities ****)
   557 
   558 exception ARCLAUSE of string;
   559  
   560 
   561 type class = string; 
   562 type tcons = string; 
   563 
   564 
   565 datatype arLit = TConsLit of bool * (class * tcons * string list) | TVarLit of bool * (class * string);
   566  
   567 datatype arityClause =  
   568 	 ArityClause of {clause_id: clause_id,
   569 	  	         axiom_name: axiom_name,
   570 			 kind: kind,
   571 			 conclLit: arLit,
   572 			 premLits: arLit list};
   573 
   574 
   575 fun get_TVars 0 = []
   576   | get_TVars n = ("T_" ^ (Int.toString n)) :: get_TVars (n-1);
   577 
   578 
   579 
   580 fun pack_sort(_,[])  = raise ARCLAUSE("Empty Sort Found") 
   581   | pack_sort(tvar, [cls]) = [(make_type_class cls, tvar)] 
   582   | pack_sort(tvar, cls::srt) =  (make_type_class cls,tvar) :: (pack_sort(tvar, srt));
   583     
   584     
   585 fun make_TVarLit (b,(cls,str)) = TVarLit(b,(cls,str));
   586 fun make_TConsLit (b,(cls,tcons,tvars)) = TConsLit(b,(make_type_class cls,make_fixed_type_const tcons,tvars));
   587 
   588 fun make_axiom_arity_clause (tcons,n,(res,args)) =
   589    let val nargs = length args
   590        val tvars = get_TVars nargs
   591        val tvars_srts = ListPair.zip (tvars,args)
   592        val tvars_srts' = union_all(map pack_sort tvars_srts)
   593        val false_tvars_srts' = map (pair false) tvars_srts'
   594    in
   595       ArityClause {clause_id = n, kind = Axiom, 
   596                    axiom_name = tcons,
   597                    conclLit = make_TConsLit(true,(res,tcons,tvars)), 
   598                    premLits = map make_TVarLit false_tvars_srts'}
   599    end;
   600     
   601 (*The number of clauses generated from cls, including type clauses*)
   602 fun num_of_clauses (Clause cls) =
   603     let val num_tfree_lits = 
   604 	      if !keep_types then length (#tfree_type_literals cls)
   605 	      else 0
   606     in 	1 + num_tfree_lits  end;
   607 
   608 
   609 (**** Isabelle class relations ****)
   610 
   611 
   612 datatype classrelClause = 
   613 	 ClassrelClause of {clause_id: clause_id,
   614 			    subclass: class,
   615 			    superclass: class option};
   616 
   617 
   618 fun make_axiom_classrelClause n subclass superclass =
   619   ClassrelClause {clause_id = n,
   620                   subclass = subclass, superclass = superclass};
   621 
   622 
   623 fun classrelClauses_of_aux n sub [] = []
   624   | classrelClauses_of_aux n sub (sup::sups) =
   625       make_axiom_classrelClause n sub (SOME sup) :: classrelClauses_of_aux (n+1) sub sups;
   626 
   627 
   628 fun classrelClauses_of (sub,sups) = 
   629     case sups of [] => [make_axiom_classrelClause 0 sub NONE]
   630 	       | _ => classrelClauses_of_aux 0 sub sups;
   631 
   632 
   633 (***** Isabelle arities *****)
   634 
   635 
   636 fun arity_clause _ (tcons, []) = []
   637   | arity_clause n (tcons, ar::ars) =
   638       make_axiom_arity_clause (tcons,n,ar) :: 
   639       arity_clause (n+1) (tcons,ars);
   640 
   641 fun multi_arity_clause [] = []
   642   | multi_arity_clause (tcon_ar :: tcons_ars)  =
   643       arity_clause 0 tcon_ar  @  multi_arity_clause tcons_ars 
   644 
   645 fun arity_clause_thy thy =
   646   let val arities = #arities (Type.rep_tsig (Sign.tsig_of thy))
   647   in multi_arity_clause (Symtab.dest arities) end;
   648 
   649 
   650 (* Isabelle classes *)
   651 
   652 type classrelClauses = classrelClause list Symtab.table;
   653 
   654 val classrel_of = #2 o #classes o Type.rep_tsig o Sign.tsig_of;
   655 fun classrel_clauses_classrel (C: Sorts.classes) = map classrelClauses_of (Graph.dest C);
   656 val classrel_clauses_thy = List.concat o classrel_clauses_classrel o classrel_of;
   657 
   658 
   659 
   660 (****!!!! Changed for typed equality !!!!****)
   661 
   662 fun wrap_eq_type typ t = eq_typ_wrapper ^"(" ^ t ^ "," ^ typ ^ ")";
   663 
   664 (* Only need to wrap equality's arguments with "typeinfo" if the output clauses are typed && if we specifically ask for types to be included.   *)
   665 fun string_of_equality (typ,terms) =
   666       let val [tstr1,tstr2] = map string_of_term terms
   667       in
   668 	  if !keep_types andalso !special_equal 
   669 	  then "equal(" ^ (wrap_eq_type typ tstr1) ^ "," ^ 
   670 		 	  (wrap_eq_type typ tstr2) ^ ")"
   671 	  else "equal(" ^ tstr1 ^ "," ^ tstr2 ^ ")"
   672       end
   673 and string_of_term (UVar(x,_)) = x
   674   | string_of_term (Fun("equal",typ,terms)) = string_of_equality(typ,terms)
   675   | string_of_term (Fun (name,typ,[])) = name
   676   | string_of_term (Fun (name,typ,terms)) = 
   677       let val terms' = map string_of_term terms
   678       in
   679 	  if !keep_types andalso typ<>"" 
   680 	  then name ^ (paren_pack (terms' @ [typ]))
   681 	  else name ^ (paren_pack terms')
   682       end;
   683 
   684 (* before output the string of the predicate, check if the predicate corresponds to an equality or not. *)
   685 fun string_of_predicate (Predicate("equal",typ,terms)) = 
   686       string_of_equality(typ,terms)
   687   | string_of_predicate (Predicate(name,_,[])) = name 
   688   | string_of_predicate (Predicate(name,typ,terms)) = 
   689       let val terms_as_strings = map string_of_term terms
   690       in
   691 	  if !keep_types andalso typ<>""
   692 	  then name ^ (paren_pack (terms_as_strings @ [typ]))
   693 	  else name ^ (paren_pack terms_as_strings) 
   694       end;
   695 
   696 
   697 fun string_of_clausename (cls_id,ax_name) = 
   698     clause_prefix ^ ascii_of ax_name ^ "_" ^ Int.toString cls_id;
   699 
   700 fun string_of_type_clsname (cls_id,ax_name,idx) = 
   701     string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx);
   702     
   703 
   704 (********************************)
   705 (* Code for producing DFG files *)
   706 (********************************)
   707 
   708 fun dfg_literal (Literal(pol,pred,tag)) =
   709     let val pred_string = string_of_predicate pred
   710     in
   711 	if pol then pred_string else "not(" ^pred_string ^ ")"  
   712     end;
   713 
   714 
   715 (* FIX: what does this mean? *)
   716 (*fun dfg_of_typeLit (LTVar x) = "not(" ^ x ^ ")"
   717   | dfg_of_typeLit (LTFree x) = "(" ^ x ^ ")";*)
   718 
   719 fun dfg_of_typeLit (LTVar x) =  x 
   720   | dfg_of_typeLit (LTFree x) = x ;
   721  
   722 (*Make the string of universal quantifiers for a clause*)
   723 fun forall_open ([],[]) = ""
   724   | forall_open (vars,tvars) = "forall([" ^ (commas (tvars@vars))^ "],\n"
   725 
   726 fun forall_close ([],[]) = ""
   727   | forall_close (vars,tvars) = ")"
   728 
   729 fun gen_dfg_cls (cls_id,ax_name,knd,lits,tvars,vars) = 
   730     "clause( %(" ^ knd ^ ")\n" ^ forall_open(vars,tvars) ^ 
   731     "or(" ^ lits ^ ")" ^ forall_close(vars,tvars) ^ ",\n" ^ 
   732     string_of_clausename (cls_id,ax_name) ^  ").";
   733 
   734 fun gen_dfg_type_cls (cls_id,ax_name,knd,tfree_lit,idx,tvars,vars) = 
   735     "clause( %(" ^ knd ^ ")\n" ^ forall_open(vars,tvars) ^ 
   736     "or( " ^ tfree_lit ^ ")" ^ forall_close(vars,tvars) ^ ",\n" ^ 
   737     string_of_type_clsname (cls_id,ax_name,idx) ^  ").";
   738 
   739 fun dfg_clause_aux (Clause cls) = 
   740   let val lits = map dfg_literal (#literals cls)
   741       val tvar_lits_strs = 
   742 	  if !keep_types then map dfg_of_typeLit (#tvar_type_literals cls) 
   743 	  else []
   744       val tfree_lits =
   745           if !keep_types then map dfg_of_typeLit (#tfree_type_literals cls)
   746           else []
   747   in
   748       (tvar_lits_strs @ lits, tfree_lits)
   749   end; 
   750 
   751 
   752 fun dfg_folterms (Literal(pol,pred,tag)) = 
   753   let val Predicate (predname, foltype, folterms) = pred
   754   in
   755       folterms
   756   end
   757 
   758  
   759 fun get_uvars (UVar(a,typ)) = [a] 
   760 |   get_uvars (Fun (_,typ,tlist)) = union_all(map get_uvars tlist)
   761 
   762 
   763 fun is_uvar (UVar _) = true
   764 |   is_uvar (Fun _) = false;
   765 
   766 fun uvar_name (UVar(a,_)) = a
   767 |   uvar_name (Fun (a,_,_)) = raise CLAUSE("Not a variable", Const(a,dummyT));
   768 
   769 fun mergelist [] = []
   770 |   mergelist (x::xs) = x @ mergelist xs
   771 
   772 fun dfg_vars (Clause cls) =
   773     let val lits = #literals cls
   774         val folterms = mergelist(map dfg_folterms lits)
   775     in 
   776         union_all(map get_uvars folterms)
   777     end
   778 
   779 
   780 fun dfg_tvars (Clause cls) =(#tvars cls)
   781 
   782 
   783 	
   784 (* make this return funcs and preds too? *)
   785 fun string_of_predname (Predicate("equal",typ,terms)) = "EQUALITY"
   786   | string_of_predname (Predicate(name,_,[])) = name 
   787   | string_of_predname (Predicate(name,typ,terms)) = name
   788     
   789 	
   790 (* make this return funcs and preds too? *)
   791 
   792 fun string_of_predicate (Predicate("equal",typ,terms)) =  
   793       string_of_equality(typ,terms)
   794   | string_of_predicate (Predicate(name,_,[])) = name 
   795   | string_of_predicate (Predicate(name,typ,terms)) = 
   796       let val terms_as_strings = map string_of_term terms
   797       in
   798 	  if !keep_types andalso typ<>""
   799 	  then name ^ (paren_pack  (terms_as_strings @ [typ]))
   800 	  else name ^ (paren_pack terms_as_strings) 
   801       end;
   802 
   803 
   804 fun concat_with sep []  = ""
   805   | concat_with sep [x] = "(" ^ x ^ ")"
   806   | concat_with sep (x::xs) = "(" ^ x ^ ")" ^  sep ^ (concat_with sep xs);
   807 
   808 fun dfg_pred (Literal(pol,pred,tag)) ax_name = 
   809     (string_of_predname pred) ^ " " ^ ax_name
   810 
   811 fun dfg_clause cls =
   812     let val (lits,tfree_lits) = dfg_clause_aux cls 
   813              (*"lits" includes the typing assumptions (TVars)*)
   814         val vars = dfg_vars cls
   815         val tvars = dfg_tvars cls
   816 	val knd = string_of_kind cls
   817 	val lits_str = commas lits
   818 	val cls_id = get_clause_id cls
   819 	val axname = get_axiomName cls
   820 	val cls_str = gen_dfg_cls(cls_id,axname,knd,lits_str,tvars, vars) 			
   821         fun typ_clss k [] = []
   822           | typ_clss k (tfree :: tfrees) = 
   823               (gen_dfg_type_cls(cls_id,axname,knd,tfree,k, tvars,vars)) :: 
   824               (typ_clss (k+1) tfrees)
   825     in 
   826 	cls_str :: (typ_clss 0 tfree_lits)
   827     end;
   828 
   829 fun string_of_arity (name, num) =  name ^ "," ^ (Int.toString num) 
   830 
   831 fun string_of_preds preds = 
   832   "predicates[" ^ (concat_with ", " (map string_of_arity preds)) ^ "].\n";
   833 
   834 fun string_of_funcs funcs =
   835   "functions[" ^ (concat_with ", " (map string_of_arity funcs)) ^ "].\n" ;
   836 
   837 
   838 fun string_of_symbols predstr funcstr = 
   839   "list_of_symbols.\n" ^ predstr  ^ funcstr  ^ "end_of_list.\n\n";
   840 
   841 
   842 fun string_of_axioms axstr = 
   843   "list_of_clauses(axioms,cnf).\n" ^ axstr ^ "end_of_list.\n\n";
   844 
   845 
   846 fun string_of_conjectures conjstr = 
   847   "list_of_clauses(conjectures,cnf).\n" ^ conjstr ^ "end_of_list.\n\n";
   848 
   849 fun string_of_descrip () = 
   850   "list_of_descriptions.\nname({*[ File     : ],[ Names    :]*}).\nauthor({*[ Source   :]*}).\nstatus(unknown).\ndescription({*[ Refs     :]*}).\nend_of_list.\n\n"
   851 
   852 
   853 fun string_of_start name = "%------------------------------------------------------------------------------\nbegin_problem(" ^ name ^ ").\n\n";
   854 
   855 
   856 fun string_of_end () = "end_problem.\n%------------------------------------------------------------------------------";
   857 
   858 
   859 fun clause2dfg cls =
   860     let val (lits,tfree_lits) = dfg_clause_aux cls 
   861             (*"lits" includes the typing assumptions (TVars)*)
   862 	val cls_id = get_clause_id cls
   863 	val ax_name = get_axiomName cls
   864         val vars = dfg_vars cls
   865         val tvars = dfg_tvars cls
   866         val funcs = funcs_of_cls cls
   867         val preds = preds_of_cls cls
   868 	val knd = string_of_kind cls
   869 	val lits_str = commas lits
   870 	val cls_str = gen_dfg_cls(cls_id,ax_name,knd,lits_str,tvars,vars) 
   871     in
   872 	(cls_str,tfree_lits) 
   873     end;
   874 
   875 
   876 
   877 fun tfree_dfg_clause tfree_lit =
   878   "clause( %(conjecture)\n" ^ "or( " ^ tfree_lit ^ "),\n" ^ "tfree_tcs" ^ ")."
   879 
   880 
   881 fun gen_dfg_file probname axioms conjectures funcs preds = 
   882     let val axstrs_tfrees = (map clause2dfg axioms)
   883 	val (axstrs, atfrees) = ListPair.unzip axstrs_tfrees
   884         val axstr = (space_implode "\n" axstrs) ^ "\n\n"
   885         val conjstrs_tfrees = (map clause2dfg conjectures)
   886 	val (conjstrs, atfrees) = ListPair.unzip conjstrs_tfrees
   887         val tfree_clss = map tfree_dfg_clause (union_all atfrees) 
   888         val conjstr = (space_implode "\n" (tfree_clss@conjstrs)) ^ "\n\n"
   889         val funcstr = string_of_funcs funcs
   890         val predstr = string_of_preds preds
   891     in
   892        (string_of_start probname) ^ (string_of_descrip ()) ^ 
   893        (string_of_symbols funcstr predstr) ^  
   894        (string_of_axioms axstr) ^
   895        (string_of_conjectures conjstr) ^ (string_of_end ())
   896     end;
   897    
   898 fun clauses2dfg [] probname axioms conjectures funcs preds = 
   899       let val funcs' = (union_all(map funcs_of_cls axioms)) @ funcs
   900 	  val preds' = (union_all(map preds_of_cls axioms)) @ preds
   901       in
   902 	 gen_dfg_file probname axioms conjectures funcs' preds' 
   903       end
   904  | clauses2dfg (cls::clss) probname axioms conjectures funcs preds = 
   905      let val (lits,tfree_lits) = dfg_clause_aux cls
   906 	       (*"lits" includes the typing assumptions (TVars)*)
   907 	 val cls_id = get_clause_id cls
   908 	 val ax_name = get_axiomName cls
   909 	 val vars = dfg_vars cls
   910 	 val tvars = dfg_tvars cls
   911 	 val funcs' = (funcs_of_cls cls) union funcs
   912 	 val preds' = (preds_of_cls cls) union preds
   913 	 val knd = string_of_kind cls
   914 	 val lits_str = concat_with ", " lits
   915 	 val axioms' = if knd = "axiom" then (cls::axioms) else axioms
   916 	 val conjectures' = 
   917 	     if knd = "conjecture" then (cls::conjectures) else conjectures
   918      in
   919 	 clauses2dfg clss probname axioms' conjectures' funcs' preds' 
   920      end;
   921 
   922 
   923 fun string_of_arClauseID (ArityClause {clause_id,axiom_name,...}) =
   924     arclause_prefix ^ ascii_of axiom_name ^ "_" ^ Int.toString clause_id;
   925 
   926 fun string_of_arKind (ArityClause arcls) = name_of_kind(#kind arcls);
   927 
   928 (*FIXME!!! currently is TPTP format!*)
   929 fun dfg_of_arLit (TConsLit(b,(c,t,args))) =
   930       let val pol = if b then "++" else "--"
   931 	  val arg_strs = (case args of [] => "" | _ => paren_pack args)
   932       in 
   933 	  pol ^ c ^ "(" ^ t ^ arg_strs ^ ")"
   934       end
   935   | dfg_of_arLit (TVarLit(b,(c,str))) =
   936       let val pol = if b then "++" else "--"
   937       in
   938 	  pol ^ c ^ "(" ^ str ^ ")"
   939       end;
   940     
   941 
   942 fun dfg_of_conclLit (ArityClause arcls) = dfg_of_arLit (#conclLit arcls);
   943      
   944 
   945 fun dfg_of_premLits (ArityClause arcls) = map dfg_of_arLit (#premLits arcls);
   946 		
   947 
   948 
   949 (*FIXME: would this have variables in a forall? *)
   950 
   951 fun dfg_arity_clause arcls = 
   952   let val arcls_id = string_of_arClauseID arcls
   953       val concl_lit = dfg_of_conclLit arcls
   954       val prems_lits = dfg_of_premLits arcls
   955       val knd = string_of_arKind arcls
   956       val all_lits = concl_lit :: prems_lits
   957   in
   958       "clause( %(" ^ knd ^ ")\n" ^  "or( " ^ (bracket_pack all_lits) ^ ")),\n" ^
   959        arcls_id ^  ")."
   960   end;
   961 
   962 
   963 (********************************)
   964 (* code to produce TPTP files   *)
   965 (********************************)
   966 
   967 fun tptp_literal (Literal(pol,pred,tag)) =
   968     let val pred_string = string_of_predicate pred
   969 	val tagged_pol = 
   970 	      if (tag andalso !tagged) then (if pol then "+++" else "---")
   971 	      else (if pol then "++" else "--")
   972      in
   973 	tagged_pol ^ pred_string
   974     end;
   975 
   976 
   977 
   978 fun tptp_of_typeLit (LTVar x) = "--" ^ x
   979   | tptp_of_typeLit (LTFree x) = "++" ^ x;
   980  
   981 
   982 fun gen_tptp_cls (cls_id,ax_name,knd,lits) = 
   983     "input_clause(" ^ string_of_clausename (cls_id,ax_name) ^ "," ^ 
   984     knd ^ "," ^ lits ^ ").";
   985 
   986 fun gen_tptp_type_cls (cls_id,ax_name,knd,tfree_lit,idx) = 
   987     "input_clause(" ^ string_of_type_clsname (cls_id,ax_name,idx) ^ "," ^ 
   988     knd ^ ",[" ^ tfree_lit ^ "]).";
   989 
   990 fun tptp_type_lits (Clause cls) = 
   991     let val lits = map tptp_literal (#literals cls)
   992 	val tvar_lits_strs =
   993 	      if !keep_types 
   994 	      then (map tptp_of_typeLit (#tvar_type_literals cls)) 
   995 	      else []
   996 	val tfree_lits = 
   997 	      if !keep_types
   998 	      then (map tptp_of_typeLit (#tfree_type_literals cls)) 
   999 	      else []
  1000     in
  1001 	(tvar_lits_strs @ lits, tfree_lits)
  1002     end; 
  1003 
  1004 fun tptp_clause cls =
  1005     let val (lits,tfree_lits) = tptp_type_lits cls 
  1006             (*"lits" includes the typing assumptions (TVars)*)
  1007 	val cls_id = get_clause_id cls
  1008 	val ax_name = get_axiomName cls
  1009 	val knd = string_of_kind cls
  1010 	val lits_str = bracket_pack lits
  1011 	val cls_str = gen_tptp_cls(cls_id,ax_name,knd,lits_str) 			 
  1012 	fun typ_clss k [] = []
  1013           | typ_clss k (tfree :: tfrees) = 
  1014               gen_tptp_type_cls(cls_id,ax_name,knd,tfree,k) :: 
  1015               typ_clss (k+1) tfrees
  1016     in 
  1017 	cls_str :: (typ_clss 0 tfree_lits)
  1018     end;
  1019 
  1020 fun clause2tptp cls =
  1021     let val (lits,tfree_lits) = tptp_type_lits cls 
  1022             (*"lits" includes the typing assumptions (TVars)*)
  1023 	val cls_id = get_clause_id cls
  1024 	val ax_name = get_axiomName cls
  1025 	val knd = string_of_kind cls
  1026 	val lits_str = bracket_pack lits
  1027 	val cls_str = gen_tptp_cls(cls_id,ax_name,knd,lits_str) 
  1028     in
  1029 	(cls_str,tfree_lits) 
  1030     end;
  1031 
  1032 
  1033 fun tfree_clause tfree_lit =
  1034     "input_clause(" ^ "tfree_tcs," ^ "conjecture" ^ ",[" ^ tfree_lit ^ "]).";
  1035 
  1036 
  1037 fun tptp_of_arLit (TConsLit(b,(c,t,args))) =
  1038       let val pol = if b then "++" else "--"
  1039 	  val  arg_strs = (case args of [] => "" | _ => paren_pack args)
  1040       in 
  1041 	  pol ^ c ^ "(" ^ t ^ arg_strs ^ ")"
  1042       end
  1043   | tptp_of_arLit (TVarLit(b,(c,str))) =
  1044       let val pol = if b then "++" else "--"
  1045       in
  1046 	  pol ^ c ^ "(" ^ str ^ ")"
  1047       end;
  1048     
  1049 
  1050 fun tptp_of_conclLit (ArityClause arcls) = tptp_of_arLit (#conclLit arcls);
  1051      
  1052 fun tptp_of_premLits (ArityClause arcls) = map tptp_of_arLit (#premLits arcls);
  1053 		
  1054 fun tptp_arity_clause arcls = 
  1055     let val arcls_id = string_of_arClauseID arcls
  1056 	val concl_lit = tptp_of_conclLit arcls
  1057 	val prems_lits = tptp_of_premLits arcls
  1058 	val knd = string_of_arKind arcls
  1059 	val all_lits = concl_lit :: prems_lits
  1060     in
  1061 	"input_clause(" ^ arcls_id ^ "," ^ knd ^ "," ^ 
  1062 	(bracket_pack all_lits) ^ ")."
  1063     end;
  1064 
  1065 fun tptp_classrelLits sub sup = 
  1066     let val tvar = "(T)"
  1067     in 
  1068 	case sup of NONE => "[++" ^ sub ^ tvar ^ "]"
  1069 		  | (SOME supcls) =>  "[--" ^ sub ^ tvar ^ ",++" ^ supcls ^ tvar ^ "]"
  1070     end;
  1071 
  1072 
  1073 fun tptp_classrelClause (ClassrelClause {clause_id,subclass,superclass,...}) =
  1074     let val relcls_id = clrelclause_prefix ^ ascii_of subclass ^ "_" ^ 
  1075                         Int.toString clause_id
  1076 	val lits = tptp_classrelLits (make_type_class subclass) 
  1077 	                (Option.map make_type_class superclass)
  1078     in
  1079 	"input_clause(" ^ relcls_id ^ ",axiom," ^ lits ^ ")."
  1080     end; 
  1081 
  1082 
  1083 
  1084 end;