src/HOL/Tools/res_clause.ML
author wenzelm
Fri, 17 Aug 2007 23:10:41 +0200
changeset 24310 af4af9993922
parent 24183 a46b758941a4
child 24322 dc7336b8c54c
permissions -rw-r--r--
proper signature;
removed unused const_types_of;
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
wenzelm@24310
     5
Storing/printing FOL clauses and arity clauses.
paulson@15347
     6
Typed equality is treated differently.
paulson@15347
     7
*)
paulson@15347
     8
paulson@23385
     9
(*FIXME: combine with res_hol_clause!*)
paulson@15347
    10
signature RES_CLAUSE =
wenzelm@24310
    11
sig
wenzelm@24310
    12
  val schematic_var_prefix: string
wenzelm@24310
    13
  val fixed_var_prefix: string
wenzelm@24310
    14
  val tvar_prefix: string
wenzelm@24310
    15
  val tfree_prefix: string
wenzelm@24310
    16
  val clause_prefix: string
wenzelm@24310
    17
  val const_prefix: string
wenzelm@24310
    18
  val tconst_prefix: string
wenzelm@24310
    19
  val class_prefix: string
wenzelm@24310
    20
  val union_all: ''a list list -> ''a list
wenzelm@24310
    21
  val const_trans_table: string Symtab.table
wenzelm@24310
    22
  val type_const_trans_table: string Symtab.table
wenzelm@24310
    23
  val ascii_of: string -> string
wenzelm@24310
    24
  val undo_ascii_of: string -> string
wenzelm@24310
    25
  val paren_pack : string list -> string
wenzelm@24310
    26
  val make_schematic_var : string * int -> string
wenzelm@24310
    27
  val make_fixed_var : string -> string
wenzelm@24310
    28
  val make_schematic_type_var : string * int -> string
wenzelm@24310
    29
  val make_fixed_type_var : string -> string
wenzelm@24310
    30
  val dfg_format: bool ref
wenzelm@24310
    31
  val make_fixed_const : string -> string
wenzelm@24310
    32
  val make_fixed_type_const : string -> string
wenzelm@24310
    33
  val make_type_class : string -> string
wenzelm@24310
    34
  datatype kind = Axiom | Conjecture
wenzelm@24310
    35
  val name_of_kind : kind -> string
wenzelm@24310
    36
  type axiom_name = string
wenzelm@24310
    37
  datatype typ_var = FOLTVar of indexname | FOLTFree of string
wenzelm@24310
    38
  datatype fol_type =
wenzelm@24310
    39
      AtomV of string
wenzelm@24310
    40
    | AtomF of string
wenzelm@24310
    41
    | Comp of string * fol_type list
wenzelm@24310
    42
  val string_of_fol_type : fol_type -> string
wenzelm@24310
    43
  datatype type_literal = LTVar of string * string | LTFree of string * string
wenzelm@24310
    44
  val mk_typ_var_sort: typ -> typ_var * sort
paulson@18868
    45
  exception CLAUSE of string * term
paulson@18868
    46
  val init : theory -> unit
paulson@18868
    47
  val isMeta : string -> bool
wenzelm@24310
    48
  val add_typs_aux : (typ_var * string list) list -> type_literal list * type_literal list
wenzelm@24310
    49
  val get_tvar_strs: (typ_var * sort) list -> string list
wenzelm@24310
    50
  datatype arLit =
wenzelm@24310
    51
      TConsLit of class * string * string list
wenzelm@24310
    52
    | TVarLit of class * string
wenzelm@24310
    53
  datatype arityClause = ArityClause of
wenzelm@24310
    54
   {axiom_name: axiom_name,
wenzelm@24310
    55
    kind: kind,
wenzelm@24310
    56
    conclLit: arLit,
wenzelm@24310
    57
    premLits: arLit list}
wenzelm@24310
    58
  datatype classrelClause = ClassrelClause of
wenzelm@24310
    59
   {axiom_name: axiom_name,
wenzelm@24310
    60
    subclass: class,
wenzelm@24310
    61
    superclass: class}
wenzelm@24310
    62
  val make_classrel_clauses: theory -> class list -> class list -> classrelClause list
wenzelm@24310
    63
  val make_arity_clauses: theory -> string list -> class list -> class list * arityClause list
wenzelm@24310
    64
  val add_type_sort_preds: (typ_var * string list) * int Symtab.table -> int Symtab.table
wenzelm@24310
    65
  val add_classrelClause_preds : classrelClause * int Symtab.table -> int Symtab.table
wenzelm@24310
    66
  val class_of_arityLit: arLit -> class
wenzelm@24310
    67
  val add_arityClause_preds: arityClause * int Symtab.table -> int Symtab.table
wenzelm@24310
    68
  val add_foltype_funcs: fol_type * int Symtab.table -> int Symtab.table
wenzelm@24310
    69
  val add_arityClause_funcs: arityClause * int Symtab.table -> int Symtab.table
wenzelm@24310
    70
  val init_functab: int Symtab.table
wenzelm@24310
    71
  val writeln_strs: TextIO.outstream -> string list -> unit
wenzelm@24310
    72
  val dfg_sign: bool -> string -> string
