src/HOL/Tools/Sledgehammer/sledgehammer_fol_clause.ML
author blanchet
Tue, 22 Jun 2010 16:23:29 +0200
changeset 37500 7587b6e63454
parent 37498 b426cbdb5a23
child 37509 f39464d971c4
permissions -rw-r--r--
thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
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
36393
be73a2b2443b support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents: 36378
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     4
33311
49cd8abb2863 proper header;
wenzelm
parents: 33043
diff changeset
     5
Storing/printing FOL clauses and arity clauses.  Typed equality is
49cd8abb2863 proper header;
wenzelm
parents: 33043
diff changeset
     6
treated differently.
49cd8abb2863 proper header;
wenzelm
parents: 33043
diff changeset
     7
36218
0e4a01f3e7d3 get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents: 36170
diff changeset
     8
FIXME: combine with sledgehammer_hol_clause!
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     9
*)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    10
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    11
signature SLEDGEHAMMER_FOL_CLAUSE =
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    12
sig
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    13
  val schematic_var_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    14
  val fixed_var_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    15
  val tvar_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    16
  val tfree_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    17
  val clause_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    18
  val const_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    19
  val tconst_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    20
  val class_prefix: string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    21
  val union_all: ''a list list -> ''a list
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    22
  val const_trans_table: string Symtab.table
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    23
  val type_const_trans_table: string Symtab.table
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    24
  val ascii_of: string -> string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    25
  val undo_ascii_of: string -> string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    26
  val paren_pack : string list -> string
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    27
  val make_schematic_var : string * int -> string
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
    28
  val make_fixed_var : string -> string
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
    29
  val make_schematic_type_var : string * int -> string
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    30
  val make_fixed_type_var : string -> string
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
    31
  val make_fixed_const : string -> string
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
    32
  val make_fixed_type_const : string -> string
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
    33
  val make_type_class : string -> string
