src/HOL/Tools/res_clause.ML
author paulson
Thu, 28 Jul 2005 17:54:22 +0200
changeset 16953 f025e0dc638b
parent 16925 0fd7b1438d28
child 16976 377962871f5d
permissions -rw-r--r--
uniform treatment of variable prefixes
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@15347
    11
    exception ARCLAUSE of string
paulson@15347
    12
    exception CLAUSE of string
paulson@15347
    13
    type arityClause 
paulson@15347
    14
    type classrelClause
paulson@16925
    15
    val classrelClauses_of : string * string list -> classrelClause list
paulson@15347
    16
    type clause
paulson@16925
    17
    val init : theory -> unit
paulson@15347
    18
    val keep_types : bool ref
paulson@15347
    19
    val make_axiom_arity_clause :
paulson@15347
    20
       string * (string * string list list) -> arityClause
paulson@15347
    21
    val make_axiom_classrelClause :
skalberg@15531
    22
       string * string option -> classrelClause
paulson@15347
    23
    val make_axiom_clause : Term.term -> string * int -> clause
paulson@15347
    24
    val make_conjecture_clause : Term.term -> clause
paulson@15347
    25
    val make_conjecture_clause_thm : Thm.thm -> clause
paulson@15347
    26
    val make_hypothesis_clause : Term.term -> clause
paulson@15347
    27
    val special_equal : bool ref
paulson@15347
    28
    val tptp_arity_clause : arityClause -> string
paulson@15347
    29
    val tptp_classrelClause : classrelClause -> string
paulson@15347
    30
    val tptp_clause : clause -> string list
quigley@16039
    31
    val clause_info : clause ->  string * string
paulson@15347
    32
    val tptp_clauses2str : string list -> string
paulson@15347
    33
    val typed : unit -> unit
paulson@15347
    34
    val untyped : unit -> unit
paulson@15608
    35
    val clause2tptp : clause -> string * string list
paulson@15608
    36
    val tfree_clause : string -> string
paulson@16953
    37
    val schematic_var_prefix : string
paulson@16953
    38
    val fixed_var_prefix : string
paulson@16953
    39
    val tvar_prefix : string
paulson@16953
    40
    val tfree_prefix : string
paulson@16953
    41
    val clause_prefix : string 
paulson@16953
    42
    val arclause_prefix : string
paulson@16953
    43
    val const_prefix : string
paulson@16953
    44
    val tconst_prefix : string 
paulson@16953
    45
    val class_prefix : string 
paulson@15347
    46
  end;
paulson@15347
    47
paulson@15347
    48
structure ResClause : RES_CLAUSE =
paulson@15347
    49
struct
paulson@15347
    50
paulson@15347
    51
(* Added for typed equality *)
paulson@15347
    52
val special_equal = ref false; (* by default,equality does not carry type information *)
paulson@15347
    53
val eq_typ_wrapper = "typeinfo"; (* default string *)
paulson@15347
    54
paulson@15347
    55
paulson@15347
    56
val schematic_var_prefix = "V_";
paulson@15347
    57
val fixed_var_prefix = "v_";
paulson@15347
    58
paulson@16903
    59
val tvar_prefix = "T_";
paulson@16903
    60
val tfree_prefix = "t_";
paulson@15347
    61
paulson@15347
    62
val clause_prefix = "cls_"; 
paulson@15347
    63
paulson@15347
    64
val arclause_prefix = "arcls_" 
paulson@15347
    65
paulson@16903
    66
val const_prefix = "c_";
paulson@16903
    67
val tconst_prefix = "tc_"; 
paulson@15347
    68
paulson@16199
    69
val class_prefix = "class_"; 
paulson@15347
    70
paulson@15347
    71
paulson@15347
    72
(**** some useful functions ****)
paulson@15347
    73
 
paulson@15347
    74
val const_trans_table =
paulson@15347
    75
      Symtab.make [("op =", "equal"),
paulson@15347
    76
	  	   ("op <=", "lessequals"),
paulson@15347
    77
		   ("op <", "less"),
paulson@15347
    78
		   ("op &", "and"),
paulson@15347
    79
		   ("op |", "or"),
paulson@15347
    80
		   ("op -->", "implies"),
paulson@15347
    81
		   ("op :", "in"),
paulson@15347
    82
		   ("op Un", "union"),
paulson@15347
    83
		   ("op Int", "inter")];
paulson@15347
    84
paulson@16903
    85
val type_const_trans_table =
paulson@16903
    86
      Symtab.make [("*", "t_prod"),
paulson@16903
    87
	  	   ("+", "t_sum"),
paulson@16903
    88
		   ("~=>", "t_map")];
paulson@15347
    89
paulson@15610
    90
(*Escaping of special characters.
paulson@15610
    91
  Alphanumeric characters are left unchanged.
paulson@15610
    92
  The character _ goes to __
paulson@15610
    93
  Characters in the range ASCII space to / go to _A to _P, respectively.
paulson@15610
    94
  Other printing characters go to _NNN where NNN is the decimal ASCII code.*)
paulson@15610
    95
local
paulson@15610
    96
paulson@15610
    97
val A_minus_space = Char.ord #"A" - Char.ord #" ";
paulson@15610
    98
paulson@15347
    99
fun ascii_of_c c =
paulson@15610
   100
  if Char.isAlphaNum c then String.str c
paulson@15610
   101
  else if c = #"_" then "__"
paulson@15610
   102
  else if #" " <= c andalso c <= #"/" 
