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