src/HOL/Tools/res_clause.ML
author paulson
Tue, 31 Jan 2006 10:39:13 +0100
changeset 18863 a113b6839df1
parent 18856 4669dec681f4
child 18868 7cfc21ee0370
permissions -rw-r--r--
reorganization of code to support DFG otuput
paulson@15347
     1
(*  Author: Jia Meng, Cambridge University Computer Laboratory
paulson@15347
     2
    ID: $Id$
paulson@15347
     3
    Copyright 2004 University of Cambridge
paulson@15347
     4
paulson@15347
     5
ML data structure for storing/printing FOL clauses and arity clauses.
paulson@15347
     6
Typed equality is treated differently.
paulson@15347
     7
*)
paulson@15347
     8
quigley@17150
     9
(* works for writeoutclasimp on typed *)
paulson@15347
    10
signature RES_CLAUSE =
paulson@15347
    11
  sig
paulson@17404
    12
  val keep_types : bool ref
paulson@17404
    13
  val special_equal : bool ref
paulson@17404
    14
  val tagged : bool ref
quigley@17150
    15
paulson@17404
    16
  exception ARCLAUSE of string
paulson@17404
    17
  exception CLAUSE of string * term
paulson@17404
    18
  type arityClause 
paulson@17404
    19
  type classrelClause
paulson@17404
    20
  type clause
paulson@17404
    21
  val init : theory -> unit
paulson@17404
    22
  val make_axiom_clause : Term.term -> string * int -> clause
paulson@17888
    23
  val make_conjecture_clauses : term list -> clause list
paulson@17404
    24
  val get_axiomName : clause ->  string
paulson@17404
    25
  val isTaut : clause -> bool
paulson@17404
    26
  val num_of_clauses : clause -> int
quigley@17150
    27
paulson@17845
    28
  val arity_clause_thy: theory -> arityClause list 
paulson@17845
    29
  val classrel_clauses_thy: theory -> classrelClause list 
paulson@17845
    30
paulson@17404
    31
  val tptp_arity_clause : arityClause -> string
paulson@17404
    32
  val tptp_classrelClause : classrelClause -> string
paulson@17404
    33
  val tptp_clause : clause -> string list
paulson@17404
    34
  val clause2tptp : clause -> string * string list
paulson@18863
    35
  val tptp_tfree_clause : string -> string
paulson@17404
    36
  val schematic_var_prefix : string
paulson@17404
    37
  val fixed_var_prefix : string
paulson@17404
    38
  val tvar_prefix : string
paulson@17404
    39
  val tfree_prefix : string
paulson@17404
    40
  val clause_prefix : string 
paulson@17404
    41
  val arclause_prefix : string
paulson@17404
    42
  val const_prefix : string
paulson@17404
    43
  val tconst_prefix : string 
paulson@17404
    44
  val class_prefix : string 
mengj@17908
    45
mengj@17908
    46
  val union_all : ''a list list -> ''a list
mengj@17908
    47
  val ascii_of : String.string -> String.string
mengj@17908
    48
  val paren_pack : string list -> string
mengj@17908
    49
  val bracket_pack : string list -> string
mengj@17908
    50
  val make_schematic_var : String.string * int -> string
mengj@17908
    51
  val make_fixed_var : String.string -> string
mengj@17908
    52
  val make_schematic_type_var : string * int -> string
mengj@17908
    53
  val make_fixed_type_var : string -> string
mengj@17908
    54
  val make_fixed_const : String.string -> string		
mengj@17908
    55
  val make_fixed_type_const : String.string -> string   
mengj@17908
    56
  val make_type_class : String.string -> string
mengj@17999
    57
  val isMeta : String.string -> bool
mengj@17999
    58
  
mengj@17999
    59
  type typ_var
mengj@17999
    60
  val mk_typ_var_sort : Term.typ -> typ_var * sort
mengj@17999
    61
  type type_literal
paulson@18856
    62
  val add_typs_aux : (typ_var * string list) list -> type_literal list * type_literal list
mengj@17999
    63
  val gen_tptp_cls : int * string * string * string -> string
mengj@17999
    64
  val gen_tptp_type_cls : int * string * string * string * int -> string
mengj@17999
    65
  val tptp_of_typeLit : type_literal -> string
mengj@18275
    66
paulson@18420
    67
  (*for hashing*)
paulson@18420
    68
  val clause_eq : clause * clause -> bool
paulson@18449
    69
  val hash_clause : clause -> int
paulson@18420
    70
mengj@18439
    71
  val list_ord : ('a * 'b -> order) -> 'a list * 'b list -> order
mengj@18439
    72
  type fol_type
mengj@18439
    73
  val types_ord : fol_type list * fol_type list -> order
mengj@18439
    74
  val string_of_fol_type : fol_type -> string
mengj@18439
    75
  val mk_fol_type: string * string * fol_type list -> fol_type
paulson@18863
    76
  val types_eq: fol_type list * fol_type list -> 
paulson@18863
    77
        (string*string) list * (string*string) list -> 
paulson@18863
    78
        bool * ((string*string) list * (string*string) list)
mengj@18439
    79
  val check_var_pairs: ''a * ''b -> (''a * ''b) list -> int
paulson@18863
    80
paulson@18863
    81
  val dfg_write_file: thm list -> string -> 
paulson@18863
    82
        (clause list * classrelClause list * arityClause list) -> unit
paulson@18863
    83
  val tptp_write_file: thm list -> string -> 
paulson@18863
    84
        (clause list * classrelClause list * arityClause list) -> unit
paulson@18863
    85
  val writeln_strs: TextIO.outstream -> TextIO.vector list -> unit
paulson@15347
    86
  end;
paulson@15347
    87
paulson@18420
    88
structure ResClause : RES_CLAUSE =
paulson@15347
    89
struct
paulson@15347
    90
paulson@15347
    91
(* Added for typed equality *)
paulson@15347
    92
val special_equal = ref false; (* by default,equality does not carry type information *)
paulson@15347
    93
val eq_typ_wrapper = "typeinfo"; (* default string *)
paulson@15347
    94
paulson@15347
    95
paulson@15347
    96
val schematic_var_prefix = "V_";
paulson@15347
    97
val fixed_var_prefix = "v_";
paulson@15347
    98
paulson@17230
    99
val tvar_prefix = "T_";
paulson@17230
   100
val tfree_prefix = "t_";
paulson@15347
   101
paulson@15347
   102
val clause_prefix = "cls_"; 
paulson@17525
   103
val arclause_prefix = "clsarity_" 
paulson@17525
   104
val clrelclause_prefix = "clsrel_";
paulson@15347
   105
paulson@17230
   106
val const_prefix = "c_";
paulson@17230
   107
val tconst_prefix = "tc_"; 
paulson@16199
   108
val class_prefix = "class_"; 
paulson@15347
   109
paulson@17775
   110
fun union_all xss = foldl (op union) [] xss;
paulson@17775
   111
paulson@17775
   112
(*Provide readable names for the more common symbolic functions*)
paulson@15347
   113
val const_trans_table =
paulson@15347
   114
      Symtab.make [("op =", "equal"),
paulson@15347
   115
	  	   ("op <=", "lessequals"),
paulson@15347
   116
		   ("op <", "less"),
paulson@15347
   117
		   ("op &", "and"),
paulson@15347
   118
		   ("op |", "or"),
paulson@17375
   119
		   ("op +", "plus"),
paulson@17375
   120
		   ("op -", "minus"),
paulson@17375
   121
		   ("op *", "times"),
paulson@18676
   122
		   ("Divides.op div", "div"),
paulson@18676
   123
		   ("HOL.divide", "divide"),
paulson@15347
   124
		   ("op -->", "implies"),
paulson@17375
   125
		   ("{}", "emptyset"),
paulson@15347
   126
		   ("op :", "in"),
paulson@15347
   127
		   ("op Un", "union"),
paulson@18390
   128
		   ("op Int", "inter"),
paulson@18390
   129
		   ("List.op @", "append")];
paulson@15347
   130
paulson@17230
   131
val type_const_trans_table =
paulson@18411
   132
      Symtab.make [("*", "prod"),
paulson@18411
   133
	  	   ("+", "sum"),
paulson@18411
   134
		   ("~=>", "map")];
paulson@15347
   135
paulson@15610
   136
(*Escaping of special characters.
paulson@15610
   137
  Alphanumeric characters are left unchanged.
paulson@15610
   138
  The character _ goes to __
paulson@15610
   139
  Characters in the range ASCII space to / go to _A to _P, respectively.
paulson@15610
   140
  Other printing characters go to _NNN where NNN is the decimal ASCII code.*)
paulson@15610
   141
local
paulson@15610
   142
paulson@15610
   143
val A_minus_space = Char.ord #"A" - Char.ord #" ";
paulson@15610
   144
paulson@15347
   145
fun ascii_of_c c =
paulson@15610
   146
  if Char.isAlphaNum c then String.str c
paulson@15610
   147
  else if c = #"_" then "__"
paulson@15610
   148
  else if #" " <= c andalso c <= #"/" 
paulson@15610
   149
       then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
paulson@15610
   150
  else if Char.isPrint c then ("_" ^ Int.toString (Char.ord c))
paulson@15610
   151
  else ""
paulson@15347
   152
paulson@15610
   153
in
paulson@15610
   154
paulson@15610
   155
val ascii_of = String.translate ascii_of_c;
paulson@15610
   156
paulson@15610
   157
end;
paulson@15347
   158
paulson@17525
   159
(* convert a list of strings into one single string; surrounded by brackets *)
paulson@18218
   160
fun paren_pack [] = ""   (*empty argument list*)
paulson@18218
   161
  | paren_pack strings = "(" ^ commas strings ^ ")";
paulson@17525
   162
paulson@17525
   163
