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