author | blanchet |
Fri, 25 Jun 2010 16:27:53 +0200 | |
changeset 37575 | cfc5e281740f |
parent 37572 | a899f9506f39 |
child 37577 | 5379f41a1322 |
permissions | -rw-r--r-- |
35826 | 1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_fol_clause.ML |
33311 | 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 | 4 |
|
33311 | 5 |
Storing/printing FOL clauses and arity clauses. Typed equality is |
6 |
treated differently. |
|
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 | 9 |
*) |
10 |
||
35826 | 11 |
signature SLEDGEHAMMER_FOL_CLAUSE = |
24310 | 12 |
sig |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37500
diff
changeset
|
13 |
val type_wrapper_name : string |
24310 | 14 |
val schematic_var_prefix: string |
15 |
val fixed_var_prefix: string |
|
16 |
val tvar_prefix: string |
|
17 |
val tfree_prefix: string |
|
18 |
val const_prefix: string |
|
19 |
val tconst_prefix: string |
|
20 |
val class_prefix: string |
|
21 |
val union_all: ''a list list -> ''a list |
|
37572
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
22 |
val invert_const: string -> string |
24310 | 23 |
val ascii_of: string -> string |
24 |
val undo_ascii_of: string -> string |
|
37572
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
25 |
val strip_prefix: string -> string -> string option |
24310 | 26 |
val make_schematic_var : string * int -> string |
18868 | 27 |
val make_fixed_var : string -> string |
28 |
val make_schematic_type_var : string * int -> string |
|
24310 | 29 |
val make_fixed_type_var : string -> string |
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
30 |
val make_fixed_const : string -> string |
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
31 |
val make_fixed_type_const : string -> string |
18868 | 32 |
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
|
33 |
type name = string * string |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
34 |
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
|
35 |
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
|
36 |
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
|
37 |
val nice_name : name -> name_pool option -> string * name_pool option |
24310 | 38 |
datatype kind = Axiom | 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
|
39 |
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
|
40 |
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
|
41 |
TyLitFree of string * name |
36966
adc11fb3f3aa
generate proper arity declarations for TFrees for SPASS's DFG format;
blanchet
parents:
36692
diff
changeset
|
42 |
val type_literals_for_types : typ list -> type_literal list |
24310 | 43 |
datatype arLit = |
44 |
TConsLit of class * string * string list |
|
45 |
| TVarLit of class * string |
|
35865 | 46 |
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
|
47 |
{axiom_name: string, conclLit: arLit, premLits: arLit list} |
35865 | 48 |
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
|
49 |
{axiom_name: string, subclass: class, superclass: class} |
35865 | 50 |
val make_classrel_clauses: theory -> class list -> class list -> classrel_clause list |
51 |
val make_arity_clauses: theory -> string list -> class list -> class list * arity_clause list |
|
24310 | 52 |
end |
15347 | 53 |
|
35826 | 54 |
structure Sledgehammer_FOL_Clause : SLEDGEHAMMER_FOL_CLAUSE = |
15347 | 55 |
struct |
56 |
||
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37500
diff
changeset
|
57 |
val type_wrapper_name = "ti" |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37500
diff
changeset
|
58 |
|
15347 | 59 |
val schematic_var_prefix = "V_"; |
60 |
val fixed_var_prefix = "v_"; |
|
61 |
||
17230
77e93bf303a5
fixed arities and restored changes that had gone missing
paulson
parents:
17150
diff
changeset
|
62 |
val tvar_prefix = "T_"; |
77e93bf303a5
fixed arities and restored changes that had gone missing
paulson
parents:
17150
diff
changeset
|
63 |
val tfree_prefix = "t_"; |
15347 | 64 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37500
diff
changeset
|
65 |
val classrel_clause_prefix = "clsrel_"; |
15347 | 66 |
|
17230
77e93bf303a5
fixed arities and restored changes that had gone missing
paulson
parents:
17150
diff
changeset
|
67 |
val const_prefix = "c_"; |
24310 | 68 |
val tconst_prefix = "tc_"; |
69 |
val class_prefix = "class_"; |
|
15347 | 70 |
|
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36170
diff
changeset
|
71 |
fun union_all xss = fold (union (op =)) xss [] |
17775 | 72 |
|
36493
a3357a631b96
reintroduced short names for HOL->FOL constants; other parts of the code rely on these
blanchet
parents:
36491
diff
changeset
|
73 |
(* 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
|
74 |
last nine entries of the table unless you know what you are doing. *) |
15347 | 75 |
val const_trans_table = |
35865 | 76 |
Symtab.make [(@{const_name "op ="}, "equal"), |
77 |
(@{const_name "op &"}, "and"), |
|
78 |
(@{const_name "op |"}, "or"), |
|
79 |
(@{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
|
80 |
(@{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
|
81 |
(@{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
|
82 |
(@{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
|
83 |
(@{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
|
84 |
(@{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
|
85 |
(@{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
|
86 |
(@{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
|
87 |
(@{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
|
88 |
(@{const_name False}, "False"), |
37572
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
89 |
(@{const_name If}, "If"), |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
90 |
(@{type_name "*"}, "prod"), |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
91 |
(@{type_name "+"}, "sum")] |
15347 | 92 |
|
37572
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
93 |
(* Invert the table of translations between Isabelle and ATPs. *) |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
94 |
val const_trans_table_inv = |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
95 |
Symtab.update ("fequal", @{const_name "op ="}) |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
96 |
(Symtab.make (map swap (Symtab.dest const_trans_table))) |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
97 |
|
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
98 |
val invert_const = perhaps (Symtab.lookup const_trans_table_inv) |
15347 | 99 |
|
15610 | 100 |
(*Escaping of special characters. |
101 |
Alphanumeric characters are left unchanged. |
|
102 |
The character _ goes to __ |
|
103 |
Characters in the range ASCII space to / go to _A to _P, respectively. |
|
24183 | 104 |
Other printing characters go to _nnn where nnn is the decimal ASCII code.*) |
105 |
val A_minus_space = Char.ord #"A" - Char.ord #" "; |
|
15610 | 106 |
|
24183 | 107 |
fun stringN_of_int 0 _ = "" |
108 |
| stringN_of_int k n = stringN_of_int (k-1) (n div 10) ^ Int.toString (n mod 10); |
|
15610 | 109 |
|
15347 | 110 |
fun ascii_of_c c = |
15610 | 111 |
if Char.isAlphaNum c then String.str c |
112 |
else if c = #"_" then "__" |
|
24310 | 113 |
else if #" " <= c andalso c <= #"/" |
15610 | 114 |
then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space)) |
24310 | 115 |
else if Char.isPrint c |
24183 | 116 |
then ("_" ^ stringN_of_int 3 (Char.ord c)) (*fixed width, in case more digits follow*) |
15610 | 117 |
else "" |
15347 | 118 |
|
15610 | 119 |
val ascii_of = String.translate ascii_of_c; |
120 |
||
24183 | 121 |
(** Remove ASCII armouring from names in proof files **) |
122 |
||
123 |
(*We don't raise error exceptions because this code can run inside the watcher. |
|
124 |
Also, the errors are "impossible" (hah!)*) |
|
125 |
fun undo_ascii_aux rcs [] = String.implode(rev rcs) |
|
126 |
| undo_ascii_aux rcs [#"_"] = undo_ascii_aux (#"_"::rcs) [] (*ERROR*) |
|
127 |
(*Three types of _ escapes: __, _A to _P, _nnn*) |
|
128 |
| undo_ascii_aux rcs (#"_" :: #"_" :: cs) = undo_ascii_aux (#"_"::rcs) cs |
|
24310 | 129 |
| undo_ascii_aux rcs (#"_" :: c :: cs) = |
24183 | 130 |
if #"A" <= c andalso c<= #"P" (*translation of #" " to #"/"*) |
131 |
then undo_ascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs |
|
24310 | 132 |
else |
24183 | 133 |
let val digits = List.take (c::cs, 3) handle Subscript => [] |
24310 | 134 |
in |
24183 | 135 |
case Int.fromString (String.implode digits) of |
136 |
NONE => undo_ascii_aux (c:: #"_"::rcs) cs (*ERROR*) |
|
137 |
| SOME n => undo_ascii_aux (Char.chr n :: rcs) (List.drop (cs, 2)) |
|
138 |
end |
|
139 |
| undo_ascii_aux rcs (c::cs) = undo_ascii_aux (c::rcs) cs; |
|
140 |
||
141 |
val undo_ascii_of = undo_ascii_aux [] o String.explode; |
|
15347 | 142 |
|
37572
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
143 |
(* If string s has the prefix s1, return the result of deleting it, |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
144 |
un-ASCII'd. *) |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
145 |
fun strip_prefix s1 s = |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
146 |
if String.isPrefix s1 s then |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
147 |
SOME (undo_ascii_of (String.extract (s, size s1, NONE))) |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
148 |
else |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
149 |
NONE |
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
150 |
|
16925
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
paulson
parents:
16903
diff
changeset
|
151 |
(*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
|
152 |
fun trim_type_var s = |
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
paulson
parents:
16903
diff
changeset
|
153 |
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
|
154 |
else error ("trim_type: Malformed type variable encountered: " ^ s); |
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
paulson
parents:
16903
diff
changeset
|
155 |
|
16903 | 156 |
fun ascii_of_indexname (v,0) = ascii_of v |
17525
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17422
diff
changeset
|
157 |
| ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i; |
15347 | 158 |
|
17230
77e93bf303a5
fixed arities and restored changes that had gone missing
paulson
parents:
17150
diff
changeset
|
159 |
fun make_schematic_var v = schematic_var_prefix ^ (ascii_of_indexname v); |
15347 | 160 |
fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x); |
161 |
||
24310 | 162 |
fun make_schematic_type_var (x,i) = |
16925
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
paulson
parents:
16903
diff
changeset
|
163 |
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
|
164 |
fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x)); |
15347 | 165 |
|
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
166 |
fun lookup_const c = |
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
167 |
case Symtab.lookup const_trans_table c of |
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
168 |
SOME c' => c' |
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
169 |
| NONE => ascii_of c |
23075 | 170 |
|
36062 | 171 |
(* "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
|
172 |
fun make_fixed_const @{const_name "op ="} = "equal" |
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
173 |
| 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
|
174 |
|
37572
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
175 |
fun make_fixed_type_const c = tconst_prefix ^ lookup_const c |
17150
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
16976
diff
changeset
|
176 |
|
17261 | 177 |
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
|
178 |
|
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
16976
diff
changeset
|
179 |
|
36169
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
180 |
(**** name pool ****) |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
181 |
|
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
182 |
type name = string * string |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
183 |
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
|
184 |
|
36222
0e3e49bd658d
don't use readable names if proof reconstruction is needed, because it uses the structure of names
blanchet
parents:
36221
diff
changeset
|
185 |
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
|
186 |
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
|
187 |
|
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
|
188 |
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
|
189 |
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
|
190 |
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
|
191 |
|
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
192 |
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
|
193 |
let |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
194 |
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
|
195 |
in |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
196 |
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
|
197 |
SOME full_name' => |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
198 |
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
|
199 |
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
|
200 |
| NONE => |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
201 |
(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
|
202 |
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
|
203 |
end |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
204 |
|
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
205 |
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
|
206 |
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
|
207 |
|
36222
0e3e49bd658d
don't use readable names if proof reconstruction is needed, because it uses the structure of names
blanchet
parents:
36221
diff
changeset
|
208 |
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
|
209 |
let |
37575
cfc5e281740f
exploit "Name.desymbolize" to remove some dependencies
blanchet
parents:
37572
diff
changeset
|
210 |
val s = s |> Long_Name.base_name |> Name.desymbolize false |
36169
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
211 |
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
|
212 |
val s' = |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
213 |
(s' |> rev |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
214 |
|> implode |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
215 |
|> String.translate |
36221 | 216 |
(fn c => if Char.isAlphaNum c orelse c = #"_" then String.str c |
217 |
else "")) |
|
36169
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
218 |
^ 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
|
219 |
val s' = |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
220 |
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
|
221 |
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
|
222 |
(* 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
|
223 |
("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
|
224 |
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
|
225 |
in |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
226 |
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
|
227 |
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
|
228 |
(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
|
229 |
| (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
|
230 |
| _ => s' |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
231 |
end |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
232 |
|
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
233 |
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
|
234 |
| 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
|
235 |
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
|
236 |
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
|
237 |
| 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
|
238 |
the_pool |
36169
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
239 |
|> apsnd SOME |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
240 |
|
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
241 |
(**** Definitions and functions for FOL clauses for TPTP format output ****) |
15347 | 242 |
|
21509
6c5755ad9cae
ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents:
21470
diff
changeset
|
243 |
datatype kind = Axiom | Conjecture; |
23385 | 244 |
|
15347 | 245 |
(**** Isabelle FOL clauses ****) |
246 |
||
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
|
247 |
(* 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
|
248 |
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
|
249 |
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
|
250 |
TyLitFree of string * name |
15347 | 251 |
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17375
diff
changeset
|
252 |
exception CLAUSE of string * term; |
15347 | 253 |
|
24310 | 254 |
(*Make literals for sorted type variables*) |
24940 | 255 |
fun sorts_on_typs_aux (_, []) = [] |
256 |
| sorts_on_typs_aux ((x,i), s::ss) = |
|
257 |
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
|
258 |
in |
bc3bb8e9594a
Improved and simplified the treatment of classrel/arity clauses
paulson
parents:
22383
diff
changeset
|
259 |
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
|
260 |
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
|
261 |
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
|
262 |
end; |
17150
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
16976
diff
changeset
|
263 |
|
24940 | 264 |
fun sorts_on_typs (TFree (a,s)) = sorts_on_typs_aux ((a,~1),s) |
265 |
| sorts_on_typs (TVar (v,s)) = sorts_on_typs_aux (v,s); |
|
266 |
||
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24322
diff
changeset
|
267 |
(*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
|
268 |
fun type_literals_for_types Ts = |
adc11fb3f3aa
generate proper arity declarations for TFrees for SPASS's DFG format;
blanchet
parents:
36692
diff
changeset
|
269 |
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
|
270 |
|
1ffcf4802802
Literals aren't sorted any more. Output overloaded constants' type var instantiations.
mengj
parents:
19719
diff
changeset
|
271 |
(** make axiom and conjecture clauses. **) |
1ffcf4802802
Literals aren't sorted any more. Output overloaded constants' type var instantiations.
mengj
parents:
19719
diff
changeset
|
272 |
|
15347 | 273 |
(**** Isabelle arities ****) |
274 |
||
24310 | 275 |
datatype arLit = TConsLit of class * string * string list |
22643
bc3bb8e9594a
Improved and simplified the treatment of classrel/arity clauses
paulson
parents:
22383
diff
changeset
|
276 |
| TVarLit of class * string; |
24310 | 277 |
|
35865 | 278 |
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
|
279 |
ArityClause of {axiom_name: string, conclLit: arLit, premLits: arLit list} |
15347 | 280 |
|
281 |
||
18798 | 282 |
fun gen_TVars 0 = [] |
283 |
| gen_TVars n = ("T_" ^ Int.toString n) :: gen_TVars (n-1); |
|
15347 | 284 |
|
18411
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents:
18409
diff
changeset
|
285 |
fun pack_sort(_,[]) = [] |
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents:
18409
diff
changeset
|
286 |
| 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
|
287 |
| pack_sort(tvar, cls::srt) = (cls, tvar) :: pack_sort(tvar, srt); |
24310 | 288 |
|
18411
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
paulson
parents:
18409
diff
changeset
|
289 |
(*Arity of type constructor tcon :: (arg1,...,argN)res*) |
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
290 |
fun make_axiom_arity_clause (tcons, axiom_name, (cls,args)) = |
21560 | 291 |
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
|
292 |
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
|
293 |
in |
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
294 |
ArityClause {axiom_name = axiom_name, |
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
295 |
conclLit = TConsLit (cls, make_fixed_type_const tcons, tvars), |
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
296 |
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
|
297 |
end; |
15347 | 298 |
|
299 |
||
300 |
(**** Isabelle class relations ****) |
|
301 |
||
35865 | 302 |
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
|
303 |
ClassrelClause of {axiom_name: string, subclass: class, superclass: class} |
24310 | 304 |
|
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21254
diff
changeset
|
305 |
(*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
|
306 |
fun class_pairs _ [] _ = [] |
21432
625797c592b2
Optimized class_pairs for the common case of no subclasses
paulson
parents:
21416
diff
changeset
|
307 |
| class_pairs thy subs supers = |
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36170
diff
changeset
|
308 |
let |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36170
diff
changeset
|
309 |
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
|
310 |
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
|
311 |
fun add_supers sub = fold (add_super sub) supers |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36170
diff
changeset
|
312 |
in fold add_supers subs [] end |
15347 | 313 |
|
35865 | 314 |
fun make_classrel_clause (sub,super) = |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37500
diff
changeset
|
315 |
ClassrelClause {axiom_name = classrel_clause_prefix ^ ascii_of sub ^ "_" ^ |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37500
diff
changeset
|
316 |
ascii_of super, |
24310 | 317 |
subclass = make_type_class sub, |
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21254
diff
changeset
|
318 |
superclass = make_type_class super}; |
15347 | 319 |
|
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21254
diff
changeset
|
320 |
fun make_classrel_clauses thy subs supers = |
35865 | 321 |
map make_classrel_clause (class_pairs thy subs supers); |
18868 | 322 |
|
323 |
||
324 |
(** Isabelle arities **) |
|
17845
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
paulson
parents:
17775
diff
changeset
|
325 |
|
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
326 |
fun arity_clause _ _ (_, []) = [] |
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
327 |
| arity_clause seen n (tcons, ("HOL.type",_)::ars) = (*ignore*) |
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
328 |
arity_clause seen n (tcons,ars) |
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
329 |
| 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
|
330 |
if member (op =) seen class then (*multiple arities for the same tycon, class pair*) |
37572
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
331 |
make_axiom_arity_clause (tcons, lookup_const tcons ^ "_" ^ class ^ "_" ^ Int.toString n, ar) :: |
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
332 |
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
|
333 |
else |
37572
a899f9506f39
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents:
37570
diff
changeset
|
334 |
make_axiom_arity_clause (tcons, lookup_const tcons ^ "_" ^ class, ar) :: |
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
335 |
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
|
336 |
|
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
337 |
fun multi_arity_clause [] = [] |
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
338 |
| multi_arity_clause ((tcons, ars) :: tc_arlists) = |
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
339 |
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
|
340 |
|
22643
bc3bb8e9594a
Improved and simplified the treatment of classrel/arity clauses
paulson
parents:
22383
diff
changeset
|
341 |
(*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
|
342 |
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
|
343 |
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
|
344 |
let val alg = Sign.classes_of thy |
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36170
diff
changeset
|
345 |
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
|
346 |
fun add_class tycon class = |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36170
diff
changeset
|
347 |
cons (class, domain_sorts tycon class) |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36170
diff
changeset
|
348 |
handle Sorts.CLASS_ERROR _ => I |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36170
diff
changeset
|
349 |
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
|
350 |
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
|
351 |
|
22643
bc3bb8e9594a
Improved and simplified the treatment of classrel/arity clauses
paulson
parents:
22383
diff
changeset
|
352 |
(*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
|
353 |
fun iter_type_class_pairs _ _ [] = ([], []) |
22643
bc3bb8e9594a
Improved and simplified the treatment of classrel/arity clauses
paulson
parents:
22383
diff
changeset
|
354 |
| iter_type_class_pairs thy tycons classes = |
bc3bb8e9594a
Improved and simplified the treatment of classrel/arity clauses
paulson
parents:
22383
diff
changeset
|
355 |
let val cpairs = type_class_pairs thy tycons classes |
33040 | 356 |
val newclasses = union_all (union_all (union_all (map (map #2 o #2) cpairs))) |
357 |
|> subtract (op =) classes |> subtract (op =) HOLogic.typeS |
|
24310 | 358 |
val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses |
33042 | 359 |
in (union (op =) classes' classes, union (op =) cpairs' cpairs) end; |
24310 | 360 |
|
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37479
diff
changeset
|
361 |
fun make_arity_clauses thy tycons classes = |
24310 | 362 |
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
|
363 |
in (classes', multi_arity_clause cpairs) end; |
18863 | 364 |
|
15347 | 365 |
end; |