36169
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
    34
  type name = string * string
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
    35
  type name_pool = string Symtab.table * string Symtab.table
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
    36
  val empty_name_pool : bool -> name_pool option
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
    37
  val pool_map : ('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
    38
  val nice_name : name -> name_pool option -> string * name_pool option
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    39
  datatype kind = Axiom | Conjecture
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    40
  datatype fol_type =
36169
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
    41
    TyVar of name |
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
    42
    TyFree of name |
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
    43
    TyConstr of name * fol_type list
36170
0cdb76723c88 added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents: 36169
diff changeset
    44
  val string_of_fol_type :
0cdb76723c88 added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents: 36169
diff changeset
    45
    fol_type -> name_pool option -> string * name_pool option
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
    46
  datatype type_literal =
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
    47
    TyLitVar of string * name |
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
    48
    TyLitFree of string * name
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    49
  exception CLAUSE of string * term
36966
adc11fb3f3aa generate proper arity declarations for TFrees for SPASS's DFG format;
blanchet
parents: 36692
diff changeset
    50
  val type_literals_for_types : typ list -> type_literal list
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    51
  datatype arLit =
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    52
      TConsLit of class * string * string list
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    53
    | TVarLit of class * string
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
    54
  datatype arity_clause = ArityClause of
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
    55
   {axiom_name: string, conclLit: arLit, premLits: arLit list}
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
    56
  datatype classrel_clause = ClassrelClause of
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
    57
   {axiom_name: string, subclass: class, superclass: class}
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
    58
  val make_classrel_clauses: theory -> class list -> class list -> classrel_clause list
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
    59
  val make_arity_clauses: theory -> string list -> class list -> class list * arity_clause list
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    60
  val tptp_sign: bool -> string -> string
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
    61
  val tptp_of_type_literal :
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
    62
    bool -> type_literal -> name_pool option -> string * name_pool option
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
    63
  val gen_tptp_cls : int * string * kind * string list * string list -> string
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    64
  val tptp_tfree_clause : string -> string
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
    65
  val tptp_arity_clause : arity_clause -> string
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
    66
  val tptp_classrel_clause : classrel_clause -> string
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    67
end
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    68
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    69
structure Sledgehammer_FOL_Clause : SLEDGEHAMMER_FOL_CLAUSE =
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    70
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    71
36063
cdc6855a6387 make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents: 36062
diff changeset
    72
open Sledgehammer_Util
cdc6855a6387 make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents: 36062
diff changeset
    73
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    74
val schematic_var_prefix = "V_";
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    75
val fixed_var_prefix = "v_";
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    76
17230
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
    77
val tvar_prefix = "T_";
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
    78
val tfree_prefix = "t_";
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    79
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    80
val clause_prefix = "cls_";
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    81
val arclause_prefix = "clsarity_"
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
    82
val clrelclause_prefix = "clsrel_";
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    83
17230
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
    84
val const_prefix = "c_";
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    85
val tconst_prefix = "tc_";
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    86
val class_prefix = "class_";
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    87
36218
0e4a01f3e7d3 get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents: 36170
diff changeset
    88
fun union_all xss = fold (union (op =)) xss []
17775
2679ba74411f minor code tidyig
paulson
parents: 17764
diff changeset
    89
36493
a3357a631b96 reintroduced short names for HOL->FOL constants; other parts of the code rely on these
blanchet
parents: 36491
diff changeset
    90
(* Readable names for the more common symbolic functions. Do not mess with the
37479
f6b1ee5b420b try to improve Sledgehammer/Metis's behavior in full_types mode, e.g. by handing True, False, and If better
blanchet
parents: 37414
diff changeset
    91
   last nine entries of the table unless you know what you are doing. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    92
val const_trans_table =
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
    93
  Symtab.make [(@{const_name "op ="}, "equal"),
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
    94
               (@{const_name "op &"}, "and"),
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
    95
               (@{const_name "op |"}, "or"),
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
    96
               (@{const_name "op -->"}, "implies"),
36493
a3357a631b96 reintroduced short names for HOL->FOL constants; other parts of the code rely on these
blanchet
parents: 36491
diff changeset
    97
               (@{const_name "op :"}, "in"),
a3357a631b96 reintroduced short names for HOL->FOL constants; other parts of the code rely on these
blanchet
parents: 36491
diff changeset
    98
               (@{const_name fequal}, "fequal"),
a3357a631b96 reintroduced short names for HOL->FOL constants; other parts of the code rely on these
blanchet
parents: 36491
diff changeset
    99
               (@{const_name COMBI}, "COMBI"),
a3357a631b96 reintroduced short names for HOL->FOL constants; other parts of the code rely on these
blanchet
parents: 36491
diff changeset
   100
               (@{const_name COMBK}, "COMBK"),
a3357a631b96 reintroduced short names for HOL->FOL constants; other parts of the code rely on these
blanchet
parents: 36491
diff changeset
   101
               (@{const_name COMBB}, "COMBB"),
a3357a631b96 reintroduced short names for HOL->FOL constants; other parts of the code rely on these
blanchet
parents: 36491
diff changeset
   102
               (@{const_name COMBC}, "COMBC"),
37479
f6b1ee5b420b try to improve Sledgehammer/Metis's behavior in full_types mode, e.g. by handing True, False, and If better
blanchet
parents: 37414
diff changeset
   103
               (@{const_name COMBS}, "COMBS"),
f6b1ee5b420b try to improve Sledgehammer/Metis's behavior in full_types mode, e.g. by handing True, False, and If better
blanchet
parents: 37414
diff changeset
   104
               (@{const_name True}, "True"),
f6b1ee5b420b try to improve Sledgehammer/Metis's behavior in full_types mode, e.g. by handing True, False, and If better
blanchet
parents: 37414
diff changeset
   105
               (@{const_name False}, "False"),
f6b1ee5b420b try to improve Sledgehammer/Metis's behavior in full_types mode, e.g. by handing True, False, and If better
blanchet
parents: 37414
diff changeset
   106
               (@{const_name If}, "If")]
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   107
17230
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
   108
val type_const_trans_table =
36476
a04cf4704668 in Sledgehammer "debug" mode, the names of most variables are already short and sweet, so most of the entries of the "const_trans_table" don't have a raison d'etre anymore
blanchet
parents: 36393
diff changeset
   109
  Symtab.make [(@{type_name "*"}, "prod"),
a04cf4704668 in Sledgehammer "debug" mode, the names of most variables are already short and sweet, so most of the entries of the "const_trans_table" don't have a raison d'etre anymore
blanchet
parents: 36393
diff changeset
   110
               (@{type_name "+"}, "sum")]
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   111
15610
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   112
(*Escaping of special characters.
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   113
  Alphanumeric characters are left unchanged.
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   114
  The character _ goes to __
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   115
  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
   116
  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
   117
val A_minus_space = Char.ord #"A" - Char.ord #" ";
15610
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   118
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   119
fun stringN_of_int 0 _ = ""
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   120
  | 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
   121
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   122
fun ascii_of_c c =
15610
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   123
  if Char.isAlphaNum c then String.str c
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   124
  else if c = #"_" then "__"
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   125
  else if #" " <= c andalso c <= #"/"
15610
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   126
       then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   127
  else if Char.isPrint c
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   128
       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
   129
  else ""
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   130
15610
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   131
val ascii_of = String.translate ascii_of_c;
f855fd163b62 more concise ASCII escaping
paulson
parents: 15608
diff changeset
   132
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   133
(** Remove ASCII armouring from names in proof files **)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   134
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   135
(*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
   136
  Also, the errors are "impossible" (hah!)*)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   137
fun undo_ascii_aux rcs [] = String.implode(rev rcs)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   138
  | undo_ascii_aux rcs [#"_"] = undo_ascii_aux (#"_"::rcs) []  (*ERROR*)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   139
      (*Three types of _ escapes: __, _A to _P, _nnn*)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   140
  | undo_ascii_aux rcs (#"_" :: #"_" :: cs) = undo_ascii_aux (#"_"::rcs) cs
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   141
  | undo_ascii_aux rcs (#"_" :: c :: cs) =
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   142
      if #"A" <= c andalso c<= #"P"  (*translation of #" " to #"/"*)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   143
      then undo_ascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   144
      else
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   145
        let val digits = List.take (c::cs, 3) handle Subscript => []
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   146
        in
24183
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   147
            case Int.fromString (String.implode digits) of
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   148
                NONE => undo_ascii_aux (c:: #"_"::rcs) cs  (*ERROR*)
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   149
              | 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
   150
        end
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   151
  | 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
   152
a46b758941a4 Code to undo the function ascii_of
paulson
parents: 23881
diff changeset
   153
val undo_ascii_of = undo_ascii_aux [] o String.explode;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   154
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
   155
(* 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
   156
fun paren_pack [] = ""   (*empty argument list*)
9a7ffce389c3 new treatment of polymorphic types, using Sign.const_typargs
paulson
parents: 18199
diff changeset
   157
  | paren_pack strings = "(" ^ commas strings ^ ")";
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
   158
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   159
fun tptp_clause strings = "(" ^ space_implode " | " strings ^ ")"
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
   160
16925
0fd7b1438d28 simpler variable names, and no types for monomorphic constants
paulson
parents: 16903
diff changeset
   161
(*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
   162
fun trim_type_var s =
0fd7b1438d28 simpler variable names, and no types for monomorphic constants
paulson
parents: 16903
diff changeset
   163
  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
   164
  else error ("trim_type: Malformed type variable encountered: " ^ s);
0fd7b1438d28 simpler variable names, and no types for monomorphic constants
paulson
parents: 16903
diff changeset
   165
16903
bf42a9071ad1 streamlined the tptp output
paulson
parents: 16794
diff changeset
   166
fun ascii_of_indexname (v,0) = ascii_of v
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17422
diff changeset
   167
  | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   168
17230
77e93bf303a5 fixed arities and restored changes that had gone missing
paulson
parents: 17150
diff changeset
   169
fun make_schematic_var v = schematic_var_prefix ^ (ascii_of_indexname v);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   170
fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x);
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   171
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   172
fun make_schematic_type_var (x,i) =
16925
0fd7b1438d28 simpler variable names, and no types for monomorphic constants
paulson
parents: 16903
diff changeset
   173
      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
   174
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
   175
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   176
fun lookup_const c =
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   177
  case Symtab.lookup const_trans_table c of
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   178
    SOME c' => c'
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   179
  | NONE => ascii_of c
23075
69e30a7e8880 Some hacks for SPASS format
paulson
parents: 23029
diff changeset
   180
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   181
fun lookup_type_const c =
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   182
  case Symtab.lookup type_const_trans_table c of
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   183
    SOME c' => c'
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   184
  | NONE => ascii_of c
18411
2d3165a0fb40 No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents: 18409
diff changeset
   185
36062
194cb6e3c13f get rid of Polyhash, since it's no longer used
blanchet
parents: 35865
diff changeset
   186
(* "op =" MUST BE "equal" because it's built into ATPs. *)
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   187
fun make_fixed_const @{const_name "op ="} = "equal"
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   188
  | make_fixed_const c = const_prefix ^ lookup_const c
18411
2d3165a0fb40 No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents: 18409
diff changeset
   189
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   190
fun make_fixed_type_const c = tconst_prefix ^ lookup_type_const c
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   191
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17234
diff changeset
   192
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
   193
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   194
36169
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   195
(**** name pool ****)
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   196
 
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   197
type name = string * string
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   198
type name_pool = string Symtab.table * string Symtab.table
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   199
36222
0e3e49bd658d don't use readable names if proof reconstruction is needed, because it uses the structure of names
blanchet
parents: 36221
diff changeset
   200
fun empty_name_pool readable_names =
0e3e49bd658d don't use readable names if proof reconstruction is needed, because it uses the structure of names
blanchet
parents: 36221
diff changeset
   201
  if readable_names then SOME (`I Symtab.empty) else NONE
36169
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   202
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   203
fun pool_fold f xs z = pair z #> fold_rev (fn x => uncurry (f x)) xs
36169
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   204
fun pool_map f xs =
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   205
  pool_fold (fn x => fn ys => fn pool => f x pool |>> (fn y => y :: ys)) xs []
36169
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   206
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   207
fun add_nice_name full_name nice_prefix j the_pool =
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   208
  let
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   209
    val nice_name = nice_prefix ^ (if j = 0 then "" else "_" ^ Int.toString j)
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   210
  in
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   211
    case Symtab.lookup (snd the_pool) nice_name of
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   212
      SOME full_name' =>
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   213
      if full_name = full_name' then (nice_name, the_pool)
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   214
      else add_nice_name full_name nice_prefix (j + 1) the_pool
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   215
    | NONE =>
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   216
      (nice_name, (Symtab.update_new (full_name, nice_name) (fst the_pool),
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   217
                   Symtab.update_new (nice_name, full_name) (snd the_pool)))
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   218
  end
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   219
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   220
fun translate_first_char f s =
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   221
  String.str (f (String.sub (s, 0))) ^ String.extract (s, 1, NONE)
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   222
36222
0e3e49bd658d don't use readable names if proof reconstruction is needed, because it uses the structure of names
blanchet
parents: 36221
diff changeset
   223
fun readable_name full_name s =
36169
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   224
  let
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   225
    val s = s |> Long_Name.base_name
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   226
              |> fold remove_all ["\<^sub>", "\<^bsub>", "\<^esub>", "\<^isub>"]
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   227
    val s' = s |> explode |> rev |> dropwhile (curry (op =) "'")
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   228
    val s' =
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   229
      (s' |> rev
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   230
          |> implode
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   231
          |> String.translate
36221
3abbae8a10cd allow "_" in TPTP names in debug mode
blanchet
parents: 36218
diff changeset
   232
                 (fn c => if Char.isAlphaNum c orelse c = #"_" then String.str c
3abbae8a10cd allow "_" in TPTP names in debug mode
blanchet
parents: 36218
diff changeset
   233
                          else ""))
36169
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   234
      ^ replicate_string (String.size s - length s') "_"
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   235
    val s' =
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   236
      if s' = "" orelse not (Char.isAlpha (String.sub (s', 0))) then "X" ^ s'
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   237
      else s'
36476
a04cf4704668 in Sledgehammer "debug" mode, the names of most variables are already short and sweet, so most of the entries of the "const_trans_table" don't have a raison d'etre anymore
blanchet
parents: 36393
diff changeset
   238
    (* Avoid "equal", since it's built into ATPs; and "op" is very ambiguous
a04cf4704668 in Sledgehammer "debug" mode, the names of most variables are already short and sweet, so most of the entries of the "const_trans_table" don't have a raison d'etre anymore
blanchet
parents: 36393
diff changeset
   239
       ("op &", "op |", etc.). *)
a04cf4704668 in Sledgehammer "debug" mode, the names of most variables are already short and sweet, so most of the entries of the "const_trans_table" don't have a raison d'etre anymore
blanchet
parents: 36393
diff changeset
   240
    val s' = if s' = "equal" orelse s' = "op" then full_name else s'
36169
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   241
  in
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   242
    case (Char.isLower (String.sub (full_name, 0)),
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   243
          Char.isLower (String.sub (s', 0))) of
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   244
      (true, false) => translate_first_char Char.toLower s'
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   245
    | (false, true) => translate_first_char Char.toUpper s'
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   246
    | _ => s'
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   247
  end
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   248
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   249
fun nice_name (full_name, _) NONE = (full_name, NONE)
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   250
  | nice_name (full_name, desired_name) (SOME the_pool) =
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   251
    case Symtab.lookup (fst the_pool) full_name of
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   252
      SOME nice_name => (nice_name, SOME the_pool)
36222
0e3e49bd658d don't use readable names if proof reconstruction is needed, because it uses the structure of names
blanchet
parents: 36221
diff changeset
   253
    | NONE => add_nice_name full_name (readable_name full_name desired_name) 0
0e3e49bd658d don't use readable names if proof reconstruction is needed, because it uses the structure of names
blanchet
parents: 36221
diff changeset
   254
                            the_pool
36169
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   255
              |> apsnd SOME
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   256
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   257
(**** Definitions and functions for FOL clauses for TPTP format output ****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   258
21509
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   259
datatype kind = Axiom | Conjecture;
23385
0ef4f9fc0d09 Deleted unused code
paulson
parents: 23075
diff changeset
   260
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   261
(**** Isabelle FOL clauses ****)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   262
36168
0a6ed065683d give more sensible names to "fol_type" constructors, as requested by a FIXME comment
blanchet
parents: 36167
diff changeset
   263
datatype fol_type =
36169
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   264
  TyVar of name |
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   265
  TyFree of name |
27b1cc58715e store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents: 36168
diff changeset
   266
  TyConstr of name * fol_type list
18402
aaba095cf62b 1. changed fol_type, it's not a string type anymore.
mengj
parents: 18390
diff changeset
   267
36170
0cdb76723c88 added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents: 36169
diff changeset
   268
fun string_of_fol_type (TyVar sp) pool = nice_name sp pool
0cdb76723c88 added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents: 36169
diff changeset
   269
  | string_of_fol_type (TyFree sp) pool = nice_name sp pool
0cdb76723c88 added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents: 36169
diff changeset
   270
  | string_of_fol_type (TyConstr (sp, tys)) pool =
0cdb76723c88 added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents: 36169
diff changeset
   271
    let
0cdb76723c88 added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents: 36169
diff changeset
   272
      val (s, pool) = nice_name sp pool
0cdb76723c88 added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents: 36169
diff changeset
   273
      val (ss, pool) = pool_map string_of_fol_type tys pool
0cdb76723c88 added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents: 36169
diff changeset
   274
    in (s ^ paren_pack ss, pool) end
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   275
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   276
(* The first component is the type class; the second is a TVar or TFree. *)
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   277
datatype type_literal =
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   278
  TyLitVar of string * name |
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   279
  TyLitFree of string * name
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   280
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17375
diff changeset
   281
exception CLAUSE of string * term;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   282
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   283
(*Make literals for sorted type variables*)
24940
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   284
fun sorts_on_typs_aux (_, [])   = []
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   285
  | sorts_on_typs_aux ((x,i),  s::ss) =
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   286
      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
   287
      in
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   288
          if s = "HOL.type" then sorts
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   289
          else if i = ~1 then TyLitFree (make_type_class s, `make_fixed_type_var x) :: sorts
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   290
          else TyLitVar (make_type_class s, (make_schematic_type_var (x,i), x)) :: sorts
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   291
      end;
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   292
24940
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   293
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
   294
  | sorts_on_typs (TVar (v,s))  = sorts_on_typs_aux (v,s);
8f9dea697b8d getting rid of type typ_var
paulson
parents: 24937
diff changeset
   295
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24322
diff changeset
   296
(*Given a list of sorted type variables, return a list of type literals.*)
36966
adc11fb3f3aa generate proper arity declarations for TFrees for SPASS's DFG format;
blanchet
parents: 36692
diff changeset
   297
fun type_literals_for_types Ts =
adc11fb3f3aa generate proper arity declarations for TFrees for SPASS's DFG format;
blanchet
parents: 36692
diff changeset
   298
  fold (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
   299
1ffcf4802802 Literals aren't sorted any more. Output overloaded constants' type var instantiations.
mengj
parents: 19719
diff changeset
   300
(** make axiom and conjecture clauses. **)
1ffcf4802802 Literals aren't sorted any more. Output overloaded constants' type var instantiations.
mengj
parents: 19719
diff changeset
   301
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   302
(**** Isabelle arities ****)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   303
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   304
datatype arLit = TConsLit of class * string * string list
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   305
               | TVarLit of class * string;
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   306
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
   307
datatype arity_clause =
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   308
  ArityClause of {axiom_name: string, conclLit: arLit, premLits: arLit list}
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   309
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   310
18798
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   311
fun gen_TVars 0 = []
ca02a2077955 tidying up SPASS output
paulson
parents: 18676
diff changeset
   312
  | gen_TVars n = ("T_" ^ Int.toString n) :: gen_TVars (n-1);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   313
18411
2d3165a0fb40 No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents: 18409
diff changeset
   314
fun pack_sort(_,[])  = []
2d3165a0fb40 No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents: 18409
diff changeset
   315
  | 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
   316
  | pack_sort(tvar, cls::srt) =  (cls, tvar) :: pack_sort(tvar, srt);
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   317
18411
2d3165a0fb40 No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents: 18409
diff changeset
   318
(*Arity of type constructor tcon :: (arg1,...,argN)res*)
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   319
fun make_axiom_arity_clause (tcons, axiom_name, (cls,args)) =
21560
d92389321fa7 tidied code
paulson
parents: 21509
diff changeset
   320
   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
   321
       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
   322
   in
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   323
     ArityClause {axiom_name = axiom_name, 
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   324
                  conclLit = TConsLit (cls, make_fixed_type_const tcons, tvars),
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   325
                  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
   326
   end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   327
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   328
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   329
(**** Isabelle class relations ****)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   330
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
   331
datatype classrel_clause =
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   332
  ClassrelClause of {axiom_name: string, subclass: class, superclass: class}
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   333
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21254
diff changeset
   334
(*Generate all pairs (sub,super) such that sub is a proper subclass of super in theory thy.*)
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   335
fun class_pairs _ [] _ = []
21432
625797c592b2 Optimized class_pairs for the common case of no subclasses
paulson
parents: 21416
diff changeset
   336
  | class_pairs thy subs supers =
36218
0e4a01f3e7d3 get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents: 36170
diff changeset
   337
      let
0e4a01f3e7d3 get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents: 36170
diff changeset
   338
        val class_less = Sorts.class_less (Sign.classes_of thy)
0e4a01f3e7d3 get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents: 36170
diff changeset
   339
        fun add_super sub super = class_less (sub, super) ? cons (sub, super)
0e4a01f3e7d3 get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents: 36170
diff changeset
   340
        fun add_supers sub = fold (add_super sub) supers
0e4a01f3e7d3 get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents: 36170
diff changeset
   341
      in fold add_supers subs [] end
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   342
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
   343
fun make_classrel_clause (sub,super) =
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21254
diff changeset
   344
  ClassrelClause {axiom_name = clrelclause_prefix ^ ascii_of sub ^ "_" ^ ascii_of super,
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   345
                  subclass = make_type_class sub,
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21254
diff changeset
   346
                  superclass = make_type_class super};
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   347
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21254
diff changeset
   348
fun make_classrel_clauses thy subs supers =
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35827
diff changeset
   349
  map make_classrel_clause (class_pairs thy subs supers);
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   350
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   351
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   352
(** Isabelle arities **)
17845
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   353
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   354
fun arity_clause _ _ (_, []) = []
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   355
  | arity_clause seen n (tcons, ("HOL.type",_)::ars) =  (*ignore*)
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   356
      arity_clause seen n (tcons,ars)
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   357
  | arity_clause seen n (tcons, (ar as (class,_)) :: ars) =
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36556
diff changeset
   358
      if member (op =) seen class then (*multiple arities for the same tycon, class pair*)
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   359
          make_axiom_arity_clause (tcons, lookup_type_const tcons ^ "_" ^ class ^ "_" ^ Int.toString n, ar) ::
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   360
          arity_clause 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
   361
      else
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   362
          make_axiom_arity_clause (tcons, lookup_type_const tcons ^ "_" ^ class, ar) ::
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   363
          arity_clause (class::seen) n (tcons,ars)
17845
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   364
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   365
fun multi_arity_clause [] = []
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   366
  | multi_arity_clause ((tcons, ars) :: tc_arlists) =
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   367
      arity_clause [] 1 (tcons, ars) @ multi_arity_clause tc_arlists
17845
1438291d57f0 deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents: 17775
diff changeset
   368
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   369
(*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
   370
  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
   371
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
   372
  let val alg = Sign.classes_of thy
36218
0e4a01f3e7d3 get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents: 36170
diff changeset
   373
      fun domain_sorts tycon = Sorts.mg_domain alg tycon o single
0e4a01f3e7d3 get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents: 36170
diff changeset
   374
      fun add_class tycon class =
0e4a01f3e7d3 get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents: 36170
diff changeset
   375
        cons (class, domain_sorts tycon class)
0e4a01f3e7d3 get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents: 36170
diff changeset
   376
        handle Sorts.CLASS_ERROR _ => I
0e4a01f3e7d3 get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents: 36170
diff changeset
   377
      fun try_classes tycon = (tycon, fold (add_class tycon) classes [])
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21290
diff changeset
   378
  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
   379
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   380
(*Proving one (tycon, class) membership may require proving others, so iterate.*)
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   381
fun iter_type_class_pairs _ _ [] = ([], [])
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   382
  | iter_type_class_pairs thy tycons classes =
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   383
      let val cpairs = type_class_pairs thy tycons classes
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   384
          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
   385
            |> subtract (op =) classes |> subtract (op =) HOLogic.typeS
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   386
          val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33038
diff changeset
   387
      in (union (op =) classes' classes, union (op =) cpairs' cpairs) end;
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   388
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   389
fun make_arity_clauses thy tycons classes =
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   390
  let val (classes', cpairs) = iter_type_class_pairs thy tycons classes
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   391
  in  (classes', multi_arity_clause cpairs)  end;
18863
a113b6839df1 reorganization of code to support DFG otuput
paulson
parents: 18856
diff changeset
   392
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   393
18869
00741f7280f7 removal of ResClause.num_of_clauses and other simplifications
paulson
parents: 18868
diff changeset
   394
(**** Produce TPTP files ****)
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   395
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   396
fun string_of_clausename (cls_id, ax_name) =
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   397
    clause_prefix ^ ascii_of ax_name ^ "_" ^ Int.toString cls_id
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   398
21509
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   399
fun tptp_sign true s = s
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   400
  | tptp_sign false s = "~ " ^ s
18868
7cfc21ee0370 working SPASS support; much tidying
paulson
parents: 18863
diff changeset
   401
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   402
fun tptp_of_type_literal pos (TyLitVar (s, name)) =
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   403
    nice_name name #>> (fn s' => tptp_sign pos (s ^ "(" ^ s' ^ ")"))
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   404
  | tptp_of_type_literal pos (TyLitFree (s, name)) =
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   405
    nice_name name #>> (fn s' => tptp_sign pos (s ^ "(" ^ s' ^ ")"))
36167
c1a35be8e476 make Sledgehammer's output more debugging friendly
blanchet
parents: 36063
diff changeset
   406
c1a35be8e476 make Sledgehammer's output more debugging friendly
blanchet
parents: 36063
diff changeset
   407
fun tptp_cnf name kind formula =
c1a35be8e476 make Sledgehammer's output more debugging friendly
blanchet
parents: 36063
diff changeset
   408
  "cnf(" ^ name ^ ", " ^ kind ^ ",\n    " ^ formula ^ ").\n"
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   409
36167
c1a35be8e476 make Sledgehammer's output more debugging friendly
blanchet
parents: 36063
diff changeset
   410
fun gen_tptp_cls (cls_id, ax_name, Axiom, lits, tylits) =
c1a35be8e476 make Sledgehammer's output more debugging friendly
blanchet
parents: 36063
diff changeset
   411
      tptp_cnf (string_of_clausename (cls_id, ax_name)) "axiom"
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   412
               (tptp_clause (tylits @ lits))
36167
c1a35be8e476 make Sledgehammer's output more debugging friendly
blanchet
parents: 36063
diff changeset
   413
  | gen_tptp_cls (cls_id, ax_name, Conjecture, lits, _) =
c1a35be8e476 make Sledgehammer's output more debugging friendly
blanchet
parents: 36063
diff changeset
   414
      tptp_cnf (string_of_clausename (cls_id, ax_name)) "negated_conjecture"
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   415
               (tptp_clause lits)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   416
18863
a113b6839df1 reorganization of code to support DFG otuput
paulson
parents: 18856
diff changeset
   417
fun tptp_tfree_clause tfree_lit =
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   418
    tptp_cnf "tfree_tcs" "negated_conjecture" (tptp_clause [tfree_lit])
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   419
22643
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   420
fun tptp_of_arLit (TConsLit (c,t,args)) =
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   421
      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
   422
  | tptp_of_arLit (TVarLit (c,str)) =
bc3bb8e9594a Improved and simplified the treatment of classrel/arity clauses
paulson
parents: 22383
diff changeset
   423
      tptp_sign false (make_type_class c ^ "(" ^ str ^ ")")
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   424
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   425
fun tptp_arity_clause (ArityClause {axiom_name, conclLit, premLits, ...}) =
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   426
  tptp_cnf (arclause_prefix ^ ascii_of axiom_name) "axiom"
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   427
           (tptp_clause (map tptp_of_arLit (conclLit :: premLits)))
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   428
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
   429
fun tptp_classrelLits sub sup =
21509
6c5755ad9cae ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents: 21470
diff changeset
   430
  let val tvar = "(T)"
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36493
diff changeset
   431
  in  tptp_clause [tptp_sign false (sub^tvar), tptp_sign true (sup^tvar)]  end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   432
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   433
fun tptp_classrel_clause (ClassrelClause {axiom_name, subclass, superclass,
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   434
                                          ...}) =
36167
c1a35be8e476 make Sledgehammer's output more debugging friendly
blanchet
parents: 36063
diff changeset
   435
  tptp_cnf axiom_name "axiom" (tptp_classrelLits subclass superclass)
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
   436
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   437
end;