wenzelm@24310
    73
  val dfg_of_typeLit: type_literal -> string
wenzelm@24310
    74
  val gen_dfg_cls: int * string * string * string * string list -> string
wenzelm@24310
    75
  val string_of_preds: (string * Int.int) list -> string
wenzelm@24310
    76
  val string_of_funcs: (string * int) list -> string
wenzelm@24310
    77
  val string_of_symbols: string -> string -> string
wenzelm@24310
    78
  val string_of_start: string -> string
wenzelm@24310
    79
  val string_of_descrip : string -> string
wenzelm@24310
    80
  val dfg_tfree_clause : string -> string
wenzelm@24310
    81
  val dfg_classrelClause: classrelClause -> string
wenzelm@24310
    82
  val dfg_arity_clause: arityClause -> string
wenzelm@24310
    83
  val tptp_sign: bool -> string -> string
wenzelm@24310
    84
  val tptp_of_typeLit : type_literal -> string
wenzelm@24310
    85
  val gen_tptp_cls : int * string * kind * string list -> string
wenzelm@24310
    86
  val tptp_tfree_clause : string -> string
paulson@17404
    87
  val tptp_arity_clause : arityClause -> string
paulson@17404
    88
  val tptp_classrelClause : classrelClause -> string
wenzelm@24310
    89
end
paulson@15347
    90
wenzelm@24310
    91
structure ResClause: RES_CLAUSE =
paulson@15347
    92
struct
paulson@15347
    93
paulson@15347
    94
val schematic_var_prefix = "V_";
paulson@15347
    95
val fixed_var_prefix = "v_";
paulson@15347
    96
paulson@17230
    97
val tvar_prefix = "T_";
paulson@17230
    98
val tfree_prefix = "t_";
paulson@15347
    99
wenzelm@24310
   100
val clause_prefix = "cls_";
wenzelm@24310
   101
val arclause_prefix = "clsarity_"
paulson@17525
   102
val clrelclause_prefix = "clsrel_";
paulson@15347
   103
paulson@17230
   104
val const_prefix = "c_";
wenzelm@24310
   105
val tconst_prefix = "tc_";
wenzelm@24310
   106
val class_prefix = "class_";
paulson@15347
   107
paulson@17775
   108
fun union_all xss = foldl (op union) [] xss;
paulson@17775
   109
paulson@17775
   110
(*Provide readable names for the more common symbolic functions*)
paulson@15347
   111
val const_trans_table =
paulson@15347
   112
      Symtab.make [("op =", "equal"),
wenzelm@24310
   113
                   (@{const_name HOL.less_eq}, "lessequals"),
wenzelm@24310
   114
                   (@{const_name HOL.less}, "less"),
wenzelm@24310
   115
                   ("op &", "and"),
wenzelm@24310
   116
                   ("op |", "or"),
wenzelm@24310
   117
                   (@{const_name HOL.plus}, "plus"),
wenzelm@24310
   118
                   (@{const_name HOL.minus}, "minus"),
wenzelm@24310
   119
                   (@{const_name HOL.times}, "times"),
wenzelm@24310
   120
                   (@{const_name Divides.div}, "div"),
wenzelm@24310
   121
                   (@{const_name HOL.divide}, "divide"),
wenzelm@24310
   122
                   ("op -->", "implies"),
wenzelm@24310
   123
                   ("{}", "emptyset"),
wenzelm@24310
   124
                   ("op :", "in"),
wenzelm@24310
   125
                   ("op Un", "union"),
wenzelm@24310
   126
                   ("op Int", "inter"),
wenzelm@24310
   127
                   ("List.append", "append"),
wenzelm@24310
   128
                   ("ATP_Linkup.fequal", "fequal"),
wenzelm@24310
   129
                   ("ATP_Linkup.COMBI", "COMBI"),
wenzelm@24310
   130
                   ("ATP_Linkup.COMBK", "COMBK"),
wenzelm@24310
   131
                   ("ATP_Linkup.COMBB", "COMBB"),
wenzelm@24310
   132
                   ("ATP_Linkup.COMBC", "COMBC"),
wenzelm@24310
   133
                   ("ATP_Linkup.COMBS", "COMBS"),
wenzelm@24310
   134
                   ("ATP_Linkup.COMBB'", "COMBB_e"),
wenzelm@24310
   135
                   ("ATP_Linkup.COMBC'", "COMBC_e"),
wenzelm@24310
   136
                   ("ATP_Linkup.COMBS'", "COMBS_e")];
paulson@15347
   137
paulson@17230
   138
val type_const_trans_table =
paulson@18411
   139
      Symtab.make [("*", "prod"),
wenzelm@24310
   140
                   ("+", "sum"),
wenzelm@24310
   141
                   ("~=>", "map")];
paulson@15347
   142
paulson@15610
   143
(*Escaping of special characters.
paulson@15610
   144
  Alphanumeric characters are left unchanged.
paulson@15610
   145
  The character _ goes to __
paulson@15610
   146
  Characters in the range ASCII space to / go to _A to _P, respectively.
paulson@24183
   147
  Other printing characters go to _nnn where nnn is the decimal ASCII code.*)
paulson@24183
   148