paulson@15610
   103
       then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
paulson@15610
   104
  else if Char.isPrint c then ("_" ^ Int.toString (Char.ord c))
paulson@15610
   105
  else ""
paulson@15347
   106
paulson@15610
   107
in
paulson@15610
   108
paulson@15610
   109
val ascii_of = String.translate ascii_of_c;
paulson@15610
   110
paulson@15610
   111
end;
paulson@15347
   112
paulson@16925
   113
(*Remove the initial ' character from a type variable, if it is present*)
paulson@16925
   114
fun trim_type_var s =
paulson@16925
   115
  if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
paulson@16925
   116
  else error ("trim_type: Malformed type variable encountered: " ^ s);
paulson@16925
   117
paulson@16903
   118
fun ascii_of_indexname (v,0) = ascii_of v
paulson@16903
   119
  | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ string_of_int i;
paulson@15347
   120
paulson@16903
   121
fun make_schematic_var v = schematic_var_prefix ^ (ascii_of_indexname v);
paulson@15347
   122
fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x);
paulson@15347
   123
paulson@16903
   124
(*Type variables contain _H because the character ' translates to that.*)
paulson@16925
   125
fun make_schematic_type_var (x,i) = 
paulson@16925
   126
      tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
paulson@16925
   127
fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
paulson@15347
   128
paulson@16903
   129
fun make_fixed_const c =
paulson@16903
   130
    case Symtab.lookup (const_trans_table,c) of
paulson@16903
   131
        SOME c' => c'
paulson@16903
   132
      | NONE =>  const_prefix ^ (ascii_of c);
paulson@16903
   133
paulson@16903
   134
fun make_fixed_type_const c = 
paulson@16903
   135
    case Symtab.lookup (type_const_trans_table,c) of
paulson@16903
   136
        SOME c' => c'
paulson@16903
   137
      | NONE =>  tconst_prefix ^ (ascii_of c);
paulson@15347
   138
paulson@15347
   139
fun make_type_class clas = class_prefix ^ (ascii_of clas);
paulson@15347
   140
paulson@15347
   141
paulson@15347
   142
paulson@15347
   143
(***** definitions and functions for FOL clauses, prepared for conversion into TPTP format or SPASS format. *****)
paulson@15347
   144
paulson@15347
   145
val keep_types = ref true; (* default is true *)
paulson@15347
   146
fun untyped () = (keep_types := false);
paulson@15347
   147
fun typed () = (keep_types := true);
paulson@15347
   148
paulson@15347
   149
paulson@15347
   150
datatype kind = Axiom | Hypothesis | Conjecture;
paulson@15347
   151
fun name_of_kind Axiom = "axiom"
paulson@15347
   152
  | name_of_kind Hypothesis = "hypothesis"
paulson@15347
   153
  | name_of_kind Conjecture = "conjecture";
paulson@15347
   154
paulson@15347
   155
type clause_id = int;
paulson@15347
   156
type axiom_name = string;
paulson@15347
   157
paulson@15347
   158
paulson@15347
   159
type polarity = bool;
paulson@15347
   160
paulson@15347
   161
type indexname = Term.indexname;
paulson@15347
   162
paulson@15347
   163
paulson@15347
   164
(* "tag" is used for vampire specific syntax  *)
paulson@15347
   165
type tag = bool; 
paulson@15347
   166
paulson@15347
   167
paulson@16903
   168
val id_ref = ref 0;
paulson@15347
   169
paulson@15347
   170
fun generate_id () = 
paulson@16903
   171
    let val id = !id_ref
paulson@16903
   172
    in id_ref := id + 1; id end;
paulson@15347
   173
paulson@15347
   174
paulson@15347
   175
paulson@15347
   176
(**** Isabelle FOL clauses ****)
paulson@15347
   177
paulson@15347
   178
(* by default it is false *)
paulson@15347
   179
val tagged = ref false;
paulson@15347
   180
paulson@15347
   181
type pred_name = string;
paulson@15347
   182
type sort = Term.sort;
paulson@15347
   183
type fol_type = string;
paulson@15347
   184
paulson@15347
   185
paulson@15347
   186
datatype type_literal = LTVar of string | LTFree of string;
paulson@15347
   187
paulson@15347
   188
paulson@15347
   189
datatype folTerm = UVar of string * fol_type| Fun of string * fol_type * folTerm list;
paulson@15347
   190
datatype predicate = Predicate of pred_name * fol_type * folTerm list;
paulson@15347
   191
paulson@15347
   192
paulson@15347
   193
datatype literal = Literal of polarity * predicate * tag;
paulson@15347
   194
paulson@15347
   195
paulson@15347
   196
datatype typ_var = FOLTVar of indexname | FOLTFree of string;
paulson@15347
   197
paulson@15347
   198
paulson@15347
   199
(* ML datatype used to repsent one single clause: disjunction of literals. *)
paulson@15347
   200
datatype clause = 
paulson@15347
   201
	 Clause of {clause_id: clause_id,
paulson@15347
   202
		    axiom_name: axiom_name,
paulson@15347
   203
		    kind: kind,
paulson@15347
   204
		    literals: literal list,
paulson@15347
   205
		    types_sorts: (typ_var * sort) list, 
paulson@15347
   206
                    tvar_type_literals: type_literal list, 
paulson@15347
   207
                    tfree_type_literals: type_literal list };
paulson@15347
   208
paulson@15347
   209
paulson@15347
   210
exception CLAUSE of string;
paulson@15347
   211
paulson@15347
   212
paulson@15347
   213
paulson@15347
   214
(*** make clauses ***)
paulson@15347
   215
paulson@15347
   216
paulson@16903
   217
fun make_clause (clause_id,axiom_name,kind,literals,
paulson@16903
   218
                 types_sorts,tvar_type_literals,
paulson@16903
   219
                 tfree_type_literals) =
paulson@16903
   220
     Clause {clause_id = clause_id, axiom_name = axiom_name, kind = kind, 
paulson@16903
   221
             literals = literals, types_sorts = types_sorts,
paulson@16903
   222
             tvar_type_literals = tvar_type_literals,
paulson@16903
   223
             tfree_type_literals = tfree_type_literals};
paulson@15347
   224
paulson@16925
   225
paulson@16925
   226
(*Definitions of the current theory--to allow suppressing types.*)
paulson@16925
   227
val curr_defs = ref Defs.empty;
paulson@16925
   228
paulson@16925
   229
(*Initialize the type suppression mechanism with the current theory before
paulson@16925
   230
    producing any clauses!*)
paulson@16925
   231
fun init thy = (curr_defs := Theory.defs_of thy);
paulson@16925
   232
paulson@16925
   233
(*Types aren't needed if the constant has at most one definition and is monomorphic*)
paulson@16925
   234
fun no_types_needed s =
paulson@16925
   235
  (case Defs.fast_overloading_info (!curr_defs) s of
paulson@16925
   236
      NONE => true
paulson@16925
   237
    | SOME (T,len,_) => len <= 1 andalso null (typ_tvars T))
paulson@16925
   238
    
paulson@16925
   239
(*Flatten a type to a string while accumulating sort constraints on the TFress and
paulson@16925
   240
  TVars it contains.*)    
paulson@15347
   241
fun type_of (Type (a, [])) = (make_fixed_type_const a,[]) 
paulson@15347
   242
  | type_of (Type (a, Ts)) = 
paulson@16903
   243
      let val foltyps_ts = map type_of Ts
paulson@16903
   244
	  val (folTyps,ts) = ListPair.unzip foltyps_ts
paulson@16903
   245
	  val ts' = ResLib.flat_noDup ts
paulson@16903
   246
      in    
paulson@16903
   247
	  (((make_fixed_type_const a) ^ (ResLib.list_to_string folTyps)), ts') 
paulson@16903
   248
      end
paulson@15347
   249
  | type_of (TFree (a, s)) = (make_fixed_type_var a, [((FOLTFree a),s)])
paulson@16903
   250
  | type_of (TVar (v, s))  = (make_schematic_type_var v, [((FOLTVar v),s)]);
paulson@15390
   251
paulson@16925
   252
fun maybe_type_of c T =
paulson@16925
   253
 if no_types_needed c then ("",[]) else type_of T;
paulson@16925
   254
paulson@16903
   255
(* Any variables created via the METAHYPS tactical should be treated as
paulson@16903
   256
   universal vars, although it is represented as "Free(...)" by Isabelle *)
paulson@16903
   257
val isMeta = String.isPrefix "METAHYP1_"
paulson@15347
   258
    
paulson@16925
   259
fun pred_name_type (Const(c,T)) = (make_fixed_const c, maybe_type_of c T)
paulson@15390
   260
  | pred_name_type (Free(x,T))  = 
paulson@16903
   261
      if isMeta x then raise CLAUSE("Predicate Not First Order") 
paulson@16903
   262
      else (make_fixed_var x, type_of T)
paulson@15347
   263
  | pred_name_type (Var(_,_))   = raise CLAUSE("Predicate Not First Order")
paulson@15347
   264
  | pred_name_type _          = raise CLAUSE("Predicate input unexpected");
paulson@15347
   265
paulson@15615
   266
paulson@15615
   267
(* For type equality *)
paulson@15615
   268
(* here "arg_typ" is the type of "="'s argument's type, not the type of the equality *)
paulson@15615
   269
(* Find type of equality arg *)
paulson@15615
   270
fun eq_arg_type (Type("fun",[T,_])) = 
paulson@15615
   271
    let val (folT,_) = type_of T;
paulson@15615
   272
    in
paulson@15615
   273
	folT
paulson@15615
   274
    end;
paulson@15615
   275
paulson@16925
   276
fun fun_name_type (Const(c,T)) = (make_fixed_const c, maybe_type_of c T)
paulson@15347
   277
  | fun_name_type (Free(x,T)) = (make_fixed_var x,type_of T)
paulson@15347
   278
  | fun_name_type _ = raise CLAUSE("Function Not First Order");
paulson@15347
   279
    
paulson@15347
   280
quigley@16767
   281
(* FIX - add in funcs and stuff to this *)
paulson@15615
   282
paulson@15347
   283
fun term_of (Var(ind_nm,T)) = 
paulson@16903
   284
      let val (folType,ts) = type_of T
paulson@16903
   285
      in
paulson@16903
   286
	  (UVar(make_schematic_var ind_nm, folType), ts)
paulson@16903
   287
      end
paulson@15347
   288
  | term_of (Free(x,T)) = 
paulson@16903
   289
      let val (folType,ts) = type_of T
paulson@16903
   290
      in
paulson@16903
   291
	  if isMeta x then (UVar(make_schematic_var(x,0), folType), ts)
paulson@16903
   292
	  else (Fun(make_fixed_var x,folType,[]), ts)
paulson@16903
   293
      end
paulson@15615
   294
  | term_of (Const(c,T)) =  (* impossible to be equality *)
paulson@16903
   295
      let val (folType,ts) = type_of T
paulson@16903
   296
      in
paulson@16903
   297
	  (Fun(make_fixed_const c,folType,[]), ts)
paulson@16903
   298
      end    
paulson@15347
   299
  | term_of (app as (t $ a)) = 
paulson@16903
   300
      let val (f,args) = strip_comb app
paulson@16903
   301
	  fun term_of_aux () = 
paulson@16903
   302
	      let val (funName,(funType,ts1)) = fun_name_type f
paulson@16903
   303
		   val (args',ts2) = ListPair.unzip (map term_of args)
paulson@16903
   304
		   val ts3 = ResLib.flat_noDup (ts1::ts2)
paulson@16903
   305
	      in
paulson@16903
   306
		  (Fun(funName,funType,args'),ts3)
paulson@16903
   307
	      end
paulson@16903
   308
	  fun term_of_eq ((Const ("op =", typ)),args) =
paulson@16903
   309
	      let val arg_typ = eq_arg_type typ
paulson@16903
   310
		  val (args',ts) = ListPair.unzip (map term_of args)
paulson@16903
   311
		  val equal_name = make_fixed_const ("op =")
paulson@16903
   312
	      in
paulson@16903
   313
		  (Fun(equal_name,arg_typ,args'),ResLib.flat_noDup ts)
paulson@16903
   314
	      end
paulson@16903
   315
      in
paulson@16903
   316
	  case f of Const ("op =", typ) => term_of_eq (f,args)
paulson@16903
   317
		  | Const(_,_) => term_of_aux ()
paulson@16903
   318
		  | Free(s,_)  => if isMeta s 
paulson@16903
   319
		                  then raise CLAUSE("Function Not First Order") 
paulson@16903
   320
		                  else term_of_aux()
paulson@16903
   321
		  | _          => raise CLAUSE("Function Not First Order")
paulson@16903
   322
      end
paulson@15390
   323
  | term_of _ = raise CLAUSE("Function Not First Order"); 
paulson@15390
   324
paulson@15347
   325
paulson@15347
   326
fun pred_of_eq ((Const ("op =", typ)),args) =
paulson@15347
   327
    let val arg_typ = eq_arg_type typ 
paulson@15774
   328
	val (args',ts) = ListPair.unzip (map term_of args)
paulson@16903
   329
	val equal_name = make_fixed_const "op ="
paulson@15347
   330
    in
paulson@15347
   331
	(Predicate(equal_name,arg_typ,args'),ResLib.flat_noDup ts)
paulson@15347
   332
    end;
paulson@15347
   333
paulson@15347
   334
paulson@15347
   335
(* changed for non-equality predicate *)
paulson@15347
   336
(* The input "pred" cannot be an equality *)
paulson@15347
   337
fun pred_of_nonEq (pred,args) = 
paulson@15347
   338
    let val (predName,(predType,ts1)) = pred_name_type pred
paulson@15774
   339
	val (args',ts2) = ListPair.unzip (map term_of args)
paulson@15347
   340
	val ts3 = ResLib.flat_noDup (ts1::ts2)
paulson@15347
   341
    in
paulson@15347
   342
	(Predicate(predName,predType,args'),ts3)
paulson@15347
   343
    end;
paulson@15347
   344
paulson@15347
   345
paulson@15347
   346
(* Changed for typed equality *)
paulson@15347
   347
(* First check if the predicate is an equality or not, then call different functions for equality and non-equalities *)
paulson@15347
   348
fun predicate_of term =
paulson@15347
   349
    let val (pred,args) = strip_comb term
paulson@15347
   350
    in
paulson@15347
   351
	case pred of (Const ("op =", _)) => pred_of_eq (pred,args)
paulson@15347
   352
		   | _ => pred_of_nonEq (pred,args)
paulson@15347
   353
    end;
paulson@15347
   354
paulson@15347
   355
 
paulson@16925
   356
fun literals_of_term ((Const("Trueprop",_) $ P),lits_ts) = literals_of_term (P,lits_ts)
paulson@15347
   357
  | literals_of_term ((Const("op |",_) $ P $ Q),(lits,ts)) = 
paulson@16925
   358
      let val (lits',ts') = literals_of_term (P,(lits,ts))
paulson@16925
   359
      in
paulson@16925
   360
	  literals_of_term (Q, (lits',ts'))
paulson@16925
   361
      end
paulson@15347
   362
  | literals_of_term ((Const("Not",_) $ P),(lits,ts)) = 
paulson@16925
   363
      let val (pred,ts') = predicate_of P
paulson@16925
   364
	  val lits' = Literal(false,pred,false) :: lits
paulson@16925
   365
	  val ts'' = ResLib.no_rep_app ts ts'
paulson@16925
   366
      in
paulson@16925
   367
	  (lits',ts'')
paulson@16925
   368
      end
paulson@15347
   369
  | literals_of_term (P,(lits,ts)) = 
paulson@16925
   370
      let val (pred,ts') = predicate_of P
paulson@16925
   371
	  val lits' = Literal(true,pred,false) :: lits
paulson@16925
   372
	  val ts'' = ResLib.no_rep_app ts ts' 
paulson@16925
   373
      in
paulson@16925
   374
	  (lits',ts'')
paulson@16925
   375
      end;
paulson@15347
   376
     
paulson@15956
   377
fun literals_of_thm thm = literals_of_term (prop_of thm, ([],[]));
paulson@15347
   378
    
paulson@16199
   379
(*Make literals for sorted type variables*) 
paulson@15347
   380
fun sorts_on_typs (_, []) = []
paulson@16199
   381
  | sorts_on_typs (v, "HOL.type" :: s) =
paulson@16199
   382
      sorts_on_typs (v,s)   (*Ignore sort "type"*)
paulson@16199
   383
  | sorts_on_typs ((FOLTVar(indx)), (s::ss)) =
paulson@16199
   384
      LTVar((make_type_class s) ^ 
paulson@16903
   385
        "(" ^ (make_schematic_type_var indx) ^ ")") :: 
paulson@16199
   386
      (sorts_on_typs ((FOLTVar(indx)), ss))
paulson@16199
   387
  | sorts_on_typs ((FOLTFree(x)), (s::ss)) =
paulson@16199
   388
      LTFree((make_type_class s) ^ "(" ^ (make_fixed_type_var(x)) ^ ")") :: 
paulson@16199
   389
      (sorts_on_typs ((FOLTFree(x)), ss));
paulson@15347
   390
paulson@16199
   391
(*Given a list of sorted type variables, return two separate lists.
paulson@16199
   392
  The first is for TVars, the second for TFrees.*)
paulson@15347
   393
fun add_typs_aux [] = ([],[])
paulson@15347
   394
  | add_typs_aux ((FOLTVar(indx),s)::tss) = 
paulson@16199
   395
      let val vs = sorts_on_typs (FOLTVar(indx), s)
paulson@16199
   396
	  val (vss,fss) = add_typs_aux tss
paulson@16199
   397
      in
paulson@16199
   398
	  (ResLib.no_rep_app vs vss, fss)
paulson@16199
   399
      end
paulson@15347
   400
  | add_typs_aux ((FOLTFree(x),s)::tss) =
paulson@16199
   401
      let val fs = sorts_on_typs (FOLTFree(x), s)
paulson@16199
   402
	  val (vss,fss) = add_typs_aux tss
paulson@16199
   403
      in
paulson@16199
   404
	  (vss, ResLib.no_rep_app fs fss)
paulson@16199
   405
      end;
paulson@15347
   406
paulson@16199
   407
fun add_typs (Clause cls) = add_typs_aux (#types_sorts cls)
paulson@15347
   408
paulson@15347
   409
paulson@15347
   410
(** make axiom clauses, hypothesis clauses and conjecture clauses. **)
paulson@15347
   411
   
paulson@15347
   412
paulson@15347
   413
fun make_conjecture_clause_thm thm =
paulson@15347
   414
    let val (lits,types_sorts) = literals_of_thm thm
paulson@15347
   415
	val cls_id = generate_id()
paulson@15347
   416
	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
paulson@15347
   417
    in
paulson@15347
   418
	make_clause(cls_id,"",Conjecture,lits,types_sorts,tvar_lits,tfree_lits)
paulson@15347
   419
    end;
paulson@15347
   420
paulson@15347
   421
paulson@16903
   422
fun make_axiom_clause term (axname,cls_id) =
paulson@15347
   423
    let val (lits,types_sorts) = literals_of_term (term,([],[]))
paulson@15347
   424
	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
paulson@15347
   425
    in 
paulson@16903
   426
	make_clause(cls_id,axname,Axiom,lits,types_sorts,tvar_lits,tfree_lits)
paulson@15347
   427
    end;
paulson@15347
   428
paulson@15347
   429
paulson@15347
   430
fun make_hypothesis_clause term =
paulson@15347
   431
    let val (lits,types_sorts) = literals_of_term (term,([],[]))
paulson@15347
   432
	val cls_id = generate_id()
paulson@15347
   433
	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
paulson@15347
   434
    in
paulson@15347
   435
	make_clause(cls_id,"",Hypothesis,lits,types_sorts,tvar_lits,tfree_lits)
paulson@15347
   436
    end;
paulson@15347
   437
    
paulson@15347
   438
 
paulson@15347
   439
fun make_conjecture_clause term =
paulson@15347
   440
    let val (lits,types_sorts) = literals_of_term (term,([],[]))
paulson@15347
   441
	val cls_id = generate_id()
paulson@15347
   442
	val (tvar_lits,tfree_lits) = add_typs_aux types_sorts
paulson@15347
   443
    in
paulson@15347
   444
	make_clause(cls_id,"",Conjecture,lits,types_sorts,tvar_lits,tfree_lits)
paulson@15347
   445
    end;
paulson@15347
   446
 
paulson@15347
   447
paulson@15347
   448
 
paulson@15347
   449
(**** Isabelle arities ****)
paulson@15347
   450
paulson@15347
   451
exception ARCLAUSE of string;
paulson@15347
   452
 
paulson@15347
   453
paulson@15347
   454
type class = string; 
paulson@15347
   455
type tcons = string; 
paulson@15347
   456
paulson@15347
   457
paulson@15347
   458
datatype arLit = TConsLit of bool * (class * tcons * string list) | TVarLit of bool * (class * string);
paulson@15347
   459
 
paulson@15347
   460
datatype arityClause =  
paulson@15347
   461
	 ArityClause of {clause_id: clause_id,
paulson@15347
   462
			 kind: kind,
paulson@15347
   463
			 conclLit: arLit,
paulson@15347
   464
			 premLits: arLit list};
paulson@15347
   465
paulson@15347
   466
paulson@15347
   467
fun get_TVars 0 = []
paulson@15347
   468
  | get_TVars n = ("T_" ^ (string_of_int n)) :: get_TVars (n-1);
paulson@15347
   469
paulson@15347
   470
paulson@15347
   471
paulson@15347
   472
fun pack_sort(_,[])  = raise ARCLAUSE("Empty Sort Found") 
paulson@15347
   473
  | pack_sort(tvar, [cls]) = [(make_type_class cls, tvar)] 
paulson@15347
   474
  | pack_sort(tvar, cls::srt) =  (make_type_class cls,tvar) :: (pack_sort(tvar, srt));
paulson@15347
   475
    
paulson@15347
   476
    
paulson@15347
   477
fun make_TVarLit (b,(cls,str)) = TVarLit(b,(cls,str));
paulson@15347
   478
fun make_TConsLit (b,(cls,tcons,tvars)) = TConsLit(b,(make_type_class cls,make_fixed_type_const tcons,tvars));
paulson@15347
   479
paulson@15347
   480
paulson@15347
   481
fun make_arity_clause (clause_id,kind,conclLit,premLits) =
paulson@15347
   482
    ArityClause {clause_id = clause_id, kind = kind, conclLit = conclLit, premLits = premLits};
paulson@15347
   483
paulson@15347
   484
paulson@15347
   485
fun make_axiom_arity_clause (tcons,(res,args)) =
paulson@15347
   486
     let val cls_id = generate_id()
paulson@15347
   487
	 val nargs = length args
paulson@15347
   488
	 val tvars = get_TVars nargs
paulson@15347
   489
	 val conclLit = make_TConsLit(true,(res,tcons,tvars))
paulson@15774
   490
         val tvars_srts = ListPair.zip (tvars,args)
paulson@15347
   491
	 val tvars_srts' = ResLib.flat_noDup(map pack_sort tvars_srts)
paulson@15347
   492
         val false_tvars_srts' = ResLib.pair_ins false tvars_srts'
paulson@15347
   493
	 val premLits = map make_TVarLit false_tvars_srts'
paulson@15347
   494
     in
paulson@15347
   495
	 make_arity_clause (cls_id,Axiom,conclLit,premLits)
paulson@15347
   496
     end;
paulson@15347
   497
    
paulson@15347
   498
paulson@15347
   499
paulson@15347
   500
(**** Isabelle class relations ****)
paulson@15347
   501
paulson@15347
   502
paulson@15347
   503
datatype classrelClause = 
paulson@15347
   504
	 ClassrelClause of {clause_id: clause_id,
paulson@15347
   505
			    subclass: class,
skalberg@15531
   506
			    superclass: class option};
paulson@15347
   507
paulson@15347
   508
fun make_classrelClause (clause_id,subclass,superclass) =
paulson@15347
   509
    ClassrelClause {clause_id = clause_id,subclass = subclass, superclass = superclass};
paulson@15347
   510
paulson@15347
   511
paulson@15347
   512
fun make_axiom_classrelClause (subclass,superclass) =
paulson@15347
   513
    let val cls_id = generate_id()
paulson@15347
   514
	val sub = make_type_class subclass
skalberg@15531
   515
	val sup = case superclass of NONE => NONE 
skalberg@15531
   516
				   | SOME s => SOME (make_type_class s)
paulson@15347
   517
    in
paulson@15347
   518
	make_classrelClause(cls_id,sub,sup)
paulson@15347
   519
    end;
paulson@15347
   520
paulson@15347
   521
paulson@15347
   522
paulson@15347
   523
fun classrelClauses_of_aux (sub,[]) = []
skalberg@15531
   524
  | classrelClauses_of_aux (sub,(sup::sups)) = make_axiom_classrelClause(sub,SOME sup) :: (classrelClauses_of_aux (sub,sups));
paulson@15347
   525
paulson@15347
   526
paulson@15347
   527
fun classrelClauses_of (sub,sups) = 
skalberg@15531
   528
    case sups of [] => [make_axiom_classrelClause (sub,NONE)]
paulson@15347
   529
	       | _ => classrelClauses_of_aux (sub, sups);
paulson@15347
   530
paulson@15347
   531
paulson@15347
   532
paulson@15347
   533
(***** convert clauses to tptp format *****)
paulson@15347
   534
paulson@15347
   535
paulson@16903
   536
fun string_of_clauseID (Clause cls) = 
paulson@16903
   537
    clause_prefix ^ string_of_int (#clause_id cls);
paulson@15347
   538
paulson@15347
   539
fun string_of_kind (Clause cls) = name_of_kind (#kind cls);
paulson@15347
   540
paulson@15347
   541
fun string_of_axiomName (Clause cls) = #axiom_name cls;
paulson@15347
   542
paulson@15347
   543
(****!!!! Changed for typed equality !!!!****)
paulson@15347
   544
fun wrap_eq_type typ t = eq_typ_wrapper ^"(" ^ t ^ "," ^ typ ^ ")";
paulson@15347
   545
paulson@15347
   546
paulson@15347
   547
(****!!!! Changed for typed equality !!!!****)
paulson@15347
   548
(* Only need to wrap equality's arguments with "typeinfo" if the output clauses are typed && if we specifically ask for types to be included.   *)
paulson@15347
   549
fun string_of_equality (typ,terms) =
paulson@15347
   550
    let val [tstr1,tstr2] = map string_of_term terms
paulson@15347
   551
    in
paulson@15347
   552
	if ((!keep_types) andalso (!special_equal)) then 
paulson@15347
   553
	    "equal(" ^ (wrap_eq_type typ tstr1) ^ "," ^ (wrap_eq_type typ tstr2) ^ ")"
paulson@15347
   554
	else
paulson@15347
   555
	    "equal(" ^ tstr1 ^ "," ^ tstr2 ^ ")"
paulson@15615
   556
    end
paulson@15615
   557
paulson@15615
   558
and
paulson@15615
   559
    string_of_term (UVar(x,_)) = x
paulson@15615
   560
  | string_of_term (Fun("equal",typ,terms)) = string_of_equality(typ,terms)
paulson@15615
   561
  | string_of_term (Fun (name,typ,[])) = name
paulson@15615
   562
  | string_of_term (Fun (name,typ,terms)) = 
paulson@15615
   563
    let val terms' = map string_of_term terms
paulson@15615
   564
    in
paulson@16925
   565
        if !keep_types andalso typ<>"" then name ^ (ResLib.list_to_string (terms' @ [typ]))
paulson@15615
   566
        else name ^ (ResLib.list_to_string terms')
paulson@15347
   567
    end;
paulson@15347
   568
paulson@15347
   569
paulson@15347
   570
paulson@15347
   571
(* Changed for typed equality *)
paulson@15347
   572
(* before output the string of the predicate, check if the predicate corresponds to an equality or not. *)
paulson@16903
   573
fun string_of_predicate (Predicate("equal",typ,terms)) = 
paulson@16903
   574
       string_of_equality(typ,terms)
paulson@15347
   575
  | string_of_predicate (Predicate(name,_,[])) = name 
paulson@15347
   576
  | string_of_predicate (Predicate(name,typ,terms)) = 
paulson@16903
   577
      let val terms_as_strings = map string_of_term terms
paulson@16903
   578
      in
paulson@16925
   579
	  if !keep_types andalso typ<>""
paulson@16903
   580
	  then name ^ (ResLib.list_to_string  (terms_as_strings @ [typ]))
paulson@16903
   581
	  else name ^ (ResLib.list_to_string terms_as_strings) 
paulson@16903
   582
      end;
paulson@15347
   583
paulson@15347
   584
    
paulson@15347
   585
paulson@15347
   586
paulson@15347
   587
fun tptp_literal (Literal(pol,pred,tag)) =
paulson@15347
   588
    let val pred_string = string_of_predicate pred
paulson@16903
   589
	val tagged_pol = 
paulson@16903
   590
	      if (tag andalso !tagged) then (if pol then "+++" else "---")
paulson@16903
   591
	      else (if pol then "++" else "--")
paulson@15347
   592
     in
paulson@15347
   593
	tagged_pol ^ pred_string
paulson@15347
   594
    end;
paulson@15347
   595
paulson@15347
   596
paulson@15347
   597
paulson@15347
   598
fun tptp_of_typeLit (LTVar x) = "--" ^ x
paulson@15347
   599
  | tptp_of_typeLit (LTFree x) = "++" ^ x;
paulson@15347
   600
 
paulson@15347
   601
paulson@15347
   602
fun gen_tptp_cls (cls_id,ax_name,knd,lits) = 
paulson@16903
   603
    let val ax_str = (if ax_name = "" then "" else ("_" ^ ascii_of ax_name))
paulson@15347
   604
    in
paulson@15347
   605
	"input_clause(" ^ cls_id ^ ax_str ^ "," ^ knd ^ "," ^ lits ^ ")."
paulson@15347
   606
    end;
paulson@15347
   607
paulson@16903
   608
fun gen_tptp_type_cls (cls_id,knd,tfree_lit,idx) = 
paulson@16903
   609
    "input_clause(" ^ cls_id ^ "_tcs" ^ (string_of_int idx) ^ "," ^ 
paulson@16903
   610
    knd ^ ",[" ^ tfree_lit ^ "]).";
paulson@15347
   611
paulson@15347
   612
paulson@15347
   613
fun tptp_clause_aux (Clause cls) = 
paulson@15347
   614
    let val lits = map tptp_literal (#literals cls)
paulson@16903
   615
	val tvar_lits_strs =
paulson@16903
   616
	      if (!keep_types) 
paulson@16903
   617
	      then (map tptp_of_typeLit (#tvar_type_literals cls)) 
paulson@16903
   618
	      else []
paulson@16903
   619
	val tfree_lits = 
paulson@16903
   620
	      if (!keep_types) 
paulson@16903
   621
	      then (map tptp_of_typeLit (#tfree_type_literals cls)) 
paulson@16903
   622
	      else []
paulson@15347
   623
    in
paulson@15347
   624
	(tvar_lits_strs @ lits,tfree_lits)
paulson@15347
   625
    end; 
paulson@15347
   626
paulson@15608
   627
paulson@15347
   628
fun tptp_clause cls =
paulson@15347
   629
    let val (lits,tfree_lits) = tptp_clause_aux cls (*"lits" includes the typing assumptions (TVars)*)
paulson@15347
   630
	val cls_id = string_of_clauseID cls
paulson@15347
   631
	val ax_name = string_of_axiomName cls
paulson@15347
   632
	val knd = string_of_kind cls
paulson@15347
   633
	val lits_str = ResLib.list_to_string' lits
paulson@15347
   634
	val cls_str = gen_tptp_cls(cls_id,ax_name,knd,lits_str) 			fun typ_clss k [] = []
paulson@15347
   635
          | typ_clss k (tfree :: tfrees) = 
paulson@15347
   636
            (gen_tptp_type_cls(cls_id,knd,tfree,k)) ::  (typ_clss (k+1) tfrees)
paulson@15347
   637
    in 
paulson@15347
   638
	cls_str :: (typ_clss 0 tfree_lits)
paulson@15347
   639
    end;
paulson@15347
   640
paulson@16794
   641
fun clause_info cls = (string_of_axiomName cls, string_of_clauseID cls);
quigley@16039
   642
paulson@15347
   643
paulson@15608
   644
fun clause2tptp cls =
paulson@15608
   645
    let val (lits,tfree_lits) = tptp_clause_aux cls (*"lits" includes the typing assumptions (TVars)*)
paulson@15608
   646
	val cls_id = string_of_clauseID cls
paulson@15608
   647
	val ax_name = string_of_axiomName cls
paulson@15608
   648
	val knd = string_of_kind cls
paulson@15608
   649
	val lits_str = ResLib.list_to_string' lits
paulson@15608
   650
	val cls_str = gen_tptp_cls(cls_id,ax_name,knd,lits_str) 
paulson@15608
   651
    in
paulson@15608
   652
	(cls_str,tfree_lits) 
paulson@15608
   653
    end;
paulson@15608
   654
paulson@15608
   655
paulson@16903
   656
fun tfree_clause tfree_lit =
paulson@16903
   657
    "input_clause(" ^ "tfree_tcs," ^ "conjecture" ^ ",[" ^ tfree_lit ^ "]).";
paulson@15608
   658
paulson@15347
   659
val delim = "\n";
paulson@15347
   660
val tptp_clauses2str = ResLib.list2str_sep delim; 
paulson@15347
   661
     
paulson@15347
   662
paulson@15347
   663
fun string_of_arClauseID (ArityClause arcls) = arclause_prefix ^ string_of_int(#clause_id arcls);
paulson@15347
   664
paulson@15347
   665
paulson@15347
   666
fun string_of_arLit (TConsLit(b,(c,t,args))) =
paulson@15347
   667
    let val pol = if b then "++" else "--"
paulson@15347
   668
	val  arg_strs = (case args of [] => "" | _ => ResLib.list_to_string args)
paulson@15347
   669
    in 
paulson@15347
   670
	pol ^ c ^ "(" ^ t ^ arg_strs ^ ")"
paulson@15347
   671
    end
paulson@15347
   672
  | string_of_arLit (TVarLit(b,(c,str))) =
paulson@15347
   673
    let val pol = if b then "++" else "--"
paulson@15347
   674
    in
paulson@15347
   675
	pol ^ c ^ "(" ^ str ^ ")"
paulson@15347
   676
    end;
paulson@15347
   677
    
paulson@15347
   678
paulson@15347
   679
fun string_of_conclLit (ArityClause arcls) = string_of_arLit (#conclLit arcls);
paulson@15347
   680
     
paulson@15347
   681
paulson@15347
   682
fun strings_of_premLits (ArityClause arcls) = map string_of_arLit (#premLits arcls);
paulson@15347
   683
		
paulson@15347
   684
paulson@15347
   685
fun string_of_arKind (ArityClause arcls) = name_of_kind(#kind arcls);
paulson@15347
   686
paulson@15347
   687
fun tptp_arity_clause arcls = 
paulson@15347
   688
    let val arcls_id = string_of_arClauseID arcls
paulson@15347
   689
	val concl_lit = string_of_conclLit arcls
paulson@15347
   690
	val prems_lits = strings_of_premLits arcls
paulson@15347
   691
	val knd = string_of_arKind arcls
paulson@15347
   692
	val all_lits = concl_lit :: prems_lits
paulson@15347
   693
    in
paulson@15452
   694
	"input_clause(" ^ arcls_id ^ "," ^ knd ^ "," ^ (ResLib.list_to_string' all_lits) ^ ")."
paulson@15347
   695
	
paulson@15347
   696
    end;
paulson@15347
   697
paulson@15347
   698
paulson@15347
   699
val clrelclause_prefix = "relcls_";
paulson@15347
   700
paulson@15347
   701
paulson@15347
   702
fun tptp_classrelLits sub sup = 
paulson@15347
   703
    let val tvar = "(T)"
paulson@15347
   704
    in 
skalberg@15531
   705
	case sup of NONE => "[++" ^ sub ^ tvar ^ "]"
skalberg@15531
   706
		  | (SOME supcls) =>  "[--" ^ sub ^ tvar ^ ",++" ^ supcls ^ tvar ^ "]"
paulson@15347
   707
    end;
paulson@15347
   708
paulson@15347
   709
paulson@15347
   710
fun tptp_classrelClause (ClassrelClause cls) =
paulson@15347
   711
    let val relcls_id = clrelclause_prefix ^ string_of_int(#clause_id cls)
paulson@15347
   712
	val sub = #subclass cls
paulson@15347
   713
	val sup = #superclass cls
paulson@15347
   714
	val lits = tptp_classrelLits sub sup
paulson@15347
   715
    in
paulson@15347
   716
	"input_clause(" ^ relcls_id ^ ",axiom," ^ lits ^ ")."
paulson@15347
   717
    end; 
paulson@15347
   718
    
paulson@15347
   719
end;