fun bracket_pack strings = "[" ^ commas strings ^ "]";
paulson@17525
   164
paulson@17525
   165
paulson@16925
   166
(*Remove the initial ' character from a type variable, if it is present*)
paulson@16925
   167
fun trim_type_var s =
paulson@16925
   168
  if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
paulson@16925
   169
  else error ("trim_type: Malformed type variable encountered: " ^ s);
paulson@16925
   170
paulson@16903
   171
fun ascii_of_indexname (v,0) = ascii_of v
paulson@17525
   172
  | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
paulson@15347
   173
paulson@17230
   174
fun make_schematic_var v = schematic_var_prefix ^ (ascii_of_indexname v);
paulson@15347
   175
fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x);
paulson@15347
   176
paulson@16925
   177
fun make_schematic_type_var (x,i) = 
paulson@16925
   178
      tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
paulson@16925
   179
fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
paulson@15347
   180
paulson@18411
   181
fun lookup_const c =
wenzelm@17412
   182
    case Symtab.lookup const_trans_table c of
paulson@17230
   183
        SOME c' => c'
paulson@18411
   184
      | NONE => ascii_of c;
paulson@17230
   185
paulson@18411
   186
fun lookup_type_const c = 
wenzelm@17412
   187
    case Symtab.lookup type_const_trans_table c of
paulson@17230
   188
        SOME c' => c'
paulson@18411
   189
      | NONE => ascii_of c;
paulson@18411
   190
paulson@18411
   191