val A_minus_space = Char.ord #"A" - Char.ord #" ";
paulson@15610
   149
paulson@24183
   150
fun stringN_of_int 0 _ = ""
paulson@24183
   151
  | stringN_of_int k n = stringN_of_int (k-1) (n div 10) ^ Int.toString (n mod 10);
paulson@15610
   152
paulson@15347
   153
fun ascii_of_c c =
paulson@15610
   154
  if Char.isAlphaNum c then String.str c
paulson@15610
   155
  else if c = #"_" then "__"
wenzelm@24310
   156
  else if #" " <= c andalso c <= #"/"
paulson@15610
   157
       then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
wenzelm@24310
   158
  else if Char.isPrint c
paulson@24183
   159
       then ("_" ^ stringN_of_int 3 (Char.ord c))  (*fixed width, in case more digits follow*)
paulson@15610
   160
  else ""
paulson@15347
   161
paulson@15610
   162
val ascii_of = String.translate ascii_of_c;
paulson@15610
   163
paulson@24183
   164
(** Remove ASCII armouring from names in proof files **)
paulson@24183
   165
paulson@24183
   166
(*We don't raise error exceptions because this code can run inside the watcher.
paulson@24183
   167
  Also, the errors are "impossible" (hah!)*)
paulson@24183
   168
