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