fun make_fixed_const "op =" = "equal"   (*MUST BE "equal" because it's built-in to ATPs*)
paulson@18411
   192
  | make_fixed_const c      = const_prefix ^ lookup_const c;
paulson@18411
   193
paulson@18411
   194
fun make_fixed_type_const c = tconst_prefix ^ lookup_type_const c;
quigley@17150
   195
wenzelm@17261
   196
fun make_type_class clas = class_prefix ^ ascii_of clas;
quigley@17150
   197
quigley@17150
   198
paulson@18798
   199
(***** definitions and functions for FOL clauses, for conversion to TPTP or DFG format. *****)
paulson@15347
   200
paulson@17230
   201
val keep_types = ref true;
paulson@15347
   202
paulson@15347
   203
datatype kind = Axiom | Hypothesis | Conjecture;
paulson@15347
   204
fun name_of_kind Axiom = "axiom"
paulson@15347
   205
  | name_of_kind Hypothesis = "hypothesis"
paulson@15347
   206
  | name_of_kind Conjecture = "conjecture";
paulson@15347
   207
paulson@15347
   208
type clause_id = int;
paulson@15347
   209
type axiom_name = string;
paulson@15347
   210
paulson@15347
   211
paulson@15347
   212
type polarity = bool;
paulson@15347
   213
paulson@15347
   214
(* "tag" is used for vampire specific syntax  *)
paulson@15347
   215
type tag = bool; 
paulson@15347
   216
paulson@15347
   217
paulson@15347
   218
(**** Isabelle FOL clauses ****)
paulson@15347
   219
paulson@15347
   220
val tagged = ref false;
paulson@15347
   221
paulson@15347
   222
type pred_name = string;
mengj@18402
   223
mengj@18402
   224
datatype typ_var = FOLTVar of indexname | FOLTFree of string;
mengj@18402
   225
paulson@18798
   226
(*FIXME: give the constructors more sensible names*)
mengj@18402
   227
datatype fol_type = AtomV of string
mengj@18402
   228
		  | AtomF of string
mengj@18402
   229
		  | Comp of string * fol_type list;
mengj@18402
   230
mengj@18402
   231
fun string_of_fol_type (AtomV x) = x
mengj@18402
   232
  | string_of_fol_type (AtomF x) = x
mengj@18402
   233
  | string_of_fol_type (Comp(tcon,tps)) = 
paulson@18856
   234
      tcon ^ (paren_pack (map string_of_fol_type tps));
paulson@18856
   235
      
mengj@18439
   236
fun mk_fol_type ("Var",x,_) = AtomV(x)
mengj@18439
   237
  | mk_fol_type ("Fixed",x,_) = AtomF(x)
mengj@18439
   238
  | mk_fol_type ("Comp",con,args) = Comp(con,args)
mengj@18439
   239
mengj@18439
   240
paulson@18798
   241
(*First string is the type class; the second is a TVar or TFfree*)
paulson@18798
   242
datatype type_literal = LTVar of string * string | LTFree of string * string;
paulson@15347
   243
paulson@18856
   244
datatype fol_term = UVar of string * fol_type
paulson@18856
   245
                 | Fun of string * fol_type list * fol_term list;
paulson@18856
   246
datatype predicate = Predicate of pred_name * fol_type list * fol_term list;
paulson@15347
   247
paulson@15347
   248
datatype literal = Literal of polarity * predicate * tag;
paulson@15347
   249
mengj@17999
   250
fun mk_typ_var_sort (TFree(a,s)) = (FOLTFree a,s)
mengj@17999
   251
  | mk_typ_var_sort (TVar(v,s)) = (FOLTVar v,s);
mengj@17999
   252
mengj@17999
   253
paulson@18856
   254
(*A clause has first-order literals and other, type-related literals*)
paulson@15347
   255
datatype clause = 
paulson@15347
   256
	 Clause of {clause_id: clause_id,
paulson@15347
   257
		    axiom_name: axiom_name,
paulson@15347
   258
		    kind: kind,
paulson@15347
   259
		    literals: literal list,
paulson@15347
   260
		    types_sorts: (typ_var * sort) list, 
paulson@15347
   261
                    tvar_type_literals: type_literal list, 
paulson@18856
   262
                    tfree_type_literals: type_literal list};
paulson@15347
   263
paulson@17404
   264
exception CLAUSE of string * term;
paulson@15347
   265
paulson@18856
   266
fun isFalse (Literal (pol,Predicate(pname,_,[]),_)) =
paulson@18856
   267
      (pol andalso pname = "c_False") orelse
paulson@18856
   268
      (not pol andalso pname = "c_True")
paulson@17317
   269
  | isFalse _ = false;
paulson@15347
   270
paulson@18856
   271
fun isTrue (Literal (pol,Predicate(pname,_,[]),_)) =
paulson@18856
   272
      (pol andalso pname = "c_True") orelse
paulson@18856
   273
      (not pol andalso pname = "c_False")
paulson@17404
   274
  | isTrue _ = false;
paulson@17404
   275
  
paulson@17404
   276
fun isTaut (Clause {literals,...}) = exists isTrue literals;  
paulson@17404
   277
paulson@17234
   278
fun make_clause (clause_id,axiom_name,kind,literals,
paulson@18856
   279
                 types_sorts,tvar_type_literals,tfree_type_literals) =
paulson@17317
   280
  if forall isFalse literals 
paulson@17234
   281
  then error "Problem too trivial for resolution (empty clause)"
paulson@17234
   282
  else
paulson@17230
   283
     Clause {clause_id = clause_id, axiom_name = axiom_name, kind = kind, 
paulson@17230
   284
             literals = literals, types_sorts = types_sorts,
paulson@17230
   285
             tvar_type_literals = tvar_type_literals,
paulson@18856
   286
             tfree_type_literals = tfree_type_literals};
paulson@15347
   287
quigley@17150
   288
paulson@17317
   289
(** Some Clause destructor functions **)
paulson@17317
   290
paulson@17317
   291
fun string_of_kind (Clause cls) = name_of_kind (#kind cls);
paulson@17317
   292
paulson@17317
   293
fun get_axiomName (Clause cls) = #axiom_name cls;
paulson@17317
   294
paulson@17317
   295
fun get_clause_id (Clause cls) = #clause_id cls;
paulson@17317
   296
paulson@17317
   297
paulson@18218
   298
(*Declarations of the current theory--to allow suppressing types.*)
paulson@18218
   299
val const_typargs = ref (Library.K [] : (string*typ -> typ list));
paulson@17317
   300
paulson@18218
   301
fun num_typargs(s,T) = if !keep_types then length (!const_typargs (s,T)) else 0;
paulson@16925
   302
paulson@16925
   303
(*Initialize the type suppression mechanism with the current theory before
paulson@16925
   304
    producing any clauses!*)
paulson@18218
   305
fun init thy = (const_typargs := Sign.const_typargs thy);
paulson@16925
   306
    
quigley@17150
   307
mengj@18402
   308
(*Flatten a type to a fol_type while accumulating sort constraints on the TFrees and
paulson@16925
   309
  TVars it contains.*)    
paulson@18218
   310
fun type_of (Type (a, Ts)) = 
paulson@18798
   311
      let val (folTyps, ts) = types_of Ts 
paulson@18218
   312
	  val t = make_fixed_type_const a
paulson@18798
   313
      in (Comp(t,folTyps), ts) end
paulson@18798
   314
  | type_of (TFree (a,s)) = (AtomF(make_fixed_type_var a), [(FOLTFree a, s)]) 
paulson@18798
   315
  | type_of (TVar (v, s)) = (AtomV(make_schematic_type_var v), [(FOLTVar v, s)])
paulson@18218
   316
and types_of Ts =
paulson@18798
   317
      let val (folTyps,ts) = ListPair.unzip (map type_of Ts)
paulson@18798
   318
      in (folTyps, union_all ts) end;
paulson@15390
   319
mengj@18439
   320
paulson@18218
   321
fun const_types_of (c,T) = types_of (!const_typargs (c,T));
paulson@16925
   322
paulson@16903
   323
(* Any variables created via the METAHYPS tactical should be treated as
paulson@16903
   324
   universal vars, although it is represented as "Free(...)" by Isabelle *)
paulson@16903
   325
val isMeta = String.isPrefix "METAHYP1_"
quigley@17150
   326
paulson@18798
   327
fun pred_name_type (Const(c,T)) = (make_fixed_const c, const_types_of (c,T))
paulson@15390
   328
  | pred_name_type (Free(x,T))  = 
paulson@17404
   329
      if isMeta x then raise CLAUSE("Predicate Not First Order 1", Free(x,T)) 
paulson@18798
   330
      else (make_fixed_var x, ([],[]))
paulson@17404
   331
  | pred_name_type (v as Var _) = raise CLAUSE("Predicate Not First Order 2", v)
paulson@17404
   332
  | pred_name_type t        = raise CLAUSE("Predicate input unexpected", t);
paulson@15347
   333
paulson@15615
   334
paulson@18218
   335
(* For typed equality *)
paulson@15615
   336
(* here "arg_typ" is the type of "="'s argument's type, not the type of the equality *)
paulson@15615
   337
(* Find type of equality arg *)
paulson@15615
   338
fun eq_arg_type (Type("fun",[T,_])) = 
paulson@15615
   339
    let val (folT,_) = type_of T;
paulson@17230
   340
    in  folT  end;
paulson@15615
   341
paulson@18798
   342
fun fun_name_type (Const(c,T)) args = (make_fixed_const c, const_types_of (c,T))
paulson@18798
   343
  | fun_name_type (Free(x,T)) args  = 
paulson@18798
   344
       if isMeta x then raise CLAUSE("Function Not First Order", Free(x,T))
paulson@18798
   345
       else (make_fixed_var x, ([],[]))
paulson@17404
   346
  | fun_name_type f args = raise CLAUSE("Function Not First Order 1", f);
quigley@17150
   347
paulson@18856
   348
(*Convert a term to a fol_term while accumulating sort constraints on the TFrees and
paulson@18856
   349
  TVars it contains.*)    
paulson@15347
   350
fun term_of (Var(ind_nm,T)) = 
paulson@18798
   351
      let val (folType,ts) = type_of T
paulson@18798
   352
      in (UVar(make_schematic_var ind_nm, folType), ts) end
paulson@18798
   353
  | term_of (Free(x,T)) = 
paulson@18798
   354
      let val (folType, ts) = type_of T
paulson@17230
   355
      in
paulson@18798
   356
	  if isMeta x then (UVar(make_schematic_var(x,0),folType), ts)
paulson@18798
   357
	  else (Fun(make_fixed_var x, [folType], []), ts)
paulson@17230
   358
      end
paulson@18218
   359
  | term_of app = 
paulson@17230
   360
      let val (f,args) = strip_comb app
paulson@18798
   361
	  val (funName,(contys,ts1)) = fun_name_type f args
paulson@18798
   362
	  val (args',ts2) = terms_of args
paulson@17230
   363
      in
paulson@18218
   364
	  (Fun(funName,contys,args'), 
paulson@18798
   365
	   (union_all (ts1::ts2)))
paulson@17230
   366
      end
paulson@18798
   367
and terms_of ts = ListPair.unzip (map term_of ts)
paulson@15390
   368
paulson@18856
   369
(*Create a predicate value, again accumulating sort constraints.*)    
paulson@17404
   370
fun pred_of (Const("op =", typ), args) =
paulson@17404
   371
      let val arg_typ = eq_arg_type typ 
paulson@18798
   372
	  val (args',ts) = terms_of args
paulson@17404
   373
	  val equal_name = make_fixed_const "op ="
paulson@17404
   374
      in
paulson@18218
   375
	  (Predicate(equal_name,[arg_typ],args'),
paulson@18856
   376
	   union_all ts)
paulson@17404
   377
      end
paulson@17404
   378
  | pred_of (pred,args) = 
paulson@18856
   379
      let val (pname, (predType,ts1)) = pred_name_type pred
paulson@18798
   380
	  val (args',ts2) = terms_of args
paulson@17404
   381
      in
paulson@18856
   382
	  (Predicate(pname,predType,args'), union_all (ts1::ts2))
paulson@17404
   383
      end;
paulson@15347
   384
paulson@17404
   385
(*Treatment of literals, possibly negated or tagged*)
paulson@17404
   386
fun predicate_of ((Const("Not",_) $ P), polarity, tag) =
paulson@17404
   387
      predicate_of (P, not polarity, tag)
paulson@17404
   388
  | predicate_of ((Const("HOL.tag",_) $ P), polarity, tag) =
paulson@17404
   389
      predicate_of (P, polarity, true)
paulson@17404
   390
  | predicate_of (term,polarity,tag) =
paulson@17404
   391
        (pred_of (strip_comb term), polarity, tag);
paulson@15347
   392
paulson@17888
   393
fun literals_of_term1 args (Const("Trueprop",_) $ P) = literals_of_term1 args P
paulson@18856
   394
  | literals_of_term1 args (Const("op |",_) $ P $ Q) = 
paulson@18856
   395
      literals_of_term1 (literals_of_term1 args P) Q
paulson@18856
   396
  | literals_of_term1 (lits, ts) P =
paulson@18856
   397
      let val ((pred, ts'), polarity, tag) = predicate_of (P,true,false)
paulson@18856
   398
	  val lits' = Literal(polarity,pred,tag) :: lits
paulson@17234
   399
      in
paulson@18856
   400
	  (lits', ts union ts')
paulson@17234
   401
      end;
quigley@17150
   402
paulson@18856
   403
val literals_of_term = literals_of_term1 ([],[]);
quigley@17150
   404
mengj@18402
   405
mengj@18403
   406
fun list_ord _ ([],[]) = EQUAL
mengj@18403
   407
  | list_ord _ ([],_) = LESS
mengj@18403
   408
  | list_ord _ (_,[]) = GREATER
mengj@18403
   409
  | list_ord ord (x::xs, y::ys) =
mengj@18403
   410
    let val xy_ord = ord(x,y)
mengj@18403
   411
    in
mengj@18403
   412
	case xy_ord of EQUAL => list_ord ord (xs,ys)
mengj@18403
   413
		     | _ => xy_ord
mengj@18403
   414
    end;
mengj@18403
   415
mengj@18403
   416
fun type_ord (AtomV(_),AtomV(_)) = EQUAL
mengj@18403
   417
  | type_ord (AtomV(_),_) = LESS
mengj@18403
   418
  | type_ord (AtomF(_),AtomV(_)) = GREATER
mengj@18403
   419
  | type_ord (AtomF(f1),AtomF(f2)) = string_ord (f1,f2)
mengj@18403
   420
  | type_ord (AtomF(_),_) = LESS
mengj@18403
   421
  | type_ord (Comp(_,_),AtomV(_)) = GREATER
mengj@18403
   422
  | type_ord (Comp(_,_),AtomF(_)) = GREATER
mengj@18403
   423
  | type_ord (Comp(con1,args1),Comp(con2,args2)) = 
mengj@18403
   424
    let val con_ord = string_ord(con1,con2)
mengj@18403
   425
    in
mengj@18403
   426
	case con_ord of EQUAL => types_ord (args1,args2)
mengj@18403
   427
		      | _ => con_ord
mengj@18403
   428
    end
mengj@18403
   429
and
mengj@18403
   430
mengj@18403
   431
types_ord ([],[]) = EQUAL
mengj@18403
   432
  | types_ord (tps1,tps2) = list_ord type_ord (tps1,tps2);
mengj@18403
   433
mengj@18403
   434
mengj@18403
   435
fun term_ord (UVar(_,_),UVar(_,_)) = EQUAL
mengj@18403
   436
  | term_ord (UVar(_,_),_) = LESS
mengj@18403
   437
  | term_ord (Fun(_,_,_),UVar(_)) = GREATER
mengj@18403
   438
  | term_ord (Fun(f1,tps1,tms1),Fun(f2,tps2,tms2)) = 
paulson@18420
   439
     (case string_ord (f1,f2) of
paulson@18420
   440
         EQUAL => 
paulson@18420
   441
	   (case terms_ord (tms1,tms2) of EQUAL => types_ord (tps1,tps2)
paulson@18420
   442
	      | tms_ord => tms_ord)
paulson@18420
   443
       | fn_ord => fn_ord)
mengj@18403
   444
mengj@18403
   445
and
mengj@18403
   446
paulson@18420
   447
  terms_ord ([],[]) = EQUAL
paulson@18420
   448
    | terms_ord (tms1,tms2) = list_ord term_ord (tms1,tms2);
mengj@18403
   449
mengj@18403
   450
mengj@18403
   451
paulson@18856
   452
fun predicate_ord (Predicate(pname1,ftyps1,ftms1),Predicate(pname2,ftyps2,ftms2)) = 
paulson@18856
   453
  case string_ord (pname1,pname2) of
paulson@18856
   454
       EQUAL => (case terms_ord(ftms1,ftms2) of EQUAL => types_ord(ftyps1,ftyps2)
paulson@18856
   455
				              | ftms_ord => ftms_ord)
paulson@18856
   456
     | pname_ord => pname_ord
paulson@18856
   457
			   
mengj@18402
   458
mengj@18402
   459
fun literal_ord (Literal(false,_,_),Literal(true,_,_)) = LESS
mengj@18402
   460
  | literal_ord (Literal(true,_,_),Literal(false,_,_)) = GREATER
mengj@18402
   461
  | literal_ord (Literal(_,pred1,_),Literal(_,pred2,_)) = predicate_ord(pred1,pred2);
mengj@18402
   462
mengj@18402
   463
fun sort_lits lits = sort literal_ord lits;
mengj@18402
   464
paulson@18420
   465
mengj@18409
   466
(********** clause equivalence ******************)
mengj@18409
   467
mengj@18409
   468
fun check_var_pairs (x,y) [] = 0 
mengj@18409
   469
  | check_var_pairs (x,y) ((u,v)::w) =
mengj@18409
   470
    if (x,y) = (u,v) then 1 
mengj@18409
   471
    else
mengj@18409
   472
	if (x = u) orelse (y = v) then 2 (*conflict*)
mengj@18409
   473
	else check_var_pairs (x,y) w;
mengj@18409
   474
mengj@18409
   475
fun type_eq (AtomV(v1),AtomV(v2)) (vars,tvars) =
mengj@18409
   476
    (case check_var_pairs (v1,v2) tvars of 0 => (true,(vars,(v1,v2)::tvars))
mengj@18409
   477
					 | 1 => (true,(vars,tvars))
mengj@18409
   478
					 | 2 => (false,(vars,tvars)))
mengj@18409
   479
  | type_eq (AtomV(_),_) vtvars = (false,vtvars)
mengj@18409
   480
  | type_eq (AtomF(f1),AtomF(f2)) vtvars = (f1=f2,vtvars)
mengj@18409
   481
  | type_eq (AtomF(_),_) vtvars = (false,vtvars)
mengj@18409
   482
  | type_eq (Comp(con1,args1),Comp(con2,args2)) vtvars =
paulson@18420
   483
      let val (eq1,vtvars1) = 
paulson@18420
   484
	      if con1 = con2 then types_eq (args1,args2) vtvars
paulson@18420
   485
	      else (false,vtvars)
paulson@18420
   486
      in
paulson@18420
   487
	  (eq1,vtvars1)
paulson@18420
   488
      end
mengj@18409
   489
  | type_eq (Comp(_,_),_) vtvars = (false,vtvars)
mengj@18409
   490
mengj@18409
   491
and
mengj@18409
   492
paulson@18420
   493
    types_eq ([],[]) vtvars = (true,vtvars)
paulson@18420
   494
  | types_eq (tp1::tps1,tp2::tps2) vtvars =
paulson@18420
   495
      let val (eq1,vtvars1) = type_eq (tp1,tp2) vtvars
paulson@18420
   496
	  val (eq2,vtvars2) = if eq1 then types_eq (tps1,tps2) vtvars1
paulson@18420
   497
			      else (eq1,vtvars1)
paulson@18420
   498
      in
paulson@18420
   499
	  (eq2,vtvars2)
paulson@18420
   500
      end;
mengj@18409
   501
mengj@18409
   502
mengj@18409
   503
fun term_eq (UVar(v1,tp1),UVar(v2,tp2)) (vars,tvars) =
mengj@18409
   504
    (case check_var_pairs (v1,v2) vars of 0 => type_eq (tp1,tp2) (((v1,v2)::vars),tvars)
mengj@18409
   505
					| 1 => type_eq (tp1,tp2) (vars,tvars)
mengj@18409
   506
					| 2 => (false,(vars,tvars)))
mengj@18409
   507
  | term_eq (UVar(_,_),_) vtvars = (false,vtvars)
mengj@18409
   508
  | term_eq (Fun(f1,tps1,tms1),Fun(f2,tps2,tms2)) vtvars =
paulson@18420
   509
      let val (eq1,vtvars1) = 
paulson@18420
   510
	      if f1 = f2 then terms_eq (tms1,tms2) vtvars
paulson@18420
   511
	      else (false,vtvars)
paulson@18420
   512
	  val (eq2,vtvars2) =
paulson@18420
   513
	      if eq1 then types_eq (tps1,tps2) vtvars1
paulson@18420
   514
	      else (eq1,vtvars1)
paulson@18420
   515
      in
paulson@18420
   516
	  (eq2,vtvars2)
paulson@18420
   517
      end
mengj@18409
   518
  | term_eq (Fun(_,_,_),_) vtvars = (false,vtvars)
mengj@18409
   519
mengj@18409
   520
and
mengj@18409
   521
paulson@18420
   522
    terms_eq ([],[]) vtvars = (true,vtvars)
paulson@18420
   523
  | terms_eq (tm1::tms1,tm2::tms2) vtvars =
paulson@18420
   524
      let val (eq1,vtvars1) = term_eq (tm1,tm2) vtvars
paulson@18420
   525
	  val (eq2,vtvars2) = if eq1 then terms_eq (tms1,tms2) vtvars1
paulson@18420
   526
				     else (eq1,vtvars1)
paulson@18420
   527
      in
paulson@18420
   528
	  (eq2,vtvars2)
paulson@18420
   529
      end;
mengj@18409
   530
					     
mengj@18409
   531
paulson@18856
   532
fun pred_eq (Predicate(pname1,tps1,tms1),Predicate(pname2,tps2,tms2)) vtvars =
mengj@18409
   533
    let val (eq1,vtvars1) = 
paulson@18856
   534
	    if (pname1 = pname2) then terms_eq (tms1,tms2) vtvars
mengj@18409
   535
	    else (false,vtvars)
mengj@18409
   536
	val (eq2,vtvars2) = 
mengj@18409
   537
	    if eq1 then types_eq (tps1,tps2) vtvars1
mengj@18409
   538
	    else (eq1,vtvars1)
mengj@18409
   539
    in
mengj@18409
   540
	(eq2,vtvars2)
mengj@18409
   541
    end;
mengj@18409
   542
					      
mengj@18409
   543
mengj@18409
   544
fun lit_eq (Literal(pol1,pred1,_),Literal(pol2,pred2,_)) vtvars =
mengj@18409
   545
    if (pol1 = pol2) then pred_eq (pred1,pred2) vtvars
mengj@18409
   546
    else (false,vtvars);
mengj@18409
   547
mengj@18409
   548
(*must have the same number of literals*)
mengj@18409
   549
fun lits_eq ([],[]) vtvars = (true,vtvars)
mengj@18409
   550
  | lits_eq (l1::ls1,l2::ls2) vtvars = 
mengj@18409
   551
    let val (eq1,vtvars1) = lit_eq (l1,l2) vtvars
mengj@18409
   552
    in
mengj@18409
   553
	if eq1 then lits_eq (ls1,ls2) vtvars1
mengj@18409
   554
	else (false,vtvars1)
mengj@18409
   555
    end;
mengj@18409
   556
mengj@18409
   557
paulson@18420
   558
(*Equality of two clauses up to variable renaming*)
paulson@18798
   559
fun clause_eq (Clause{literals=lits1,...}, Clause{literals=lits2,...}) =
paulson@18798
   560
  length lits1 = length lits2 andalso #1 (lits_eq (lits1,lits2) ([],[]));
mengj@18409
   561
mengj@18409
   562
paulson@18420
   563
(*** Hash function for clauses ***)
paulson@18420
   564
paulson@18420
   565
val xor_words = List.foldl Word.xorb 0w0;
paulson@18420
   566
paulson@18449
   567
fun hashw_term (UVar(_,_), w) = w
paulson@18449
   568
  | hashw_term (Fun(f,tps,args), w) = 
paulson@18449
   569
      List.foldl hashw_term (Polyhash.hashw_string (f,w)) args;
paulson@18420
   570
  
paulson@18449
   571
fun hashw_pred (Predicate(pn,_,args), w) = 
paulson@18449
   572
    List.foldl hashw_term (Polyhash.hashw_string (pn,w)) args;
paulson@18420
   573
    
paulson@18449
   574
fun hash1_literal (Literal(true,pred,_)) = hashw_pred (pred, 0w0)
paulson@18449
   575
  | hash1_literal (Literal(false,pred,_)) = Word.notb (hashw_pred (pred, 0w0));
paulson@18420
   576
  
paulson@18798
   577
fun hash_clause (Clause{literals,...}) =
paulson@18798
   578
  Word.toIntX (xor_words (map hash1_literal literals));
mengj@18402
   579
mengj@18402
   580
paulson@18798
   581
(*Make literals for sorted type variables.  FIXME: can it use map?*) 
quigley@17150
   582
fun sorts_on_typs (_, [])   = ([]) 
paulson@16199
   583
  | sorts_on_typs (v, "HOL.type" :: s) =
paulson@18411
   584
      sorts_on_typs (v,s)                (*IGNORE sort "type"*)
paulson@18798
   585
  | sorts_on_typs ((FOLTVar indx), s::ss) =
paulson@18798
   586
      LTVar(make_type_class s, make_schematic_type_var indx) :: 
paulson@18798
   587
      sorts_on_typs ((FOLTVar indx), ss)
paulson@18798
   588
  | sorts_on_typs ((FOLTFree x), s::ss) =
paulson@18798
   589
      LTFree(make_type_class s, make_fixed_type_var x) :: 
paulson@18798
   590
      sorts_on_typs ((FOLTFree x), ss);
paulson@15347
   591
quigley@17150
   592
paulson@18798
   593
fun pred_of_sort (LTVar (s,ty)) = (s,1)
paulson@18798
   594
|   pred_of_sort (LTFree (s,ty)) = (s,1)
quigley@17150
   595
paulson@16199
   596
(*Given a list of sorted type variables, return two separate lists.
paulson@16199
   597
  The first is for TVars, the second for TFrees.*)
paulson@18856
   598
fun add_typs_aux [] = ([],[])
paulson@18856
   599
  | add_typs_aux ((FOLTVar indx,s)::tss) = 
paulson@17230
   600
      let val vs = sorts_on_typs (FOLTVar indx, s)
paulson@18856
   601
	  val (vss,fss) = add_typs_aux tss
quigley@17150
   602
      in
paulson@18856
   603
	  (vs union vss, fss)
quigley@17150
   604
      end
paulson@18856
   605
  | add_typs_aux ((FOLTFree x,s)::tss) =
paulson@17230
   606
      let val fs = sorts_on_typs (FOLTFree x, s)
paulson@18856
   607
	  val (vss,fss) = add_typs_aux tss
quigley@17150
   608
      in
paulson@18856
   609
	  (vss, fs union fss)
quigley@17150
   610
      end;
quigley@17150
   611
mengj@17999
   612
paulson@18856
   613
fun add_typs (Clause cls) = add_typs_aux (#types_sorts cls)  
quigley@17150
   614
quigley@17150
   615
quigley@17150
   616
(** make axiom clauses, hypothesis clauses and conjecture clauses. **)
quigley@17150
   617
quigley@17150
   618
fun get_tvar_strs [] = []
paulson@17230
   619
  | get_tvar_strs ((FOLTVar indx,s)::tss) = 
paulson@17230
   620
      let val vstr = make_schematic_type_var indx
quigley@17150
   621
      in
paulson@17888
   622
	  vstr ins (get_tvar_strs tss)
quigley@17150
   623
      end
paulson@17230
   624
  | get_tvar_strs((FOLTFree x,s)::tss) = distinct (get_tvar_strs tss)
paulson@15347
   625
paulson@17230
   626
fun make_axiom_clause_thm thm (ax_name,cls_id) =
paulson@18856
   627
    let val (lits,types_sorts) = literals_of_term (prop_of thm)
mengj@18402
   628
	val lits' = sort_lits lits
paulson@18856
   629
	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts  
quigley@17150
   630
    in 
paulson@17230
   631
	make_clause(cls_id,ax_name,Axiom,
paulson@18856
   632
	            lits',types_sorts,tvar_lits,tfree_lits)
quigley@17150
   633
    end;
paulson@15347
   634
quigley@17150
   635
mengj@18199
   636
(* check if a clause is FOL first*)
paulson@17888
   637
fun make_conjecture_clause n t =
mengj@18199
   638
    let val _ = check_is_fol_term t
mengj@18199
   639
	    handle TERM("check_is_fol_term",_) => raise CLAUSE("Goal is not FOL",t)
paulson@18856
   640
	val (lits,types_sorts) = literals_of_term t
paulson@18856
   641
	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts 
paulson@15347
   642
    in
paulson@17845
   643
	make_clause(n,"conjecture",Conjecture,
paulson@18856
   644
	            lits,types_sorts,tvar_lits,tfree_lits)
paulson@15347
   645
    end;
paulson@17845
   646
    
paulson@17845
   647
fun make_conjecture_clauses_aux _ [] = []
paulson@17888
   648
  | make_conjecture_clauses_aux n (t::ts) =
paulson@17888
   649
      make_conjecture_clause n t :: make_conjecture_clauses_aux (n+1) ts
paulson@17845
   650
paulson@17845
   651
val make_conjecture_clauses = make_conjecture_clauses_aux 0
paulson@15347
   652
paulson@15347
   653
mengj@18199
   654
(*before converting an axiom clause to "clause" format, check if it is FOL*)
paulson@17230
   655
fun make_axiom_clause term (ax_name,cls_id) =
mengj@18199
   656
    let val _ = check_is_fol_term term 
mengj@18199
   657
	    handle TERM("check_is_fol_term",_) => raise CLAUSE("Axiom is not FOL", term) 
paulson@18856
   658
	val (lits,types_sorts) = literals_of_term term
mengj@18402
   659
	val lits' = sort_lits lits
paulson@18856
   660
	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
paulson@15347
   661
    in 
paulson@17230
   662
	make_clause(cls_id,ax_name,Axiom,
paulson@18856
   663
	            lits',types_sorts,tvar_lits,tfree_lits)
paulson@15347
   664
    end;
paulson@15347
   665
paulson@15347
   666
paulson@15347
   667
(**** Isabelle arities ****)
paulson@15347
   668
paulson@15347
   669
exception ARCLAUSE of string;
paulson@15347
   670
 
paulson@15347
   671
type class = string; 
paulson@15347
   672
type tcons = string; 
paulson@15347
   673
paulson@15347
   674
datatype arLit = TConsLit of bool * (class * tcons * string list) | TVarLit of bool * (class * string);
paulson@15347
   675
 
paulson@15347
   676
datatype arityClause =  
paulson@15347
   677
	 ArityClause of {clause_id: clause_id,
paulson@17845
   678
	  	         axiom_name: axiom_name,
paulson@15347
   679
			 kind: kind,
paulson@15347
   680
			 conclLit: arLit,
paulson@15347
   681
			 premLits: arLit list};
paulson@15347
   682
paulson@15347
   683
paulson@18798
   684
fun gen_TVars 0 = []
paulson@18798
   685
  | gen_TVars n = ("T_" ^ Int.toString n) :: gen_TVars (n-1);
paulson@15347
   686
paulson@18411
   687
fun pack_sort(_,[])  = []
paulson@18411
   688
  | pack_sort(tvar, "HOL.type"::srt) = pack_sort(tvar, srt)   (*IGNORE sort "type"*)
paulson@18411
   689
  | pack_sort(tvar, cls::srt) =  (make_type_class cls, tvar) :: pack_sort(tvar, srt);
paulson@15347
   690
    
paulson@15347
   691
fun make_TVarLit (b,(cls,str)) = TVarLit(b,(cls,str));
paulson@15347
   692
fun make_TConsLit (b,(cls,tcons,tvars)) = TConsLit(b,(make_type_class cls,make_fixed_type_const tcons,tvars));
paulson@15347
   693
paulson@18411
   694
(*Arity of type constructor tcon :: (arg1,...,argN)res*)
paulson@18411
   695
fun make_axiom_arity_clause (tcons, n, (res,args)) =
paulson@17845
   696
   let val nargs = length args
paulson@18798
   697
       val tvars = gen_TVars nargs
paulson@17845
   698
       val tvars_srts = ListPair.zip (tvars,args)
paulson@17845
   699
       val tvars_srts' = union_all(map pack_sort tvars_srts)
paulson@17845
   700
       val false_tvars_srts' = map (pair false) tvars_srts'
paulson@17845
   701
   in
paulson@17845
   702
      ArityClause {clause_id = n, kind = Axiom, 
paulson@18411
   703
                   axiom_name = lookup_type_const tcons,
paulson@18411
   704
                   conclLit = make_TConsLit(true, (res,tcons,tvars)), 
paulson@17845
   705
                   premLits = map make_TVarLit false_tvars_srts'}
paulson@17845
   706
   end;
paulson@15347
   707
    
paulson@18420
   708
(*The number of clauses generated from cls, including type clauses. It's always 1
paulson@18420
   709
  except for conjecture clauses.*)
paulson@17305
   710
fun num_of_clauses (Clause cls) =
paulson@17305
   711
    let val num_tfree_lits = 
paulson@17305
   712
	      if !keep_types then length (#tfree_type_literals cls)
paulson@17305
   713
	      else 0
paulson@17305
   714
    in 	1 + num_tfree_lits  end;
paulson@15347
   715
paulson@15347
   716
paulson@15347
   717
(**** Isabelle class relations ****)
paulson@15347
   718
paulson@15347
   719
datatype classrelClause = 
paulson@15347
   720
	 ClassrelClause of {clause_id: clause_id,
paulson@15347
   721
			    subclass: class,
paulson@18411
   722
			    superclass: class};
paulson@15347
   723
paulson@17845
   724
fun make_axiom_classrelClause n subclass superclass =
paulson@17845
   725
  ClassrelClause {clause_id = n,
paulson@17845
   726
                  subclass = subclass, superclass = superclass};
paulson@15347
   727
paulson@18411
   728
fun classrelClauses_of_aux n sub [] = []
paulson@18411
   729
  | classrelClauses_of_aux n sub ("HOL.type"::sups) = (*Should be ignored*)
paulson@18411
   730
      classrelClauses_of_aux n sub sups
paulson@18411
   731
  | classrelClauses_of_aux n sub (sup::sups) =
paulson@18411
   732
      ClassrelClause {clause_id = n, subclass = sub, superclass = sup} 
paulson@18411
   733
      :: classrelClauses_of_aux (n+1) sub sups;
paulson@15347
   734
paulson@18411
   735
fun classrelClauses_of (sub,sups) = classrelClauses_of_aux 0 sub sups;
paulson@17845
   736
paulson@17845
   737
paulson@17845
   738
(***** Isabelle arities *****)
paulson@17845
   739
paulson@17845
   740
fun arity_clause _ (tcons, []) = []
paulson@18411
   741
  | arity_clause n (tcons, ("HOL.type",_)::ars) =  (*Should be ignored*)
paulson@18411
   742
      arity_clause n (tcons,ars)
paulson@17845
   743
  | arity_clause n (tcons, ar::ars) =
paulson@17845
   744
      make_axiom_arity_clause (tcons,n,ar) :: 
paulson@17845
   745
      arity_clause (n+1) (tcons,ars);
paulson@17845
   746
paulson@17845
   747
fun multi_arity_clause [] = []
paulson@17845
   748
  | multi_arity_clause (tcon_ar :: tcons_ars)  =
paulson@17845
   749
      arity_clause 0 tcon_ar  @  multi_arity_clause tcons_ars 
paulson@17845
   750
paulson@17845
   751
fun arity_clause_thy thy =
paulson@17845
   752
  let val arities = #arities (Type.rep_tsig (Sign.tsig_of thy))
paulson@17845
   753
  in multi_arity_clause (Symtab.dest arities) end;
paulson@17845
   754
paulson@17845
   755
paulson@17845
   756
(* Isabelle classes *)
paulson@17845
   757
paulson@17845
   758
type classrelClauses = classrelClause list Symtab.table;
paulson@17845
   759
paulson@17845
   760
val classrel_of = #2 o #classes o Type.rep_tsig o Sign.tsig_of;
paulson@18798
   761
paulson@18798
   762
fun classrel_clauses_classrel (C: Sorts.classes) =
paulson@18798
   763
  map classrelClauses_of (Graph.dest C);
paulson@18798
   764
paulson@17845
   765
val classrel_clauses_thy = List.concat o classrel_clauses_classrel o classrel_of;
paulson@17845
   766
paulson@15347
   767
paulson@15347
   768
fun wrap_eq_type typ t = eq_typ_wrapper ^"(" ^ t ^ "," ^ typ ^ ")";
paulson@15347
   769
paulson@18218
   770
(*Only need to wrap equality's arguments with "typeinfo" if the output clauses are typed 
paulson@18218
   771
 and if we specifically ask for types to be included.   *)
paulson@15347
   772
fun string_of_equality (typ,terms) =
paulson@17230
   773
      let val [tstr1,tstr2] = map string_of_term terms
mengj@18402
   774
	  val typ' = string_of_fol_type typ
paulson@17230
   775
      in
paulson@17230
   776
	  if !keep_types andalso !special_equal 
mengj@18402
   777
	  then "equal(" ^ (wrap_eq_type typ' tstr1) ^ "," ^ 
mengj@18402
   778
		 	  (wrap_eq_type typ' tstr2) ^ ")"
paulson@17230
   779
	  else "equal(" ^ tstr1 ^ "," ^ tstr2 ^ ")"
paulson@17230
   780
      end
paulson@17230
   781
and string_of_term (UVar(x,_)) = x
paulson@18218
   782
  | string_of_term (Fun("equal",[typ],terms)) = string_of_equality(typ,terms)
paulson@18218
   783
  | string_of_term (Fun (name,typs,[])) = name (*Overloaded consts like 0 don't get types!*)
paulson@18218
   784
  | string_of_term (Fun (name,typs,terms)) = 
paulson@18218
   785
      let val terms_as_strings = map string_of_term terms
mengj@18402
   786
	  val typs' = if !keep_types then map string_of_fol_type typs else []
paulson@18420
   787
      in  name ^ (paren_pack (terms_as_strings @ typs'))  end;
paulson@15347
   788
paulson@15347
   789
(* before output the string of the predicate, check if the predicate corresponds to an equality or not. *)
paulson@18218
   790
fun string_of_predicate (Predicate("equal",[typ],terms)) = string_of_equality(typ,terms)
paulson@18218
   791
  | string_of_predicate (Predicate(name,typs,terms)) = 
paulson@17230
   792
      let val terms_as_strings = map string_of_term terms
mengj@18402
   793
	  val typs' = if !keep_types then map string_of_fol_type typs else []
paulson@18420
   794
      in  name ^ (paren_pack (terms_as_strings @ typs'))  end;
paulson@17317
   795
paulson@17317
   796
fun string_of_clausename (cls_id,ax_name) = 
paulson@17525
   797
    clause_prefix ^ ascii_of ax_name ^ "_" ^ Int.toString cls_id;
paulson@17317
   798
paulson@17317
   799
fun string_of_type_clsname (cls_id,ax_name,idx) = 
paulson@17525
   800
    string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx);
paulson@18863
   801
paulson@18863
   802
(*Write a list of strings to a file*)
paulson@18863
   803
fun writeln_strs os = List.app (fn s => TextIO.output (os,s));
paulson@18863
   804
quigley@17150
   805
    
quigley@17150
   806
quigley@17150
   807
(********************************)
quigley@17150
   808
(* Code for producing DFG files *)
quigley@17150
   809
(********************************)
quigley@17150
   810
paulson@18863
   811
(*Attach sign in DFG syntax: false means negate.*)
paulson@18863
   812
fun dfg_sign true s = s
paulson@18863
   813
  | dfg_sign false s = "not(" ^ s ^ ")"  
paulson@18863
   814
paulson@18863
   815
fun dfg_literal (Literal(pol,pred,tag)) = dfg_sign pol (string_of_predicate pred)
quigley@17150
   816
paulson@18798
   817
fun dfg_of_typeLit (LTVar (s,ty)) = "not(" ^ s ^ "(" ^ ty ^ "))"
paulson@18856
   818
  | dfg_of_typeLit (LTFree (s,ty)) = s ^ "(" ^ ty ^ ")";
quigley@17150
   819
 
paulson@17230
   820
(*Make the string of universal quantifiers for a clause*)
paulson@17230
   821
fun forall_open ([],[]) = ""
paulson@17230
   822
  | forall_open (vars,tvars) = "forall([" ^ (commas (tvars@vars))^ "],\n"
quigley@17150
   823
paulson@17230
   824
fun forall_close ([],[]) = ""
paulson@17230
   825
  | forall_close (vars,tvars) = ")"
quigley@17150
   826
paulson@17230
   827
fun gen_dfg_cls (cls_id,ax_name,knd,lits,tvars,vars) = 
paulson@17317
   828
    "clause( %(" ^ knd ^ ")\n" ^ forall_open(vars,tvars) ^ 
paulson@17317
   829
    "or(" ^ lits ^ ")" ^ forall_close(vars,tvars) ^ ",\n" ^ 
paulson@18863
   830
    string_of_clausename (cls_id,ax_name) ^  ").\n\n";
quigley@17150
   831
paulson@18856
   832
(*FIXME: UNUSED*)
paulson@17317
   833
fun gen_dfg_type_cls (cls_id,ax_name,knd,tfree_lit,idx,tvars,vars) = 
paulson@17317
   834
    "clause( %(" ^ knd ^ ")\n" ^ forall_open(vars,tvars) ^ 
paulson@17317
   835
    "or( " ^ tfree_lit ^ ")" ^ forall_close(vars,tvars) ^ ",\n" ^ 
paulson@18863
   836
    string_of_type_clsname (cls_id,ax_name,idx) ^  ").\n\n";
quigley@17150
   837
quigley@17150
   838
fun dfg_clause_aux (Clause cls) = 
paulson@17230
   839
  let val lits = map dfg_literal (#literals cls)
paulson@17230
   840
      val tvar_lits_strs = 
paulson@17230
   841
	  if !keep_types then map dfg_of_typeLit (#tvar_type_literals cls) 
paulson@17230
   842
	  else []
paulson@17230
   843
      val tfree_lits =
paulson@17230
   844
          if !keep_types then map dfg_of_typeLit (#tfree_type_literals cls)
paulson@17230
   845
          else []
paulson@17230
   846
  in
paulson@17234
   847
      (tvar_lits_strs @ lits, tfree_lits)
paulson@17230
   848
  end; 
quigley@17150
   849
quigley@17150
   850
quigley@17150
   851
fun dfg_folterms (Literal(pol,pred,tag)) = 
paulson@18856
   852
  let val Predicate (_, _, folterms) = pred
paulson@18218
   853
  in  folterms  end
quigley@17150
   854
paulson@17404
   855
fun get_uvars (UVar(a,typ)) = [a] 
paulson@17775
   856
|   get_uvars (Fun (_,typ,tlist)) = union_all(map get_uvars tlist)
quigley@17150
   857
paulson@17404
   858
fun is_uvar (UVar _) = true
paulson@17404
   859
|   is_uvar (Fun _) = false;
quigley@17150
   860
paulson@17404
   861
fun uvar_name (UVar(a,_)) = a
paulson@17404
   862
|   uvar_name (Fun (a,_,_)) = raise CLAUSE("Not a variable", Const(a,dummyT));
quigley@17150
   863
quigley@17150
   864
fun dfg_vars (Clause cls) =
paulson@17404
   865
    let val lits = #literals cls
paulson@18218
   866
        val folterms = List.concat (map dfg_folterms lits)
quigley@17150
   867
    in 
paulson@17775
   868
        union_all(map get_uvars folterms)
quigley@17150
   869
    end
quigley@17150
   870
paulson@18218
   871
fun string_of_predname (Predicate("equal",_,terms)) = "EQUALITY"
paulson@18218
   872
  | string_of_predname (Predicate(name,_,terms)) = name
quigley@17150
   873
paulson@17234
   874
fun dfg_pred (Literal(pol,pred,tag)) ax_name = 
paulson@17234
   875
    (string_of_predname pred) ^ " " ^ ax_name
quigley@17150
   876
paulson@18798
   877
fun clause2dfg (cls as Clause{axiom_name,clause_id,kind,types_sorts,...}) =
paulson@17230
   878
    let val (lits,tfree_lits) = dfg_clause_aux cls 
paulson@18798
   879
            (*"lits" includes the typing assumptions (TVars)*)
quigley@17150
   880
        val vars = dfg_vars cls
paulson@18798
   881
        val tvars = get_tvar_strs types_sorts
paulson@18798
   882
	val knd = name_of_kind kind
paulson@17234
   883
	val lits_str = commas lits
paulson@18798
   884
	val cls_str = gen_dfg_cls(clause_id,axiom_name,knd,lits_str,tvars,vars) 
paulson@18798
   885
    in (cls_str, tfree_lits) end;
quigley@17150
   886
paulson@18798
   887
fun string_of_arity (name, num) =  "(" ^ name ^ "," ^ Int.toString num ^ ")"
quigley@17150
   888
paulson@18856
   889
fun string_of_preds [] = ""
paulson@18856
   890
  | string_of_preds preds = "predicates[" ^ commas(map string_of_arity preds) ^ "].\n";
quigley@17150
   891
paulson@18856
   892
fun string_of_funcs [] = ""
paulson@18856
   893
  | string_of_funcs funcs = "functions[" ^ commas(map string_of_arity funcs) ^ "].\n" ;
quigley@17150
   894
paulson@17234
   895
fun string_of_symbols predstr funcstr = 
paulson@17234
   896
  "list_of_symbols.\n" ^ predstr  ^ funcstr  ^ "end_of_list.\n\n";
quigley@17150
   897
paulson@18863
   898
fun write_axioms (out, strs) = 
paulson@18863
   899
  (TextIO.output (out, "list_of_clauses(axioms,cnf).\n");
paulson@18863
   900
   writeln_strs out strs;
paulson@18863
   901
   TextIO.output (out, "end_of_list.\n\n"));
quigley@17150
   902
paulson@18863
   903
fun write_conjectures (out, strs) = 
paulson@18863
   904
  (TextIO.output (out, "list_of_clauses(conjectures,cnf).\n");
paulson@18863
   905
   writeln_strs out strs;
paulson@18863
   906
   TextIO.output (out, "end_of_list.\n\n"));
quigley@17150
   907
paulson@18798
   908
fun string_of_start name = "begin_problem(" ^ name ^ ").\n\n";
quigley@17150
   909
paulson@18863
   910
fun string_of_descrip name = 
paulson@18863
   911
  "list_of_descriptions.\nname({*" ^ name ^ "*}).\nauthor({*Isabelle*}).\nstatus(unknown).\ndescription({*auto-generated*}).\nend_of_list.\n\n"
quigley@17150
   912
paulson@18863
   913
fun dfg_tfree_clause tfree_lit =
paulson@18863
   914
  "clause( %(conjecture)\n" ^ "or( " ^ tfree_lit ^ "),\n" ^ "tfree_tcs" ^ ").\n\n"
quigley@17150
   915
paulson@18798
   916
paulson@18856
   917
(*** Find all occurrences of predicates in types, terms, literals... ***)
paulson@18856
   918
paulson@18856
   919
(*FIXME: multiple-arity checking doesn't work, as update_new is the wrong 
paulson@18856
   920
  function (it flags repeated declarations of a function, even with the same arity)*)
paulson@18856
   921
paulson@18856
   922
fun update_many (tab, keypairs) = foldl (uncurry Symtab.update) tab keypairs;
paulson@18856
   923
paulson@18856
   924
fun add_predicate_preds (Predicate(pname,tys,tms), preds) = 
paulson@18856
   925
  if pname = "equal" then preds (*equality is built-in and requires no declaration*)
paulson@18856
   926
  else Symtab.update (pname, length tys + length tms) preds
paulson@18856
   927
paulson@18856
   928
fun add_literal_preds (Literal(_,pred,_), preds) = add_predicate_preds (pred,preds)
paulson@18856
   929
paulson@18856
   930
fun add_type_sort_preds ((FOLTVar indx,s), preds) = 
paulson@18856
   931
      update_many (preds, map pred_of_sort (sorts_on_typs (FOLTVar indx, s)))
paulson@18856
   932
  | add_type_sort_preds ((FOLTFree x,s), preds) =
paulson@18856
   933
      update_many (preds, map pred_of_sort (sorts_on_typs (FOLTFree x, s)));
paulson@18856
   934
paulson@18856
   935
fun add_clause_preds (Clause {literals, types_sorts, ...}, preds) =
paulson@18856
   936
  foldl add_literal_preds (foldl add_type_sort_preds preds types_sorts) literals
paulson@18856
   937
  handle Symtab.DUP a => raise ERROR ("predicate " ^ a ^ " has multiple arities")
paulson@18856
   938
paulson@18856
   939
val preds_of_clauses = Symtab.dest o (foldl add_clause_preds Symtab.empty)
paulson@18856
   940
paulson@18856
   941
paulson@18856
   942
(*** Find all occurrences of functions in types, terms, literals... ***)
paulson@18798
   943
paulson@18798
   944
fun add_foltype_funcs (AtomV _, funcs) = funcs
paulson@18856
   945
  | add_foltype_funcs (AtomF a, funcs) = Symtab.update (a,0) funcs
paulson@18798
   946
  | add_foltype_funcs (Comp(a,tys), funcs) = 
paulson@18856
   947
      foldl add_foltype_funcs (Symtab.update (a, length tys) funcs) tys;
paulson@18798
   948
paulson@18798
   949
fun add_folterm_funcs (UVar _, funcs) = funcs
paulson@18856
   950
  | add_folterm_funcs (Fun(a,tys,[]), funcs) = Symtab.update (a,0) funcs
paulson@18798
   951
      (*A constant is a special case: it has no type argument even if overloaded*)
paulson@18798
   952
  | add_folterm_funcs (Fun(a,tys,tms), funcs) = 
paulson@18798
   953
      foldl add_foltype_funcs 
paulson@18856
   954
	    (foldl add_folterm_funcs (Symtab.update (a, length tys + length tms) funcs) 
paulson@18856
   955
	           tms) 
paulson@18798
   956
	    tys
paulson@18798
   957
paulson@18798
   958
fun add_predicate_funcs (Predicate(_,tys,tms), funcs) = 
paulson@18798
   959
    foldl add_foltype_funcs (foldl add_folterm_funcs funcs tms) tys;
paulson@18798
   960
paulson@18798
   961
fun add_literal_funcs (Literal(_,pred,_), funcs) = add_predicate_funcs (pred,funcs)
paulson@18798
   962
paulson@18798
   963
fun add_clause_funcs (Clause {literals, ...}, funcs) =
paulson@18856
   964
  foldl add_literal_funcs funcs literals
paulson@18856
   965
  handle Symtab.DUP a => raise ERROR ("function " ^ a ^ " has multiple arities")
paulson@18798
   966
paulson@18856
   967
val funcs_of_clauses = Symtab.dest o (foldl add_clause_funcs Symtab.empty)
paulson@18798
   968
paulson@17845
   969
fun string_of_arClauseID (ArityClause {clause_id,axiom_name,...}) =
paulson@17845
   970
    arclause_prefix ^ ascii_of axiom_name ^ "_" ^ Int.toString clause_id;
quigley@17150
   971
quigley@17150
   972
fun string_of_arKind (ArityClause arcls) = name_of_kind(#kind arcls);
quigley@17150
   973
paulson@18863
   974
fun dfg_of_arLit (TConsLit(pol,(c,t,args))) =
paulson@18863
   975
      dfg_sign pol (c ^ "(" ^ t ^ paren_pack args ^ ")")
paulson@18863
   976
  | dfg_of_arLit (TVarLit(pol,(c,str))) =
paulson@18863
   977
      dfg_sign pol (c ^ "(" ^ str ^ ")")
paulson@17525
   978
    
paulson@17525
   979
fun dfg_of_conclLit (ArityClause arcls) = dfg_of_arLit (#conclLit arcls);
paulson@17525
   980
     
paulson@17525
   981
fun dfg_of_premLits (ArityClause arcls) = map dfg_of_arLit (#premLits arcls);
paulson@17525
   982
		
paulson@18863
   983
fun dfg_classrelLits sub sup = 
paulson@18863
   984
    let val tvar = "(T)"
paulson@18863
   985
    in 
paulson@18863
   986
	"not(" ^ sub ^ tvar ^ "), " ^ sup ^ tvar
paulson@18863
   987
    end;
paulson@17525
   988
paulson@18863
   989
fun dfg_classrelClause (ClassrelClause {clause_id,subclass,superclass,...}) =
paulson@18863
   990
  let val relcls_id = clrelclause_prefix ^ ascii_of subclass ^ "_" ^ 
paulson@18863
   991
		      Int.toString clause_id
paulson@18863
   992
      val lits = dfg_classrelLits (make_type_class subclass) (make_type_class superclass)
paulson@18863
   993
  in
paulson@18863
   994
      "clause(\nor( " ^ lits ^ ")),\n" ^
paulson@18863
   995
      relcls_id ^ ").\n\n"
paulson@18863
   996
  end; 
quigley@17150
   997
quigley@17150
   998
fun dfg_arity_clause arcls = 
paulson@17525
   999
  let val arcls_id = string_of_arClauseID arcls
paulson@17525
  1000
      val concl_lit = dfg_of_conclLit arcls
paulson@17525
  1001
      val prems_lits = dfg_of_premLits arcls
paulson@17525
  1002
      val knd = string_of_arKind arcls
paulson@17525
  1003
      val all_lits = concl_lit :: prems_lits
paulson@17525
  1004
  in
paulson@18863
  1005
      "clause( %(" ^ knd ^ ")\n" ^  "or( " ^ bracket_pack all_lits ^ ")),\n" ^
paulson@18863
  1006
       arcls_id ^ ").\n\n"
paulson@18863
  1007
  end;
paulson@18863
  1008
paulson@18863
  1009
(* write out a subgoal in DFG format to the file "xxxx_N"*)
paulson@18863
  1010
fun dfg_write_file ths filename (axclauses,classrel_clauses,arity_clauses) = 
paulson@18863
  1011
  let 
paulson@18863
  1012
    val conjectures = make_conjecture_clauses (map prop_of ths)
paulson@18863
  1013
    val (dfg_clss,tfree_litss) = ListPair.unzip (map clause2dfg conjectures)
paulson@18863
  1014
    val classrel_cls = map dfg_classrelClause classrel_clauses
paulson@18863
  1015
    val arity_cls = map dfg_arity_clause arity_clauses
paulson@18863
  1016
    val clss = conjectures @ axclauses
paulson@18863
  1017
    val funcs = (funcs_of_clauses clss)
paulson@18863
  1018
    and preds = (preds_of_clauses clss)
paulson@18863
  1019
    val out = TextIO.openOut filename
paulson@18863
  1020
    and probname = Path.pack (Path.base (Path.unpack filename))
paulson@18863
  1021
    val axstrs_tfrees = (map clause2dfg axclauses)
paulson@18863
  1022
    val (axstrs, _) = ListPair.unzip axstrs_tfrees
paulson@18863
  1023
    val tfree_clss = map dfg_tfree_clause (union_all tfree_litss) 
paulson@18863
  1024
    val funcstr = string_of_funcs funcs
paulson@18863
  1025
    val predstr = string_of_preds preds
paulson@18863
  1026
  in
paulson@18863
  1027
      TextIO.output (out, string_of_start probname); 
paulson@18863
  1028
      TextIO.output (out, string_of_descrip probname); 
paulson@18863
  1029
      TextIO.output (out, string_of_symbols funcstr predstr); 
paulson@18863
  1030
      write_axioms (out, axstrs); 
paulson@18863
  1031
      write_conjectures (out, tfree_clss@dfg_clss); 
paulson@18863
  1032
      TextIO.output (out, "end_problem.\n");
paulson@18863
  1033
      TextIO.closeOut out
paulson@17525
  1034
  end;
quigley@17150
  1035
quigley@17150
  1036
quigley@17150
  1037
(********************************)
quigley@17150
  1038
(* code to produce TPTP files   *)
quigley@17150
  1039
(********************************)
quigley@17150
  1040
paulson@15347
  1041
fun tptp_literal (Literal(pol,pred,tag)) =
paulson@15347
  1042
    let val pred_string = string_of_predicate pred
paulson@17230
  1043
	val tagged_pol = 
paulson@17230
  1044
	      if (tag andalso !tagged) then (if pol then "+++" else "---")
paulson@17230
  1045
	      else (if pol then "++" else "--")
paulson@15347
  1046
     in
paulson@15347
  1047
	tagged_pol ^ pred_string
paulson@15347
  1048
    end;
paulson@15347
  1049
paulson@18798
  1050
fun tptp_of_typeLit (LTVar (s,ty)) = "--" ^ s ^ "(" ^ ty ^ ")"
paulson@18798
  1051
  | tptp_of_typeLit (LTFree (s,ty)) = "++" ^ s ^ "(" ^ ty ^ ")";
paulson@15347
  1052
 
paulson@15347
  1053
paulson@15347
  1054
fun gen_tptp_cls (cls_id,ax_name,knd,lits) = 
paulson@17317
  1055
    "input_clause(" ^ string_of_clausename (cls_id,ax_name) ^ "," ^ 
paulson@18863
  1056
    knd ^ "," ^ lits ^ ").\n";
paulson@15347
  1057
paulson@17317
  1058
fun gen_tptp_type_cls (cls_id,ax_name,knd,tfree_lit,idx) = 
paulson@17317
  1059
    "input_clause(" ^ string_of_type_clsname (cls_id,ax_name,idx) ^ "," ^ 
paulson@18863
  1060
    knd ^ ",[" ^ tfree_lit ^ "]).\n";
paulson@15347
  1061
paulson@17422
  1062
fun tptp_type_lits (Clause cls) = 
paulson@15347
  1063
    let val lits = map tptp_literal (#literals cls)
paulson@17230
  1064
	val tvar_lits_strs =
paulson@17230
  1065
	      if !keep_types 
paulson@17230
  1066
	      then (map tptp_of_typeLit (#tvar_type_literals cls)) 
paulson@17230
  1067
	      else []
paulson@17230
  1068
	val tfree_lits = 
paulson@17230
  1069
	      if !keep_types
paulson@17230
  1070
	      then (map tptp_of_typeLit (#tfree_type_literals cls)) 
paulson@17230
  1071
	      else []
paulson@15347
  1072
    in
paulson@17305
  1073
	(tvar_lits_strs @ lits, tfree_lits)
paulson@15347
  1074
    end; 
paulson@15347
  1075
paulson@15347
  1076
fun tptp_clause cls =
paulson@17422
  1077
    let val (lits,tfree_lits) = tptp_type_lits cls 
paulson@17305
  1078
            (*"lits" includes the typing assumptions (TVars)*)
paulson@17317
  1079
	val cls_id = get_clause_id cls
paulson@17317
  1080
	val ax_name = get_axiomName cls
paulson@15347
  1081
	val knd = string_of_kind cls
paulson@17525
  1082
	val lits_str = bracket_pack lits
paulson@17305
  1083
	val cls_str = gen_tptp_cls(cls_id,ax_name,knd,lits_str) 			 
paulson@17305
  1084
	fun typ_clss k [] = []
paulson@15347
  1085
          | typ_clss k (tfree :: tfrees) = 
paulson@17317
  1086
              gen_tptp_type_cls(cls_id,ax_name,knd,tfree,k) :: 
paulson@17317
  1087
              typ_clss (k+1) tfrees
paulson@15347
  1088
    in 
paulson@15347
  1089
	cls_str :: (typ_clss 0 tfree_lits)
paulson@15347
  1090
    end;
paulson@15347
  1091
paulson@15608
  1092
fun clause2tptp cls =
paulson@17422
  1093
    let val (lits,tfree_lits) = tptp_type_lits cls 
paulson@17305
  1094
            (*"lits" includes the typing assumptions (TVars)*)
paulson@17317
  1095
	val cls_id = get_clause_id cls
paulson@17317
  1096
	val ax_name = get_axiomName cls
paulson@15608
  1097
	val knd = string_of_kind cls
paulson@17525
  1098
	val lits_str = bracket_pack lits
paulson@15608
  1099
	val cls_str = gen_tptp_cls(cls_id,ax_name,knd,lits_str) 
paulson@15608
  1100
    in
paulson@15608
  1101
	(cls_str,tfree_lits) 
paulson@15608
  1102
    end;
paulson@15608
  1103
paulson@15608
  1104
paulson@18863
  1105
fun tptp_tfree_clause tfree_lit =
paulson@18863
  1106
    "input_clause(" ^ "tfree_tcs," ^ "conjecture" ^ ",[" ^ tfree_lit ^ "]).\n";
paulson@15608
  1107
paulson@15347
  1108
paulson@17525
  1109
fun tptp_of_arLit (TConsLit(b,(c,t,args))) =
paulson@17525
  1110
      let val pol = if b then "++" else "--"
paulson@18218
  1111
	  val  arg_strs = paren_pack args
paulson@17525
  1112
      in 
paulson@17525
  1113
	  pol ^ c ^ "(" ^ t ^ arg_strs ^ ")"
paulson@17525
  1114
      end
paulson@17525
  1115
  | tptp_of_arLit (TVarLit(b,(c,str))) =
paulson@17525
  1116
      let val pol = if b then "++" else "--"
paulson@17525
  1117
      in
paulson@17525
  1118
	  pol ^ c ^ "(" ^ str ^ ")"
paulson@17525
  1119
      end;
paulson@15347
  1120
    
paulson@15347
  1121
paulson@17525
  1122
fun tptp_of_conclLit (ArityClause arcls) = tptp_of_arLit (#conclLit arcls);
paulson@15347
  1123
     
paulson@17525
  1124
fun tptp_of_premLits (ArityClause arcls) = map tptp_of_arLit (#premLits arcls);
paulson@15347
  1125
		
paulson@15347
  1126
fun tptp_arity_clause arcls = 
paulson@15347
  1127
    let val arcls_id = string_of_arClauseID arcls
paulson@17525
  1128
	val concl_lit = tptp_of_conclLit arcls
paulson@17525
  1129
	val prems_lits = tptp_of_premLits arcls
paulson@15347
  1130
	val knd = string_of_arKind arcls
paulson@15347
  1131
	val all_lits = concl_lit :: prems_lits
paulson@15347
  1132
    in
paulson@17317
  1133
	"input_clause(" ^ arcls_id ^ "," ^ knd ^ "," ^ 
paulson@18863
  1134
	(bracket_pack all_lits) ^ ").\n"
paulson@15347
  1135
    end;
paulson@15347
  1136
paulson@15347
  1137
fun tptp_classrelLits sub sup = 
paulson@15347
  1138
    let val tvar = "(T)"
paulson@15347
  1139
    in 
paulson@18411
  1140
	"[--" ^ sub ^ tvar ^ ",++" ^ sup ^ tvar ^ "]"
paulson@15347
  1141
    end;
paulson@15347
  1142
paulson@17845
  1143
fun tptp_classrelClause (ClassrelClause {clause_id,subclass,superclass,...}) =
paulson@17845
  1144
    let val relcls_id = clrelclause_prefix ^ ascii_of subclass ^ "_" ^ 
paulson@17845
  1145
                        Int.toString clause_id
paulson@18411
  1146
	val lits = tptp_classrelLits (make_type_class subclass) (make_type_class superclass)
paulson@15347
  1147
    in
paulson@18863
  1148
	"input_clause(" ^ relcls_id ^ ",axiom," ^ lits ^ ").\n"
paulson@15347
  1149
    end; 
quigley@17150
  1150
paulson@18863
  1151
(* write out a subgoal as tptp clauses to the file "xxxx_N"*)
paulson@18863
  1152
fun tptp_write_file ths filename (axclauses,classrel_clauses,arity_clauses) =
paulson@18863
  1153
  let
paulson@18863
  1154
    val clss = make_conjecture_clauses (map prop_of ths)
paulson@18863
  1155
    val (tptp_clss,tfree_litss) = ListPair.unzip (map clause2tptp clss)
paulson@18863
  1156
    val tfree_clss = map tptp_tfree_clause (foldl (op union_string) [] tfree_litss)
paulson@18863
  1157
    val classrel_cls = map tptp_classrelClause classrel_clauses
paulson@18863
  1158
    val arity_cls = map tptp_arity_clause arity_clauses
paulson@18863
  1159
    val out = TextIO.openOut filename
paulson@18863
  1160
  in
paulson@18863
  1161
    List.app (writeln_strs out o tptp_clause) axclauses;
paulson@18863
  1162
    writeln_strs out tfree_clss;
paulson@18863
  1163
    writeln_strs out tptp_clss;
paulson@18863
  1164
    writeln_strs out classrel_cls;
paulson@18863
  1165
    writeln_strs out arity_cls;
paulson@18863
  1166
    TextIO.closeOut out
paulson@18863
  1167
  end;
paulson@18863
  1168
paulson@18863
  1169
paulson@15347
  1170
end;