fun undo_ascii_aux rcs [] = String.implode(rev rcs)
paulson@24183
   169
  | undo_ascii_aux rcs [#"_"] = undo_ascii_aux (#"_"::rcs) []  (*ERROR*)
paulson@24183
   170
      (*Three types of _ escapes: __, _A to _P, _nnn*)
paulson@24183
   171
  | undo_ascii_aux rcs (#"_" :: #"_" :: cs) = undo_ascii_aux (#"_"::rcs) cs
wenzelm@24310
   172
  | undo_ascii_aux rcs (#"_" :: c :: cs) =
paulson@24183
   173
      if #"A" <= c andalso c<= #"P"  (*translation of #" " to #"/"*)
paulson@24183
   174
      then undo_ascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs
wenzelm@24310
   175
      else
paulson@24183
   176
        let val digits = List.take (c::cs, 3) handle Subscript => []
wenzelm@24310
   177
        in
paulson@24183
   178
            case Int.fromString (String.implode digits) of
paulson@24183
   179
                NONE => undo_ascii_aux (c:: #"_"::rcs) cs  (*ERROR*)
paulson@24183
   180
              | SOME n => undo_ascii_aux (Char.chr n :: rcs) (List.drop (cs, 2))
paulson@24183
   181
        end
paulson@24183
   182
  | undo_ascii_aux rcs (c::cs) = undo_ascii_aux (c::rcs) cs;
paulson@24183
   183
paulson@24183
   184
val undo_ascii_of = undo_ascii_aux [] o String.explode;
paulson@15347
   185
paulson@17525
   186
(* convert a list of strings into one single string; surrounded by brackets *)
paulson@18218
   187
fun paren_pack [] = ""   (*empty argument list*)
paulson@18218
   188
  | paren_pack strings = "(" ^ commas strings ^ ")";
paulson@17525
   189
paulson@21509
   190
(*TSTP format uses (...) rather than the old [...]*)
paulson@21509
   191
fun tptp_pack strings = "(" ^ space_implode " | " strings ^ ")";
paulson@17525
   192
paulson@17525
   193
paulson@16925
   194
(*Remove the initial ' character from a type variable, if it is present*)
paulson@16925
   195
fun trim_type_var s =
paulson@16925
   196
  if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
paulson@16925
   197
  else error ("trim_type: Malformed type variable encountered: " ^ s);
paulson@16925
   198
paulson@16903
   199
fun ascii_of_indexname (v,0) = ascii_of v
paulson@17525
   200
  | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
paulson@15347
   201
paulson@17230
   202
fun make_schematic_var v = schematic_var_prefix ^ (ascii_of_indexname v);
paulson@15347
   203
fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x);
paulson@15347
   204
wenzelm@24310
   205
fun make_schematic_type_var (x,i) =
paulson@16925
   206
      tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
paulson@16925
   207
fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
paulson@15347
   208
paulson@23075
   209
(*HACK because SPASS truncates identifiers to 63 characters :-(( *)
paulson@23075
   210
val dfg_format = ref false;
paulson@23075
   211
paulson@23075
   212
(*32-bit hash,so we expect no collisions unless there are around 65536 long identifiers...*)
paulson@23075
   213
fun controlled_length s =
wenzelm@24310
   214
  if size s > 60 andalso !dfg_format
paulson@23075
   215
  then Word.toString (Polyhash.hashw_string(s,0w0))
paulson@23075
   216
  else s;
paulson@23075
   217
paulson@18411
   218
fun lookup_const c =
wenzelm@17412
   219
    case Symtab.lookup const_trans_table c of
paulson@17230
   220
        SOME c' => c'
paulson@23075
   221
      | NONE => controlled_length (ascii_of c);
paulson@17230
   222
wenzelm@24310
   223
fun lookup_type_const c =
wenzelm@17412
   224
    case Symtab.lookup type_const_trans_table c of
paulson@17230
   225
        SOME c' => c'
paulson@23075
   226
      | NONE => controlled_length (ascii_of c);
paulson@18411
   227
paulson@18411
   228
fun make_fixed_const "op =" = "equal"   (*MUST BE "equal" because it's built-in to ATPs*)
paulson@18411
   229
  | make_fixed_const c      = const_prefix ^ lookup_const c;
paulson@18411
   230
paulson@18411
   231
fun make_fixed_type_const c = tconst_prefix ^ lookup_type_const c;
quigley@17150
   232
wenzelm@17261
   233
fun make_type_class clas = class_prefix ^ ascii_of clas;
quigley@17150
   234
quigley@17150
   235
paulson@18798
   236
(***** definitions and functions for FOL clauses, for conversion to TPTP or DFG format. *****)
paulson@15347
   237
paulson@21509
   238
datatype kind = Axiom | Conjecture;
paulson@23385
   239
paulson@15347
   240
fun name_of_kind Axiom = "axiom"
paulson@21509
   241
  | name_of_kind Conjecture = "negated_conjecture";
paulson@15347
   242
paulson@15347
   243
type axiom_name = string;
paulson@15347
   244
paulson@15347
   245
(**** Isabelle FOL clauses ****)
paulson@15347
   246
mengj@18402
   247
datatype typ_var = FOLTVar of indexname | FOLTFree of string;
mengj@18402
   248
paulson@18798
   249
(*FIXME: give the constructors more sensible names*)
mengj@18402
   250
datatype fol_type = AtomV of string
wenzelm@24310
   251
                  | AtomF of string
wenzelm@24310
   252
                  | Comp of string * fol_type list;
mengj@18402
   253
mengj@18402
   254
fun string_of_fol_type (AtomV x) = x
mengj@18402
   255
  | string_of_fol_type (AtomF x) = x
wenzelm@24310
   256
  | string_of_fol_type (Comp(tcon,tps)) =
paulson@18856
   257
      tcon ^ (paren_pack (map string_of_fol_type tps));
wenzelm@24310
   258
paulson@18798
   259
(*First string is the type class; the second is a TVar or TFfree*)
paulson@18798
   260
datatype type_literal = LTVar of string * string | LTFree of string * string;
paulson@15347
   261
mengj@17999
   262
fun mk_typ_var_sort (TFree(a,s)) = (FOLTFree a,s)
mengj@17999
   263
  | mk_typ_var_sort (TVar(v,s)) = (FOLTVar v,s);
mengj@17999
   264
mengj@17999
   265
paulson@17404
   266
exception CLAUSE of string * term;
paulson@15347
   267
paulson@17317
   268
paulson@18218
   269
(*Declarations of the current theory--to allow suppressing types.*)
paulson@18218
   270
val const_typargs = ref (Library.K [] : (string*typ -> typ list));
paulson@17317
   271
paulson@21790
   272
fun num_typargs(s,T) = length (!const_typargs (s,T));
paulson@16925
   273
paulson@16925
   274
(*Initialize the type suppression mechanism with the current theory before
paulson@16925
   275
    producing any clauses!*)
paulson@18218
   276
fun init thy = (const_typargs := Sign.const_typargs thy);
wenzelm@24310
   277
quigley@17150
   278
mengj@18402
   279
(*Flatten a type to a fol_type while accumulating sort constraints on the TFrees and
wenzelm@24310
   280
  TVars it contains.*)
wenzelm@24310
   281
fun type_of (Type (a, Ts)) =
wenzelm@24310
   282
      let val (folTyps, ts) = types_of Ts
wenzelm@24310
   283
          val t = make_fixed_type_const a
paulson@18798
   284
      in (Comp(t,folTyps), ts) end
wenzelm@24310
   285
  | type_of (TFree (a,s)) = (AtomF(make_fixed_type_var a), [(FOLTFree a, s)])
paulson@18798
   286
  | type_of (TVar (v, s)) = (AtomV(make_schematic_type_var v), [(FOLTVar v, s)])
paulson@18218
   287
and types_of Ts =
paulson@18798
   288
      let val (folTyps,ts) = ListPair.unzip (map type_of Ts)
paulson@18798
   289
      in (folTyps, union_all ts) end;
paulson@15390
   290
mengj@18439
   291
paulson@16903
   292
(* Any variables created via the METAHYPS tactical should be treated as
paulson@16903
   293
   universal vars, although it is represented as "Free(...)" by Isabelle *)
paulson@16903
   294
val isMeta = String.isPrefix "METAHYP1_"
quigley@17150
   295
wenzelm@24310
   296
(*Make literals for sorted type variables*)
paulson@22643
   297
fun sorts_on_typs (_, [])   = []
wenzelm@24310
   298
  | sorts_on_typs (v,  s::ss) =
paulson@22643
   299
      let val sorts = sorts_on_typs (v, ss)
paulson@22643
   300
      in
paulson@22643
   301
          if s = "HOL.type" then sorts
paulson@22643
   302
          else case v of
paulson@22643
   303
              FOLTVar indx => LTVar(make_type_class s, make_schematic_type_var indx) :: sorts
paulson@22643
   304
            | FOLTFree x => LTFree(make_type_class s, make_fixed_type_var x) :: sorts
paulson@22643
   305
      end;
quigley@17150
   306
paulson@18798
   307
fun pred_of_sort (LTVar (s,ty)) = (s,1)
paulson@22643
   308
  | pred_of_sort (LTFree (s,ty)) = (s,1)
quigley@17150
   309
paulson@16199
   310
(*Given a list of sorted type variables, return two separate lists.
paulson@16199
   311
  The first is for TVars, the second for TFrees.*)
paulson@18856
   312
fun add_typs_aux [] = ([],[])
wenzelm@24310
   313
  | add_typs_aux ((FOLTVar indx, s) :: tss) =
paulson@17230
   314
      let val vs = sorts_on_typs (FOLTVar indx, s)
wenzelm@24310
   315
          val (vss,fss) = add_typs_aux tss
paulson@23385
   316
      in  (vs union vss, fss)  end
paulson@22643
   317
  | add_typs_aux ((FOLTFree x, s) :: tss) =
paulson@17230
   318
      let val fs = sorts_on_typs (FOLTFree x, s)
wenzelm@24310
   319
          val (vss,fss) = add_typs_aux tss
paulson@23385
   320
      in  (vss, fs union fss)  end;
quigley@17150
   321
mengj@20015
   322
mengj@20015
   323
(** make axiom and conjecture clauses. **)
mengj@20015
   324
mengj@20015
   325
fun get_tvar_strs [] = []
wenzelm@24310
   326
  | get_tvar_strs ((FOLTVar indx,s)::tss) =
haftmann@20854
   327
      insert (op =) (make_schematic_type_var indx) (get_tvar_strs tss)
mengj@20015
   328
  | get_tvar_strs((FOLTFree x,s)::tss) = get_tvar_strs tss
mengj@20015
   329
wenzelm@24310
   330
mengj@19354
   331
paulson@15347
   332
(**** Isabelle arities ****)
paulson@15347
   333
paulson@15347
   334
exception ARCLAUSE of string;
paulson@15347
   335
wenzelm@24310
   336
datatype arLit = TConsLit of class * string * string list
paulson@22643
   337
               | TVarLit of class * string;
wenzelm@24310
   338
wenzelm@24310
   339
datatype arityClause =
wenzelm@24310
   340
         ArityClause of {axiom_name: axiom_name,
wenzelm@24310
   341
                         kind: kind,
wenzelm@24310
   342
                         conclLit: arLit,
wenzelm@24310
   343
                         premLits: arLit list};
paulson@15347
   344
paulson@15347
   345
paulson@18798
   346
fun gen_TVars 0 = []
paulson@18798
   347
  | gen_TVars n = ("T_" ^ Int.toString n) :: gen_TVars (n-1);
paulson@15347
   348
paulson@18411
   349
fun pack_sort(_,[])  = []
paulson@18411
   350
  | pack_sort(tvar, "HOL.type"::srt) = pack_sort(tvar, srt)   (*IGNORE sort "type"*)
paulson@22643
   351
  | pack_sort(tvar, cls::srt) =  (cls, tvar) :: pack_sort(tvar, srt);
wenzelm@24310
   352
paulson@18411
   353
(*Arity of type constructor tcon :: (arg1,...,argN)res*)
paulson@22643
   354
fun make_axiom_arity_clause (tcons, axiom_name, (cls,args)) =
paulson@21560
   355
   let val tvars = gen_TVars (length args)
paulson@17845
   356
       val tvars_srts = ListPair.zip (tvars,args)
paulson@17845
   357
   in
wenzelm@24310
   358
      ArityClause {axiom_name = axiom_name, kind = Axiom,
wenzelm@24310
   359
                   conclLit = TConsLit (cls, make_fixed_type_const tcons, tvars),
paulson@22643
   360
                   premLits = map TVarLit (union_all(map pack_sort tvars_srts))}
paulson@17845
   361
   end;
paulson@15347
   362
paulson@15347
   363
paulson@15347
   364
(**** Isabelle class relations ****)
paulson@15347
   365
wenzelm@24310
   366
datatype classrelClause =
wenzelm@24310
   367
         ClassrelClause of {axiom_name: axiom_name,
wenzelm@24310
   368
                            subclass: class,
wenzelm@24310
   369
                            superclass: class};
wenzelm@24310
   370
paulson@21290
   371
(*Generate all pairs (sub,super) such that sub is a proper subclass of super in theory thy.*)
paulson@21432
   372
fun class_pairs thy [] supers = []
paulson@21432
   373
  | class_pairs thy subs supers =
paulson@21432
   374
      let val class_less = Sorts.class_less(Sign.classes_of thy)
wenzelm@24310
   375
          fun add_super sub (super,pairs) =
wenzelm@24310
   376
                if class_less (sub,super) then (sub,super)::pairs else pairs
wenzelm@24310
   377
          fun add_supers (sub,pairs) = foldl (add_super sub) pairs supers
paulson@21432
   378
      in  foldl add_supers [] subs  end;
paulson@15347
   379
paulson@21290
   380
fun make_classrelClause (sub,super) =
paulson@21290
   381
  ClassrelClause {axiom_name = clrelclause_prefix ^ ascii_of sub ^ "_" ^ ascii_of super,
wenzelm@24310
   382
                  subclass = make_type_class sub,
paulson@21290
   383
                  superclass = make_type_class super};
paulson@15347
   384
paulson@21290
   385
fun make_classrel_clauses thy subs supers =
paulson@21290
   386
  map make_classrelClause (class_pairs thy subs supers);
paulson@18868
   387
paulson@18868
   388
paulson@18868
   389
(** Isabelle arities **)
paulson@17845
   390
paulson@21373
   391
fun arity_clause _ _ (tcons, []) = []
paulson@21373
   392
  | arity_clause seen n (tcons, ("HOL.type",_)::ars) =  (*ignore*)
paulson@21373
   393
      arity_clause seen n (tcons,ars)
paulson@21373
   394
  | arity_clause seen n (tcons, (ar as (class,_)) :: ars) =
paulson@21373
   395
      if class mem_string seen then (*multiple arities for the same tycon, class pair*)
wenzelm@24310
   396
          make_axiom_arity_clause (tcons, lookup_type_const tcons ^ "_" ^ class ^ "_" ^ Int.toString n, ar) ::
wenzelm@24310
   397
          arity_clause seen (n+1) (tcons,ars)
paulson@21373
   398
      else
wenzelm@24310
   399
          make_axiom_arity_clause (tcons, lookup_type_const tcons ^ "_" ^ class, ar) ::
wenzelm@24310
   400
          arity_clause (class::seen) n (tcons,ars)
paulson@17845
   401
paulson@17845
   402
fun multi_arity_clause [] = []
paulson@19155
   403
  | multi_arity_clause ((tcons,ars) :: tc_arlists) =
wenzelm@24310
   404
      arity_clause [] 1 (tcons, ars)  @  multi_arity_clause tc_arlists
paulson@17845
   405
paulson@22643
   406
(*Generate all pairs (tycon,class,sorts) such that tycon belongs to class in theory thy
paulson@22643
   407
  provided its arguments have the corresponding sorts.*)
paulson@21373
   408
fun type_class_pairs thy tycons classes =
paulson@21373
   409
  let val alg = Sign.classes_of thy
paulson@21373
   410
      fun domain_sorts (tycon,class) = Sorts.mg_domain alg tycon [class]
wenzelm@24310
   411
      fun add_class tycon (class,pairs) =
wenzelm@24310
   412
            (class, domain_sorts(tycon,class))::pairs
paulson@21373
   413
            handle Sorts.CLASS_ERROR _ => pairs
paulson@21373
   414
      fun try_classes tycon = (tycon, foldl (add_class tycon) [] classes)
paulson@21373
   415
  in  map try_classes tycons  end;
paulson@21373
   416
paulson@22643
   417
(*Proving one (tycon, class) membership may require proving others, so iterate.*)
paulson@22643
   418
fun iter_type_class_pairs thy tycons [] = ([], [])
paulson@22643
   419
  | iter_type_class_pairs thy tycons classes =
paulson@22643
   420
      let val cpairs = type_class_pairs thy tycons classes
paulson@22643
   421
          val newclasses = union_all (union_all (union_all (map (map #2 o #2) cpairs))) \\ classes \\ HOLogic.typeS
wenzelm@24310
   422
          val _ = if null newclasses then ()
paulson@22643
   423
                  else Output.debug (fn _ => "New classes: " ^ space_implode ", " newclasses)
wenzelm@24310
   424
          val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
paulson@22643
   425
      in  (classes' union classes, cpairs' union cpairs)  end;
wenzelm@24310
   426
paulson@22643
   427
fun make_arity_clauses thy tycons classes =
wenzelm@24310
   428
  let val (classes', cpairs) = iter_type_class_pairs thy tycons classes
paulson@22643
   429
  in  (classes', multi_arity_clause cpairs)  end;
paulson@17845
   430
paulson@17845
   431
paulson@18868
   432
(**** Find occurrences of predicates in clauses ****)
paulson@17845
   433
wenzelm@24310
   434
(*FIXME: multiple-arity checking doesn't work, as update_new is the wrong
paulson@18868
   435
  function (it flags repeated declarations of a function, even with the same arity)*)
paulson@17845
   436
paulson@18868
   437
fun update_many (tab, keypairs) = foldl (uncurry Symtab.update) tab keypairs;
paulson@18798
   438
wenzelm@24310
   439
fun add_type_sort_preds ((FOLTVar indx,s), preds) =
paulson@18868
   440
      update_many (preds, map pred_of_sort (sorts_on_typs (FOLTVar indx, s)))
paulson@18868
   441
  | add_type_sort_preds ((FOLTFree x,s), preds) =
paulson@18868
   442
      update_many (preds, map pred_of_sort (sorts_on_typs (FOLTFree x, s)));
paulson@18868
   443
paulson@18868
   444
fun add_classrelClause_preds (ClassrelClause {subclass,superclass,...}, preds) =
paulson@18868
   445
  Symtab.update (subclass,1) (Symtab.update (superclass,1) preds);
paulson@18868
   446
paulson@22643
   447
fun class_of_arityLit (TConsLit (tclass, _, _)) = tclass
paulson@22643
   448
  | class_of_arityLit (TVarLit (tclass, _)) = tclass;
paulson@21373
   449
paulson@21373
   450
fun add_arityClause_preds (ArityClause {conclLit,premLits,...}, preds) =
paulson@22643
   451
  let val classes = map (make_type_class o class_of_arityLit) (conclLit::premLits)
paulson@21373
   452
      fun upd (class,preds) = Symtab.update (class,1) preds
paulson@21373
   453
  in  foldl upd preds classes  end;
paulson@18868
   454
paulson@18868
   455
(*** Find occurrences of functions in clauses ***)
paulson@18868
   456
paulson@18868
   457
fun add_foltype_funcs (AtomV _, funcs) = funcs
paulson@18868
   458
  | add_foltype_funcs (AtomF a, funcs) = Symtab.update (a,0) funcs
wenzelm@24310
   459
  | add_foltype_funcs (Comp(a,tys), funcs) =
paulson@18868
   460
      foldl add_foltype_funcs (Symtab.update (a, length tys) funcs) tys;
paulson@18868
   461
paulson@20038
   462
(*TFrees are recorded as constants*)
paulson@20038
   463
fun add_type_sort_funcs ((FOLTVar _, _), funcs) = funcs
wenzelm@24310
   464
  | add_type_sort_funcs ((FOLTFree a, _), funcs) =
paulson@20038
   465
      Symtab.update (make_fixed_type_var a, 0) funcs
paulson@20038
   466
paulson@18868
   467
fun add_arityClause_funcs (ArityClause {conclLit,...}, funcs) =
paulson@22643
   468
  let val TConsLit (_, tcons, tvars) = conclLit
paulson@18868
   469
  in  Symtab.update (tcons, length tvars) funcs  end;
paulson@18868
   470
paulson@23075
   471
(*This type can be overlooked because it is built-in...*)
paulson@23075
   472
val init_functab = Symtab.update ("tc_itself", 1) Symtab.empty;
paulson@23075
   473
paulson@18868
   474
paulson@18868
   475
(**** String-oriented operations ****)
paulson@15347
   476
wenzelm@24310
   477
fun string_of_clausename (cls_id,ax_name) =
paulson@17525
   478
    clause_prefix ^ ascii_of ax_name ^ "_" ^ Int.toString cls_id;
paulson@17317
   479
wenzelm@24310
   480
fun string_of_type_clsname (cls_id,ax_name,idx) =
paulson@17525
   481
    string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx);
paulson@18863
   482
wenzelm@24310
   483
fun writeln_strs os = List.app (fn s => TextIO.output (os, s));
paulson@18863
   484
wenzelm@24310
   485
paulson@18868
   486
(**** Producing DFG files ****)
quigley@17150
   487
paulson@18863
   488
(*Attach sign in DFG syntax: false means negate.*)
paulson@18863
   489
fun dfg_sign true s = s
wenzelm@24310
   490
  | dfg_sign false s = "not(" ^ s ^ ")"
paulson@18863
   491
paulson@18798
   492
fun dfg_of_typeLit (LTVar (s,ty)) = "not(" ^ s ^ "(" ^ ty ^ "))"
paulson@18856
   493
  | dfg_of_typeLit (LTFree (s,ty)) = s ^ "(" ^ ty ^ ")";
wenzelm@24310
   494
paulson@18868
   495
(*Enclose the clause body by quantifiers, if necessary*)
wenzelm@24310
   496
fun dfg_forall [] body = body
paulson@18868
   497
  | dfg_forall vars body = "forall([" ^ commas vars ^ "],\n" ^ body ^ ")"
quigley@17150
   498
wenzelm@24310
   499
fun gen_dfg_cls (cls_id, ax_name, knd, lits, vars) =
wenzelm@24310
   500
    "clause( %(" ^ knd ^ ")\n" ^
wenzelm@24310
   501
    dfg_forall vars ("or(" ^ lits ^ ")") ^ ",\n" ^
paulson@18863
   502
    string_of_clausename (cls_id,ax_name) ^  ").\n\n";
quigley@17150
   503
paulson@18798
   504
fun string_of_arity (name, num) =  "(" ^ name ^ "," ^ Int.toString num ^ ")"
quigley@17150
   505
paulson@18856
   506
fun string_of_preds [] = ""
paulson@18856
   507
  | string_of_preds preds = "predicates[" ^ commas(map string_of_arity preds) ^ "].\n";
quigley@17150
   508
paulson@18856
   509
fun string_of_funcs [] = ""
paulson@18856
   510
  | string_of_funcs funcs = "functions[" ^ commas(map string_of_arity funcs) ^ "].\n" ;
quigley@17150
   511
wenzelm@24310
   512
fun string_of_symbols predstr funcstr =
paulson@17234
   513
  "list_of_symbols.\n" ^ predstr  ^ funcstr  ^ "end_of_list.\n\n";
quigley@17150
   514
paulson@18798
   515
fun string_of_start name = "begin_problem(" ^ name ^ ").\n\n";
quigley@17150
   516
wenzelm@24310
   517
fun string_of_descrip name =
wenzelm@24310
   518
  "list_of_descriptions.\nname({*" ^ name ^
paulson@18868
   519
  "*}).\nauthor({*Isabelle*}).\nstatus(unknown).\ndescription({*auto-generated*}).\nend_of_list.\n\n"
quigley@17150
   520
paulson@18863
   521
fun dfg_tfree_clause tfree_lit =
paulson@21509
   522
  "clause( %(negated_conjecture)\n" ^ "or( " ^ tfree_lit ^ "),\n" ^ "tfree_tcs" ^ ").\n\n"
quigley@17150
   523
paulson@22643
   524
fun dfg_of_arLit (TConsLit (c,t,args)) =
paulson@22643
   525
      dfg_sign true (make_type_class c ^ "(" ^ t ^ paren_pack args ^ ")")
paulson@22643
   526
  | dfg_of_arLit (TVarLit (c,str)) =
paulson@22643
   527
      dfg_sign false (make_type_class c ^ "(" ^ str ^ ")")
wenzelm@24310
   528
paulson@20038
   529
fun dfg_classrelLits sub sup =  "not(" ^ sub ^ "(T)), " ^ sup ^ "(T)";
paulson@17525
   530
paulson@18868
   531
fun dfg_classrelClause (ClassrelClause {axiom_name,subclass,superclass,...}) =
paulson@18868
   532
  "clause(forall([T],\nor( " ^ dfg_classrelLits subclass superclass ^ ")),\n" ^
paulson@18868
   533
  axiom_name ^ ").\n\n";
wenzelm@24310
   534
paulson@21560
   535
fun string_of_ar axiom_name = arclause_prefix ^ ascii_of axiom_name;
paulson@21560
   536
wenzelm@24310
   537
fun dfg_arity_clause (ArityClause{axiom_name,kind,conclLit,premLits,...}) =
paulson@22643
   538
  let val TConsLit (_,_,tvars) = conclLit
paulson@18868
   539
      val lits = map dfg_of_arLit (conclLit :: premLits)
paulson@18863
   540
  in
wenzelm@24310
   541
      "clause( %(" ^ name_of_kind kind ^ ")\n" ^
paulson@18868
   542
      dfg_forall tvars ("or( " ^ commas lits ^ ")") ^ ",\n" ^
paulson@21560
   543
      string_of_ar axiom_name ^ ").\n\n"
paulson@18863
   544
  end;
paulson@18863
   545
quigley@17150
   546
paulson@18869
   547
(**** Produce TPTP files ****)
paulson@18868
   548
paulson@18868
   549
(*Attach sign in TPTP syntax: false means negate.*)
paulson@21509
   550
fun tptp_sign true s = s
paulson@21509
   551
  | tptp_sign false s = "~ " ^ s
paulson@18868
   552
paulson@21509
   553
fun tptp_of_typeLit (LTVar (s,ty))  = tptp_sign false (s ^ "(" ^ ty ^ ")")
paulson@21509
   554
  | tptp_of_typeLit (LTFree (s,ty)) = tptp_sign true  (s ^ "(" ^ ty ^ ")");
wenzelm@24310
   555
wenzelm@24310
   556
fun gen_tptp_cls (cls_id,ax_name,knd,lits) =
wenzelm@24310
   557
    "cnf(" ^ string_of_clausename (cls_id,ax_name) ^ "," ^
paulson@21509
   558
    name_of_kind knd ^ "," ^ tptp_pack lits ^ ").\n";
paulson@15347
   559
paulson@18863
   560
fun tptp_tfree_clause tfree_lit =
paulson@21509
   561
    "cnf(" ^ "tfree_tcs," ^ "negated_conjecture" ^ "," ^ tptp_pack[tfree_lit] ^ ").\n";
wenzelm@24310
   562
paulson@22643
   563
fun tptp_of_arLit (TConsLit (c,t,args)) =
paulson@22643
   564
      tptp_sign true (make_type_class c ^ "(" ^ t ^ paren_pack args ^ ")")
paulson@22643
   565
  | tptp_of_arLit (TVarLit (c,str)) =
paulson@22643
   566
      tptp_sign false (make_type_class c ^ "(" ^ str ^ ")")
wenzelm@24310
   567
wenzelm@24310
   568
fun tptp_arity_clause (ArityClause{axiom_name,kind,conclLit,premLits,...}) =
wenzelm@24310
   569
  "cnf(" ^ string_of_ar axiom_name ^ "," ^ name_of_kind kind ^ "," ^
paulson@21560
   570
  tptp_pack (map tptp_of_arLit (conclLit :: premLits)) ^ ").\n";
paulson@15347
   571
wenzelm@24310
   572
fun tptp_classrelLits sub sup =
paulson@21509
   573
  let val tvar = "(T)"
paulson@21509
   574
  in  tptp_pack [tptp_sign false (sub^tvar), tptp_sign true (sup^tvar)]  end;
paulson@15347
   575
paulson@18868
   576
fun tptp_classrelClause (ClassrelClause {axiom_name,subclass,superclass,...}) =
wenzelm@24310
   577
  "cnf(" ^ axiom_name ^ ",axiom," ^ tptp_classrelLits subclass superclass ^ ").\n"
quigley@17150
   578
paulson@15347
   579
end;