src/HOL/Tools/Sledgehammer/sledgehammer_fol_clause.ML
author blanchet
Wed, 17 Mar 2010 19:26:05 +0100
changeset 35826 1590abc3d42a
parent 35825 a6aad5a70ed4
child 35827 f552152d7747
permissions -rw-r--r--
renamed Sledgehammer structures
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_fol_clause.ML
33311
49cd8abb2863 proper header;
wenzelm
parents: 33043
diff changeset
     2
    Author:     Jia Meng, Cambridge University Computer Laboratory
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     3
33311
49cd8abb2863 proper header;
wenzelm
parents: 33043
diff changeset
     4
Storing/printing FOL clauses and arity clauses.  Typed equality is
49cd8abb2863 proper header;
wenzelm
parents: 33043
diff changeset
     5
treated differently.
49cd8abb2863 proper header;
wenzelm
parents: 33043
diff changeset
     6
49cd8abb2863 proper header;
wenzelm
parents: 33043
diff changeset
     7
FIXME: combine with res_hol_clause!
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     8
*)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     9
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    10
signature SLEDGEHAMMER_FOL_CLAUSE =
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    11
sig
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    12
  val schematic_var_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    13
  val fixed_var_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    14
  val tvar_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    15
  val tfree_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    16
  val clause_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    17
  val const_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    18
  val tconst_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    19
  val class_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    20
  val union_all: ''a list list -> ''a list
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    21
  val const_trans_table: string Symtab.table
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    22
  val type_const_trans_table: string Symtab.table
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    23
  val ascii_of: string -> string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    24
  val undo_ascii_of: string -> string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    25
  val paren_pack : string list -> string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    26
  val make_schematic_var : string * int -> string
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
    27
  val make_fixed_var : string -> string
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
    28
  val make_schematic_type_var : string * int -> string
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    29
  val make_fixed_type_var : string -> string
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
    30
  val make_fixed_const : bool -> string -> string
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
    31
  val make_fixed_type_const : bool -> string -> string
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
    32
  val make_type_class : string -> string
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    33
  datatype kind = Axiom | Conjecture
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    34
  type axiom_name = string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    35
  datatype fol_type =
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    36
      AtomV of string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    37
    | AtomF of string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    38
    | Comp of string * fol_type list
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
    39
  val string_of_fol_type : fol_type -> string
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    40
  datatype type_literal = LTVar of string * string | LTFree of string * string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    41
  exception CLAUSE of string * term
24940
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
    42
  val add_typs : typ list -> type_literal list
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
    43
  val get_tvar_strs: typ list -> string list
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    44
  datatype arLit =
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    45
      TConsLit of class * string * string list
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    46
    | TVarLit of class * string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    47
  datatype arityClause = ArityClause of
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
    48
   {axiom_name: axiom_name, conclLit: arLit, premLits: arLit list}
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    49
  datatype classrelClause = ClassrelClause of
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
    50
   {axiom_name: axiom_name, subclass: class, superclass: class}
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    51
  val make_classrel_clauses: theory -> class list -> class list -> classrelClause list
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
    52
  val make_arity_clauses_dfg: bool -> theory -> string list -> class list -> class list * arityClause list
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    53
  val make_arity_clauses: theory -> string list -> class list -> class list * arityClause list
24940
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
    54
  val add_type_sort_preds: typ * int Symtab.table -> int Symtab.table
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    55
  val add_classrelClause_preds : classrelClause * int Symtab.table -> int Symtab.table
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    56
  val class_of_arityLit: arLit -> class
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    57
  val add_arityClause_preds: arityClause * int Symtab.table -> int Symtab.table
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    58
  val add_foltype_funcs: fol_type * int Symtab.table -> int Symtab.table
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    59
  val add_arityClause_funcs: arityClause * int Symtab.table -> int Symtab.table
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    60
  val init_functab: int Symtab.table
19719
837025cc6317 Changed input interface of dfg_write_file.
mengj
parents: 19642
diff changeset
    61
  val dfg_sign: bool -> string -> string
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
    62
  val dfg_of_typeLit: bool -> type_literal -> string
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
    63
  val gen_dfg_cls: int * string * kind * string list * string list * string list -> string
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    64
  val string_of_preds: (string * Int.int) list -> string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    65
  val string_of_funcs: (string * int) list -> string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    66
  val string_of_symbols: string -> string -> string
19719
837025cc6317 Changed input interface of dfg_write_file.
mengj
parents: 19642
diff changeset
    67
  val string_of_start: string -> string
837025cc6317 Changed input interface of dfg_write_file.
mengj
parents: 19642
diff changeset
    68
  val string_of_descrip : string -> string
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    69
  val dfg_tfree_clause : string -> string
19719
837025cc6317 Changed input interface of dfg_write_file.
mengj
parents: 19642
diff changeset
    70
  val dfg_classrelClause: classrelClause -> string
837025cc6317 Changed input interface of dfg_write_file.
mengj
parents: 19642
diff changeset
    71
  val dfg_arity_clause: arityClause -> string
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    72
  val tptp_sign: bool -> string -> string
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
    73
  val tptp_of_typeLit : bool -> type_literal -> string
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
    74
  val gen_tptp_cls : int * string * kind * string list * string list -> string
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    75
  val tptp_tfree_clause : string -> string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    76
  val tptp_arity_clause : arityClause -> string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    77
  val tptp_classrelClause : classrelClause -> string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    78
end
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    79
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    80
structure Sledgehammer_FOL_Clause : SLEDGEHAMMER_FOL_CLAUSE =
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    81
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    82
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    83
val schematic_var_prefix = "V_";
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    84
val fixed_var_prefix = "v_";
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    85
17230
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
    86
val tvar_prefix = "T_";
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
    87
val tfree_prefix = "t_";
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    88
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    89
val clause_prefix = "cls_";
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    90
val arclause_prefix = "clsarity_"
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
    91
val clrelclause_prefix = "clsrel_";
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    92
17230
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
    93
val const_prefix = "c_";
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    94
val tconst_prefix = "tc_";
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    95
val class_prefix = "class_";
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    96
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33038
diff changeset
    97
fun union_all xss = List.foldl (uncurry (union (op =))) [] xss;
17775
2679ba74411f minor code tidyig
paulson
parents: 17764
diff changeset
    98
2679ba74411f minor code tidyig
paulson
parents: 17764
diff changeset
    99
(*Provide readable names for the more common symbolic functions*)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   100
val const_trans_table =
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29676
diff changeset
   101
      Symtab.make [(@{const_name "op ="}, "equal"),
35092
cfe605c54e50 moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents: 34974
diff changeset
   102
                   (@{const_name Orderings.less_eq}, "lessequals"),
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29676
diff changeset
   103
                   (@{const_name "op &"}, "and"),
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29676
diff changeset
   104
                   (@{const_name "op |"}, "or"),
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29676
diff changeset
   105
                   (@{const_name "op -->"}, "implies"),
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29676
diff changeset
   106
                   (@{const_name "op :"}, "in"),
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   107
                   ("ATP_Linkup.fequal", "fequal"),
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   108
                   ("ATP_Linkup.COMBI", "COMBI"),
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   109
                   ("ATP_Linkup.COMBK", "COMBK"),
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   110
                   ("ATP_Linkup.COMBB", "COMBB"),
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   111
                   ("ATP_Linkup.COMBC", "COMBC"),
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   112
                   ("ATP_Linkup.COMBS", "COMBS"),
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   113
                   ("ATP_Linkup.COMBB'", "COMBB_e"),
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   114
                   ("ATP_Linkup.COMBC'", "COMBC_e"),
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   115
                   ("ATP_Linkup.COMBS'", "COMBS_e")];
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   116
17230
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
   117
val type_const_trans_table =
18411
2d3165a0fb40 No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents: 18409
diff changeset
   118
      Symtab.make [("*", "prod"),
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   119
                   ("+", "sum"),
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   120
                   ("~=>", "map")];
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   121
15610
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   122
(*Escaping of special characters.
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   123
  Alphanumeric characters are left unchanged.
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   124
  The character _ goes to __
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   125
  Characters in the range ASCII space to / go to _A to _P, respectively.
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   126
  Other printing characters go to _nnn where nnn is the decimal ASCII code.*)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   127
val A_minus_space = Char.ord #"A" - Char.ord #" ";
15610
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   128
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   129
fun stringN_of_int 0 _ = ""
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   130
  | stringN_of_int k n = stringN_of_int (k-1) (n div 10) ^ Int.toString (n mod 10);
15610
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   131
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   132
fun ascii_of_c c =
15610
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   133
  if Char.isAlphaNum c then String.str c
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   134
  else if c = #"_" then "__"
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   135
  else if #" " <= c andalso c <= #"/"
15610
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   136
       then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   137
  else if Char.isPrint c
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   138
       then ("_" ^ stringN_of_int 3 (Char.ord c))  (*fixed width, in case more digits follow*)
15610
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   139
  else ""
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   140
15610
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   141
val ascii_of = String.translate ascii_of_c;
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   142
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   143
(** Remove ASCII armouring from names in proof files **)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   144
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   145
(*We don't raise error exceptions because this code can run inside the watcher.
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   146
  Also, the errors are "impossible" (hah!)*)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   147
fun undo_ascii_aux rcs [] = String.implode(rev rcs)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   148
  | undo_ascii_aux rcs [#"_"] = undo_ascii_aux (#"_"::rcs) []  (*ERROR*)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   149
      (*Three types of _ escapes: __, _A to _P, _nnn*)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   150
  | undo_ascii_aux rcs (#"_" :: #"_" :: cs) = undo_ascii_aux (#"_"::rcs) cs
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   151
  | undo_ascii_aux rcs (#"_" :: c :: cs) =
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   152
      if #"A" <= c andalso c<= #"P"  (*translation of #" " to #"/"*)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   153
      then undo_ascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   154
      else
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   155
        let val digits = List.take (c::cs, 3) handle Subscript => []
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   156
        in
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   157
            case Int.fromString (String.implode digits) of
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   158
                NONE => undo_ascii_aux (c:: #"_"::rcs) cs  (*ERROR*)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   159
              | SOME n => undo_ascii_aux (Char.chr n :: rcs) (List.drop (cs, 2))
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   160
        end
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   161
  | undo_ascii_aux rcs (c::cs) = undo_ascii_aux (c::rcs) cs;
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   162
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   163
val undo_ascii_of = undo_ascii_aux [] o String.explode;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   164
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
   165
(* convert a list of strings into one single string; surrounded by brackets *)
18218
9a7ffce389c3 new treatment of polymorphic types, using Sign.const_typargs
paulson
parents: 18199
diff changeset
   166
fun paren_pack [] = ""   (*empty argument list*)
9a7ffce389c3 new treatment of polymorphic types, using Sign.const_typargs
paulson
parents: 18199
diff changeset
   167
  | paren_pack strings = "(" ^ commas strings ^ ")";
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
   168
21509
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   169
(*TSTP format uses (...) rather than the old [...]*)
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   170
fun tptp_pack strings = "(" ^ space_implode " | " strings ^ ")";
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
   171
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
   172
16925
0fd7b1438d28 simpler variable names, and no types for monomorphic constants
paulson
parents: 16903
diff changeset
   173
(*Remove the initial ' character from a type variable, if it is present*)
0fd7b1438d28 simpler variable names, and no types for monomorphic constants
paulson
parents: 16903
diff changeset
   174
fun trim_type_var s =
0fd7b1438d28 simpler variable names, and no types for monomorphic constants
paulson
parents: 16903
diff changeset
   175
  if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
0fd7b1438d28 simpler variable names, and no types for monomorphic constants
paulson
parents: 16903
diff changeset
   176
  else error ("trim_type: Malformed type variable encountered: " ^ s);
0fd7b1438d28 simpler variable names, and no types for monomorphic constants
paulson
parents: 16903
diff changeset
   177
16903
bf42a9071ad1 streamlined the tptp output
paulson
parents: 16794
diff changeset
   178
fun ascii_of_indexname (v,0) = ascii_of v
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
   179
  | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   180
17230
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
   181
fun make_schematic_var v = schematic_var_prefix ^ (ascii_of_indexname v);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   182
fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x);
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   183
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   184
fun make_schematic_type_var (x,i) =
16925
0fd7b1438d28 simpler variable names, and no types for monomorphic constants
paulson
parents: 16903
diff changeset
   185
      tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
0fd7b1438d28 simpler variable names, and no types for monomorphic constants
paulson
parents: 16903
diff changeset
   186
fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   187
23075
69e30a7e8880 Some hacks for SPASS format
paulson
parents: 23029
diff changeset
   188
(*HACK because SPASS truncates identifiers to 63 characters :-(( *)
69e30a7e8880 Some hacks for SPASS format
paulson
parents: 23029
diff changeset
   189
(*32-bit hash,so we expect no collisions unless there are around 65536 long identifiers...*)
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   190
fun controlled_length dfg_format s =
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   191
  if size s > 60 andalso dfg_format
23075
69e30a7e8880 Some hacks for SPASS format
paulson
parents: 23029
diff changeset
   192
  then Word.toString (Polyhash.hashw_string(s,0w0))
69e30a7e8880 Some hacks for SPASS format
paulson
parents: 23029
diff changeset
   193
  else s;
69e30a7e8880 Some hacks for SPASS format
paulson
parents: 23029
diff changeset
   194
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   195
fun lookup_const dfg c =
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17404
diff changeset
   196
    case Symtab.lookup const_trans_table c of
17230
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
   197
        SOME c' => c'
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   198
      | NONE => controlled_length dfg (ascii_of c);
17230
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
   199
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   200
fun lookup_type_const dfg c =
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17404
diff changeset
   201
    case Symtab.lookup type_const_trans_table c of
17230
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
   202
        SOME c' => c'
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   203
      | NONE => controlled_length dfg (ascii_of c);
18411
2d3165a0fb40 No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents: 18409
diff changeset
   204
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   205
fun make_fixed_const _ "op =" = "equal"   (*MUST BE "equal" because it's built-in to ATPs*)
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   206
  | make_fixed_const dfg c      = const_prefix ^ lookup_const dfg c;
18411
2d3165a0fb40 No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents: 18409
diff changeset
   207
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   208
fun make_fixed_type_const dfg c = tconst_prefix ^ lookup_type_const dfg c;
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   209
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17234
diff changeset
   210
fun make_type_class clas = class_prefix ^ ascii_of clas;
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   211
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   212
18798
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   213
(***** definitions and functions for FOL clauses, for conversion to TPTP or DFG format. *****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   214
21509
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   215
datatype kind = Axiom | Conjecture;
23385
0ef4f9fc0d09 Deleted unused code
paulson
parents: 23075
diff changeset
   216
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   217
type axiom_name = string;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   218
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   219
(**** Isabelle FOL clauses ****)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   220
18798
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   221
(*FIXME: give the constructors more sensible names*)
18402
aaba095cf62b 1. changed fol_type, it's not a string type anymore.
mengj
parents: 18390
diff changeset
   222
datatype fol_type = AtomV of string
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   223
                  | AtomF of string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   224
                  | Comp of string * fol_type list;
18402
aaba095cf62b 1. changed fol_type, it's not a string type anymore.
mengj
parents: 18390
diff changeset
   225
aaba095cf62b 1. changed fol_type, it's not a string type anymore.
mengj
parents: 18390
diff changeset
   226
fun string_of_fol_type (AtomV x) = x
aaba095cf62b 1. changed fol_type, it's not a string type anymore.
mengj
parents: 18390
diff changeset
   227
  | string_of_fol_type (AtomF x) = x
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   228
  | string_of_fol_type (Comp(tcon,tps)) =
18856
4669dec681f4 tidy-up of res_clause.ML, removing the "predicates" field
paulson
parents: 18798
diff changeset
   229
      tcon ^ (paren_pack (map string_of_fol_type tps));
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   230
18798
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   231
(*First string is the type class; the second is a TVar or TFfree*)
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   232
datatype type_literal = LTVar of string * string | LTFree of string * string;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   233
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17375
diff changeset
   234
exception CLAUSE of string * term;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   235
24940
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   236
fun atomic_type (TFree (a,_)) = AtomF(make_fixed_type_var a)
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   237
  | atomic_type (TVar (v,_))  = AtomV(make_schematic_type_var v);
17317
3f12de2e2e6e Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
paulson
parents: 17312
diff changeset
   238
18402
aaba095cf62b 1. changed fol_type, it's not a string type anymore.
mengj
parents: 18390
diff changeset
   239
(*Flatten a type to a fol_type while accumulating sort constraints on the TFrees and
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   240
  TVars it contains.*)
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   241
fun type_of dfg (Type (a, Ts)) =
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   242
      let val (folTyps, ts) = types_of dfg Ts
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   243
          val t = make_fixed_type_const dfg a
18798
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   244
      in (Comp(t,folTyps), ts) end
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   245
  | type_of dfg T = (atomic_type T, [T])
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   246
and types_of dfg Ts =
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   247
      let val (folTyps,ts) = ListPair.unzip (map (type_of dfg) Ts)
18798
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   248
      in (folTyps, union_all ts) end;
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15347
diff changeset
   249
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   250
(*Make literals for sorted type variables*)
24940
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   251
fun sorts_on_typs_aux (_, [])   = []
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   252
  | sorts_on_typs_aux ((x,i),  s::ss) =
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   253
      let val sorts = sorts_on_typs_aux ((x,i), ss)
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   254
      in
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   255
          if s = "HOL.type" then sorts
24940
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   256
          else if i = ~1 then LTFree(make_type_class s, make_fixed_type_var x) :: sorts
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   257
          else LTVar(make_type_class s, make_schematic_type_var (x,i)) :: sorts
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   258
      end;
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   259
24940
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   260
fun sorts_on_typs (TFree (a,s)) = sorts_on_typs_aux ((a,~1),s)
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   261
  | sorts_on_typs (TVar (v,s))  = sorts_on_typs_aux (v,s);
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   262
18798
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   263
fun pred_of_sort (LTVar (s,ty)) = (s,1)
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   264
  | pred_of_sort (LTFree (s,ty)) = (s,1)
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   265
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   266
(*Given a list of sorted type variables, return a list of type literals.*)
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33038
diff changeset
   267
fun add_typs Ts = List.foldl (uncurry (union (op =))) [] (map sorts_on_typs Ts);
20015
1ffcf4802802 Literals aren't sorted any more. Output overloaded constants' type var instantiations.
mengj
parents: 19719
diff changeset
   268
29676
cfa3378decf7 Updated comments.
paulson
parents: 25243
diff changeset
   269
(*The correct treatment of TFrees like 'a in lemmas (axiom clauses) is not clear.
cfa3378decf7 Updated comments.
paulson
parents: 25243
diff changeset
   270
  * Ignoring them leads to unsound proofs, since we do nothing to ensure that 'a
cfa3378decf7 Updated comments.
paulson
parents: 25243
diff changeset
   271
    in a lemma has the same sort as 'a in the conjecture.
cfa3378decf7 Updated comments.
paulson
parents: 25243
diff changeset
   272
  * Deleting such clauses will lead to problems with locales in other use of local results
cfa3378decf7 Updated comments.
paulson
parents: 25243
diff changeset
   273
    where 'a is fixed. Probably we should delete clauses unless the sorts agree.
cfa3378decf7 Updated comments.
paulson
parents: 25243
diff changeset
   274
  * Currently we include a class constraint in the clause, exactly as with TVars.
cfa3378decf7 Updated comments.
paulson
parents: 25243
diff changeset
   275
*)
cfa3378decf7 Updated comments.
paulson
parents: 25243
diff changeset
   276
20015
1ffcf4802802 Literals aren't sorted any more. Output overloaded constants' type var instantiations.
mengj
parents: 19719
diff changeset
   277
(** make axiom and conjecture clauses. **)
1ffcf4802802 Literals aren't sorted any more. Output overloaded constants' type var instantiations.
mengj
parents: 19719
diff changeset
   278
1ffcf4802802 Literals aren't sorted any more. Output overloaded constants' type var instantiations.
mengj
parents: 19719
diff changeset
   279
fun get_tvar_strs [] = []
24940
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   280
  | get_tvar_strs ((TVar (indx,s))::Ts) =
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   281
      insert (op =) (make_schematic_type_var indx) (get_tvar_strs Ts)
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   282
  | get_tvar_strs((TFree _)::Ts) = get_tvar_strs Ts
20015
1ffcf4802802 Literals aren't sorted any more. Output overloaded constants' type var instantiations.
mengj
parents: 19719
diff changeset
   283
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   284
19354
aebf9dddccd7 tptp_write_file accepts axioms as thm.
mengj
parents: 19277
diff changeset
   285
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   286
(**** Isabelle arities ****)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   287
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   288
exception ARCLAUSE of string;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   289
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   290
datatype arLit = TConsLit of class * string * string list
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   291
               | TVarLit of class * string;
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   292
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   293
datatype arityClause =
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   294
         ArityClause of {axiom_name: axiom_name,
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   295
                         conclLit: arLit,
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   296
                         premLits: arLit list};
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   297
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   298
18798
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   299
fun gen_TVars 0 = []
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   300
  | gen_TVars n = ("T_" ^ Int.toString n) :: gen_TVars (n-1);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   301
18411
2d3165a0fb40 No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents: 18409
diff changeset
   302
fun pack_sort(_,[])  = []
2d3165a0fb40 No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents: 18409
diff changeset
   303
  | pack_sort(tvar, "HOL.type"::srt) = pack_sort(tvar, srt)   (*IGNORE sort "type"*)
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   304
  | pack_sort(tvar, cls::srt) =  (cls, tvar) :: pack_sort(tvar, srt);
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   305
18411
2d3165a0fb40 No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents: 18409
diff changeset
   306
(*Arity of type constructor tcon :: (arg1,...,argN)res*)
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   307
fun make_axiom_arity_clause dfg (tcons, axiom_name, (cls,args)) =
21560
d92389321fa7 tidied code
paulson
parents: 21509
diff changeset
   308
   let val tvars = gen_TVars (length args)
17845
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   309
       val tvars_srts = ListPair.zip (tvars,args)
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   310
   in
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   311
      ArityClause {axiom_name = axiom_name, 
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   312
                   conclLit = TConsLit (cls, make_fixed_type_const dfg tcons, tvars),
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   313
                   premLits = map TVarLit (union_all(map pack_sort tvars_srts))}
17845
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   314
   end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   315
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   316
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   317
(**** Isabelle class relations ****)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   318
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   319
datatype classrelClause =
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   320
         ClassrelClause of {axiom_name: axiom_name,
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   321
                            subclass: class,
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   322
                            superclass: class};
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   323
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21254
diff changeset
   324
(*Generate all pairs (sub,super) such that sub is a proper subclass of super in theory thy.*)
21432
625797c592b2 Optimized class_pairs for the common case of no subclasses
paulson
parents: 21416
diff changeset
   325
fun class_pairs thy [] supers = []
625797c592b2 Optimized class_pairs for the common case of no subclasses
paulson
parents: 21416
diff changeset
   326
  | class_pairs thy subs supers =
625797c592b2 Optimized class_pairs for the common case of no subclasses
paulson
parents: 21416
diff changeset
   327
      let val class_less = Sorts.class_less(Sign.classes_of thy)
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   328
          fun add_super sub (super,pairs) =
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   329
                if class_less (sub,super) then (sub,super)::pairs else pairs
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30151
diff changeset
   330
          fun add_supers (sub,pairs) = List.foldl (add_super sub) pairs supers
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30151
diff changeset
   331
      in  List.foldl add_supers [] subs  end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   332
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21254
diff changeset
   333
fun make_classrelClause (sub,super) =
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21254
diff changeset
   334
  ClassrelClause {axiom_name = clrelclause_prefix ^ ascii_of sub ^ "_" ^ ascii_of super,
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   335
                  subclass = make_type_class sub,
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21254
diff changeset
   336
                  superclass = make_type_class super};
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   337
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21254
diff changeset
   338
fun make_classrel_clauses thy subs supers =
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21254
diff changeset
   339
  map make_classrelClause (class_pairs thy subs supers);
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   340
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   341
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   342
(** Isabelle arities **)
17845
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   343
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   344
fun arity_clause dfg _ _ (tcons, []) = []
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   345
  | arity_clause dfg seen n (tcons, ("HOL.type",_)::ars) =  (*ignore*)
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   346
      arity_clause dfg seen n (tcons,ars)
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   347
  | arity_clause dfg seen n (tcons, (ar as (class,_)) :: ars) =
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21290
diff changeset
   348
      if class mem_string seen then (*multiple arities for the same tycon, class pair*)
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   349
          make_axiom_arity_clause dfg (tcons, lookup_type_const dfg tcons ^ "_" ^ class ^ "_" ^ Int.toString n, ar) ::
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   350
          arity_clause dfg seen (n+1) (tcons,ars)
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21290
diff changeset
   351
      else
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   352
          make_axiom_arity_clause dfg (tcons, lookup_type_const dfg tcons ^ "_" ^ class, ar) ::
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   353
          arity_clause dfg (class::seen) n (tcons,ars)
17845
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   354
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   355
fun multi_arity_clause dfg [] = []
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   356
  | multi_arity_clause dfg ((tcons,ars) :: tc_arlists) =
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   357
      arity_clause dfg [] 1 (tcons, ars)  @  multi_arity_clause dfg tc_arlists
17845
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   358
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   359
(*Generate all pairs (tycon,class,sorts) such that tycon belongs to class in theory thy
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   360
  provided its arguments have the corresponding sorts.*)
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21290
diff changeset
   361
fun type_class_pairs thy tycons classes =
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21290
diff changeset
   362
  let val alg = Sign.classes_of thy
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21290
diff changeset
   363
      fun domain_sorts (tycon,class) = Sorts.mg_domain alg tycon [class]
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   364
      fun add_class tycon (class,pairs) =
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   365
            (class, domain_sorts(tycon,class))::pairs
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21290
diff changeset
   366
            handle Sorts.CLASS_ERROR _ => pairs
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30151
diff changeset
   367
      fun try_classes tycon = (tycon, List.foldl (add_class tycon) [] classes)
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21290
diff changeset
   368
  in  map try_classes tycons  end;
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21290
diff changeset
   369
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   370
(*Proving one (tycon, class) membership may require proving others, so iterate.*)
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   371
fun iter_type_class_pairs thy tycons [] = ([], [])
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   372
  | iter_type_class_pairs thy tycons classes =
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   373
      let val cpairs = type_class_pairs thy tycons classes
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   374
          val newclasses = union_all (union_all (union_all (map (map #2 o #2) cpairs)))
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   375
            |> subtract (op =) classes |> subtract (op =) HOLogic.typeS
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   376
          val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33038
diff changeset
   377
      in (union (op =) classes' classes, union (op =) cpairs' cpairs) end;
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   378
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   379
fun make_arity_clauses_dfg dfg thy tycons classes =
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   380
  let val (classes', cpairs) = iter_type_class_pairs thy tycons classes
30151
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   381
  in  (classes', multi_arity_clause dfg cpairs)  end;
629f3a92863e removed global ref dfg_format
immler@in.tum.de
parents: 29676
diff changeset
   382
val make_arity_clauses = make_arity_clauses_dfg false;
17845
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   383
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   384
(**** Find occurrences of predicates in clauses ****)
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   385
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   386
(*FIXME: multiple-arity checking doesn't work, as update_new is the wrong
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   387
  function (it flags repeated declarations of a function, even with the same arity)*)
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   388
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30151
diff changeset
   389
fun update_many (tab, keypairs) = List.foldl (uncurry Symtab.update) tab keypairs;
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   390
24940
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   391
fun add_type_sort_preds (T, preds) =
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   392
      update_many (preds, map pred_of_sort (sorts_on_typs T));
17845
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   393
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   394
fun add_classrelClause_preds (ClassrelClause {subclass,superclass,...}, preds) =
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   395
  Symtab.update (subclass,1) (Symtab.update (superclass,1) preds);
17845
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   396
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   397
fun class_of_arityLit (TConsLit (tclass, _, _)) = tclass
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   398
  | class_of_arityLit (TVarLit (tclass, _)) = tclass;
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21290
diff changeset
   399
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21290
diff changeset
   400
fun add_arityClause_preds (ArityClause {conclLit,premLits,...}, preds) =
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   401
  let val classes = map (make_type_class o class_of_arityLit) (conclLit::premLits)
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21290
diff changeset
   402
      fun upd (class,preds) = Symtab.update (class,1) preds
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30151
diff changeset
   403
  in  List.foldl upd preds classes  end;
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   404
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   405
(*** Find occurrences of functions in clauses ***)
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   406
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   407
fun add_foltype_funcs (AtomV _, funcs) = funcs
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   408
  | add_foltype_funcs (AtomF a, funcs) = Symtab.update (a,0) funcs
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   409
  | add_foltype_funcs (Comp(a,tys), funcs) =
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30151
diff changeset
   410
      List.foldl add_foltype_funcs (Symtab.update (a, length tys) funcs) tys;
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   411
20038
73231d03a2ac Some tidying.
paulson
parents: 20022
diff changeset
   412
(*TFrees are recorded as constants*)
24940
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   413
fun add_type_sort_funcs (TVar _, funcs) = funcs
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   414
  | add_type_sort_funcs (TFree (a, _), funcs) =
20038
73231d03a2ac Some tidying.
paulson
parents: 20022
diff changeset
   415
      Symtab.update (make_fixed_type_var a, 0) funcs
73231d03a2ac Some tidying.
paulson
parents: 20022
diff changeset
   416
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   417
fun add_arityClause_funcs (ArityClause {conclLit,...}, funcs) =
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   418
  let val TConsLit (_, tcons, tvars) = conclLit
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   419
  in  Symtab.update (tcons, length tvars) funcs  end;
17845
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   420
23075
69e30a7e8880 Some hacks for SPASS format
paulson
parents: 23029
diff changeset
   421
(*This type can be overlooked because it is built-in...*)
69e30a7e8880 Some hacks for SPASS format
paulson
parents: 23029
diff changeset
   422
val init_functab = Symtab.update ("tc_itself", 1) Symtab.empty;
69e30a7e8880 Some hacks for SPASS format
paulson
parents: 23029
diff changeset
   423
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   424
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   425
(**** String-oriented operations ****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   426
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   427
fun string_of_clausename (cls_id,ax_name) =
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
   428
    clause_prefix ^ ascii_of ax_name ^ "_" ^ Int.toString cls_id;
17317
3f12de2e2e6e Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
paulson
parents: 17312
diff changeset
   429
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   430
fun string_of_type_clsname (cls_id,ax_name,idx) =
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
   431
    string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx);
18863
a113b6839df1 reorganization of code to support DFG otuput
paulson
parents: 18856
diff changeset
   432
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   433
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   434
(**** Producing DFG files ****)
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   435
18863
a113b6839df1 reorganization of code to support DFG otuput
paulson
parents: 18856
diff changeset
   436
(*Attach sign in DFG syntax: false means negate.*)
a113b6839df1 reorganization of code to support DFG otuput
paulson
parents: 18856
diff changeset
   437
fun dfg_sign true s = s
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   438
  | dfg_sign false s = "not(" ^ s ^ ")"
18863
a113b6839df1 reorganization of code to support DFG otuput
paulson
parents: 18856
diff changeset
   439
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   440
fun dfg_of_typeLit pos (LTVar (s,ty))  = dfg_sign pos (s ^ "(" ^ ty ^ ")")
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   441
  | dfg_of_typeLit pos (LTFree (s,ty)) = dfg_sign pos (s ^ "(" ^ ty ^ ")");
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   442
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   443
(*Enclose the clause body by quantifiers, if necessary*)
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   444
fun dfg_forall [] body = body
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   445
  | dfg_forall vars body = "forall([" ^ commas vars ^ "],\n" ^ body ^ ")"
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   446
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   447
fun gen_dfg_cls (cls_id, ax_name, Axiom, lits, tylits, vars) =
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   448
      "clause( %(axiom)\n" ^
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   449
      dfg_forall vars ("or(" ^ commas (tylits@lits) ^ ")") ^ ",\n" ^
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   450
      string_of_clausename (cls_id,ax_name) ^  ").\n\n"
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   451
  | gen_dfg_cls (cls_id, ax_name, Conjecture, lits, _, vars) =
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   452
      "clause( %(negated_conjecture)\n" ^
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   453
      dfg_forall vars ("or(" ^ commas lits ^ ")") ^ ",\n" ^
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   454
      string_of_clausename (cls_id,ax_name) ^  ").\n\n";
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   455
18798
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   456
fun string_of_arity (name, num) =  "(" ^ name ^ "," ^ Int.toString num ^ ")"
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   457
18856
4669dec681f4 tidy-up of res_clause.ML, removing the "predicates" field
paulson
parents: 18798
diff changeset
   458
fun string_of_preds [] = ""
4669dec681f4 tidy-up of res_clause.ML, removing the "predicates" field
paulson
parents: 18798
diff changeset
   459
  | string_of_preds preds = "predicates[" ^ commas(map string_of_arity preds) ^ "].\n";
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   460
18856
4669dec681f4 tidy-up of res_clause.ML, removing the "predicates" field
paulson
parents: 18798
diff changeset
   461
fun string_of_funcs [] = ""
4669dec681f4 tidy-up of res_clause.ML, removing the "predicates" field
paulson
parents: 18798
diff changeset
   462
  | string_of_funcs funcs = "functions[" ^ commas(map string_of_arity funcs) ^ "].\n" ;
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   463
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   464
fun string_of_symbols predstr funcstr =
17234
12a9393c5d77 further tidying up of Isabelle-ATP link
paulson
parents: 17230
diff changeset
   465
  "list_of_symbols.\n" ^ predstr  ^ funcstr  ^ "end_of_list.\n\n";
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   466
18798
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   467
fun string_of_start name = "begin_problem(" ^ name ^ ").\n\n";
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   468
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   469
fun string_of_descrip name =
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   470
  "list_of_descriptions.\nname({*" ^ name ^
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   471
  "*}).\nauthor({*Isabelle*}).\nstatus(unknown).\ndescription({*auto-generated*}).\nend_of_list.\n\n"
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   472
18863
a113b6839df1 reorganization of code to support DFG otuput
paulson
parents: 18856
diff changeset
   473
fun dfg_tfree_clause tfree_lit =
21509
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   474
  "clause( %(negated_conjecture)\n" ^ "or( " ^ tfree_lit ^ "),\n" ^ "tfree_tcs" ^ ").\n\n"
18863
a113b6839df1 reorganization of code to support DFG otuput
paulson
parents: 18856
diff changeset
   475
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   476
fun dfg_of_arLit (TConsLit (c,t,args)) =
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   477
      dfg_sign true (make_type_class c ^ "(" ^ t ^ paren_pack args ^ ")")
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   478
  | dfg_of_arLit (TVarLit (c,str)) =
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   479
      dfg_sign false (make_type_class c ^ "(" ^ str ^ ")")
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   480
20038
73231d03a2ac Some tidying.
paulson
parents: 20022
diff changeset
   481
fun dfg_classrelLits sub sup =  "not(" ^ sub ^ "(T)), " ^ sup ^ "(T)";
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
   482
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   483
fun dfg_classrelClause (ClassrelClause {axiom_name,subclass,superclass,...}) =
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   484
  "clause(forall([T],\nor( " ^ dfg_classrelLits subclass superclass ^ ")),\n" ^
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   485
  axiom_name ^ ").\n\n";
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   486
21560
d92389321fa7 tidied code
paulson
parents: 21509
diff changeset
   487
fun string_of_ar axiom_name = arclause_prefix ^ ascii_of axiom_name;
d92389321fa7 tidied code
paulson
parents: 21509
diff changeset
   488
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   489
fun dfg_arity_clause (ArityClause{axiom_name,conclLit,premLits,...}) =
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   490
  let val TConsLit (_,_,tvars) = conclLit
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   491
      val lits = map dfg_of_arLit (conclLit :: premLits)
18863
a113b6839df1 reorganization of code to support DFG otuput
paulson
parents: 18856
diff changeset
   492
  in
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   493
      "clause( %(axiom)\n" ^
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   494
      dfg_forall tvars ("or( " ^ commas lits ^ ")") ^ ",\n" ^
21560
d92389321fa7 tidied code
paulson
parents: 21509
diff changeset
   495
      string_of_ar axiom_name ^ ").\n\n"
18863
a113b6839df1 reorganization of code to support DFG otuput
paulson
parents: 18856
diff changeset
   496
  end;
a113b6839df1 reorganization of code to support DFG otuput
paulson
parents: 18856
diff changeset
   497
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   498
18869
00741f7280f7 removal of ResClause.num_of_clauses and other simplifications
paulson
parents: 18868
diff changeset
   499
(**** Produce TPTP files ****)
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   500
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   501
(*Attach sign in TPTP syntax: false means negate.*)
21509
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   502
fun tptp_sign true s = s
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   503
  | tptp_sign false s = "~ " ^ s
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   504
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   505
fun tptp_of_typeLit pos (LTVar (s,ty))  = tptp_sign pos (s ^ "(" ^ ty ^ ")")
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   506
  | tptp_of_typeLit pos (LTFree (s,ty)) = tptp_sign pos  (s ^ "(" ^ ty ^ ")");
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   507
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   508
fun gen_tptp_cls (cls_id,ax_name,Axiom,lits,tylits) =
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   509
      "cnf(" ^ string_of_clausename (cls_id,ax_name) ^ ",axiom," ^ 
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   510
               tptp_pack (tylits@lits) ^ ").\n"
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   511
  | gen_tptp_cls (cls_id,ax_name,Conjecture,lits,_) =
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   512
      "cnf(" ^ string_of_clausename (cls_id,ax_name) ^ ",negated_conjecture," ^ 
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   513
               tptp_pack lits ^ ").\n";
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   514
18863
a113b6839df1 reorganization of code to support DFG otuput
paulson
parents: 18856
diff changeset
   515
fun tptp_tfree_clause tfree_lit =
21509
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   516
    "cnf(" ^ "tfree_tcs," ^ "negated_conjecture" ^ "," ^ tptp_pack[tfree_lit] ^ ").\n";
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   517
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   518
fun tptp_of_arLit (TConsLit (c,t,args)) =
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   519
      tptp_sign true (make_type_class c ^ "(" ^ t ^ paren_pack args ^ ")")
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   520
  | tptp_of_arLit (TVarLit (c,str)) =
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   521
      tptp_sign false (make_type_class c ^ "(" ^ str ^ ")")
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   522
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   523
fun tptp_arity_clause (ArityClause{axiom_name,conclLit,premLits,...}) =
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   524
  "cnf(" ^ string_of_ar axiom_name ^ ",axiom," ^
21560
d92389321fa7 tidied code
paulson
parents: 21509
diff changeset
   525
  tptp_pack (map tptp_of_arLit (conclLit :: premLits)) ^ ").\n";
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   526
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   527
fun tptp_classrelLits sub sup =
21509
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   528
  let val tvar = "(T)"
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   529
  in  tptp_pack [tptp_sign false (sub^tvar), tptp_sign true (sup^tvar)]  end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   530
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   531
fun tptp_classrelClause (ClassrelClause {axiom_name,subclass,superclass,...}) =
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   532
  "cnf(" ^ axiom_name ^ ",axiom," ^ tptp_classrelLits subclass superclass ^ ").\n"
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   533
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   534
end;