author | lcp |
Thu, 06 Apr 1995 10:55:06 +0200 | |
changeset 1000 | 0ad2b1da57ff |
parent 986 | c978bb4e9a55 |
child 1159 | 998a5c3451bf |
permissions | -rw-r--r-- |
19 | 1 |
(* Title: Pure/sign.ML |
0 | 2 |
ID: $Id$ |
251 | 3 |
Author: Lawrence C Paulson and Markus Wenzel |
0 | 4 |
|
251 | 5 |
The abstract type "sg" of signatures. |
0 | 6 |
*) |
7 |
||
19 | 8 |
signature SIGN = |
0 | 9 |
sig |
10 |
structure Symtab: SYMTAB |
|
11 |
structure Syntax: SYNTAX |
|
251 | 12 |
structure Type: TYPE |
143 | 13 |
sharing Symtab = Type.Symtab |
562 | 14 |
local open Type Syntax in |
251 | 15 |
type sg |
16 |
val rep_sg: sg -> |
|
17 |
{tsig: type_sig, |
|
18 |
const_tab: typ Symtab.table, |
|
386 | 19 |
syn: syntax, |
251 | 20 |
stamps: string ref list} |
21 |
val subsig: sg * sg -> bool |
|
22 |
val eq_sg: sg * sg -> bool |
|
386 | 23 |
val is_draft: sg -> bool |
24 |
val const_type: sg -> string -> typ option |
|
402 | 25 |
val classes: sg -> class list |
583 | 26 |
val subsort: sg -> sort * sort -> bool |
27 |
val norm_sort: sg -> sort -> sort |
|
251 | 28 |
val print_sg: sg -> unit |
562 | 29 |
val pretty_sg: sg -> Pretty.T |
251 | 30 |
val pprint_sg: sg -> pprint_args -> unit |
386 | 31 |
val pretty_term: sg -> term -> Pretty.T |
32 |
val pretty_typ: sg -> typ -> Pretty.T |
|
620 | 33 |
val pretty_sort: sort -> Pretty.T |
251 | 34 |
val string_of_term: sg -> term -> string |
35 |
val string_of_typ: sg -> typ -> string |
|
36 |
val pprint_term: sg -> term -> pprint_args -> unit |
|
37 |
val pprint_typ: sg -> typ -> pprint_args -> unit |
|
38 |
val certify_typ: sg -> typ -> typ |
|
39 |
val certify_term: sg -> term -> term * typ * int |
|
40 |
val read_typ: sg * (indexname -> sort option) -> string -> typ |
|
562 | 41 |
val infer_types: sg -> (indexname -> typ option) -> |
959 | 42 |
(indexname -> sort option) -> string list -> bool |
949
83c588d6fee9
Changed treatment of during type inference internally generated type
nipkow
parents:
926
diff
changeset
|
43 |
-> term list * typ -> int * term * (indexname * typ) list |
562 | 44 |
val add_classes: (class * class list) list -> sg -> sg |
421 | 45 |
val add_classrel: (class * class) list -> sg -> sg |
386 | 46 |
val add_defsort: sort -> sg -> sg |
47 |
val add_types: (string * int * mixfix) list -> sg -> sg |
|
48 |
val add_tyabbrs: (string * string list * string * mixfix) list -> sg -> sg |
|
49 |
val add_tyabbrs_i: (string * string list * typ * mixfix) list -> sg -> sg |
|
50 |
val add_arities: (string * sort list * sort) list -> sg -> sg |
|
51 |
val add_consts: (string * string * mixfix) list -> sg -> sg |
|
52 |
val add_consts_i: (string * typ * mixfix) list -> sg -> sg |
|
53 |
val add_syntax: (string * string * mixfix) list -> sg -> sg |
|
54 |
val add_syntax_i: (string * typ * mixfix) list -> sg -> sg |
|
55 |
val add_trfuns: |
|
56 |
(string * (ast list -> ast)) list * |
|
57 |
(string * (term list -> term)) list * |
|
58 |
(string * (term list -> term)) list * |
|
59 |
(string * (ast list -> ast)) list -> sg -> sg |
|
60 |
val add_trrules: xrule list -> sg -> sg |
|
61 |
val add_name: string -> sg -> sg |
|
62 |
val make_draft: sg -> sg |
|
251 | 63 |
val merge: sg * sg -> sg |
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
64 |
val proto_pure: sg |
251 | 65 |
val pure: sg |
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
66 |
val cpure: sg |
386 | 67 |
val const_of_class: class -> string |
68 |
val class_of_const: string -> class |
|
251 | 69 |
end |
0 | 70 |
end; |
71 |
||
251 | 72 |
functor SignFun(structure Syntax: SYNTAX and Type: TYPE): SIGN = |
143 | 73 |
struct |
0 | 74 |
|
75 |
structure Symtab = Type.Symtab; |
|
76 |
structure Syntax = Syntax; |
|
386 | 77 |
structure BasicSyntax: BASIC_SYNTAX = Syntax; |
251 | 78 |
structure Pretty = Syntax.Pretty; |
79 |
structure Type = Type; |
|
562 | 80 |
open BasicSyntax Type; |
0 | 81 |
|
143 | 82 |
|
251 | 83 |
(** datatype sg **) |
84 |
||
85 |
(*the "ref" in stamps ensures that no two signatures are identical -- it is |
|
86 |
impossible to forge a signature*) |
|
143 | 87 |
|
19 | 88 |
datatype sg = |
143 | 89 |
Sg of { |
251 | 90 |
tsig: type_sig, (*order-sorted signature of types*) |
143 | 91 |
const_tab: typ Symtab.table, (*types of constants*) |
92 |
syn: Syntax.syntax, (*syntax for parsing and printing*) |
|
93 |
stamps: string ref list}; (*unique theory indentifier*) |
|
0 | 94 |
|
95 |
fun rep_sg (Sg args) = args; |
|
402 | 96 |
val tsig_of = #tsig o rep_sg; |
0 | 97 |
|
206
0d624d1ba9cc
added subsig: sg * sg -> bool to test if one signature is contained in another.
nipkow
parents:
200
diff
changeset
|
98 |
|
402 | 99 |
(* stamps *) |
100 |
||
386 | 101 |
fun is_draft (Sg {stamps = ref "#" :: _, ...}) = true |
102 |
| is_draft _ = false; |
|
103 |
||
251 | 104 |
fun subsig (Sg {stamps = s1, ...}, Sg {stamps = s2, ...}) = s1 subset s2; |
0 | 105 |
|
266 | 106 |
fun eq_sg (Sg {stamps = s1, ...}, Sg {stamps = s2, ...}) = eq_set (s1, s2); |
0 | 107 |
|
108 |
||
402 | 109 |
(* consts *) |
110 |
||
386 | 111 |
fun const_type (Sg {const_tab, ...}) c = |
112 |
Symtab.lookup (const_tab, c); |
|
113 |
||
114 |
||
620 | 115 |
(* classes and sorts *) |
402 | 116 |
|
117 |
val classes = #classes o Type.rep_tsig o tsig_of; |
|
118 |
||
119 |
val subsort = Type.subsort o tsig_of; |
|
120 |
val norm_sort = Type.norm_sort o tsig_of; |
|
121 |
||
620 | 122 |
fun pretty_sort [c] = Pretty.str c |
123 |
| pretty_sort cs = Pretty.str_list "{" "}" cs; |
|
124 |
||
402 | 125 |
|
0 | 126 |
|
251 | 127 |
(** print signature **) |
128 |
||
386 | 129 |
val stamp_names = rev o map !; |
130 |
||
251 | 131 |
fun print_sg sg = |
132 |
let |
|
133 |
fun prt_typ syn ty = Pretty.quote (Syntax.pretty_typ syn ty); |
|
134 |
||
135 |
fun pretty_subclass (cl, cls) = Pretty.block |
|
136 |
[Pretty.str (cl ^ " <"), Pretty.brk 1, pretty_sort cls]; |
|
137 |
||
138 |
fun pretty_default cls = Pretty.block |
|
139 |
[Pretty.str "default:", Pretty.brk 1, pretty_sort cls]; |
|
140 |
||
963 | 141 |
fun pretty_ty (ty, n) = Pretty.str (ty ^ " " ^ string_of_int n); |
251 | 142 |
|
143 |
fun pretty_abbr syn (ty, (vs, rhs)) = Pretty.block |
|
620 | 144 |
[prt_typ syn (Type (ty, map (fn v => TVar ((v, 0), [])) vs)), |
251 | 145 |
Pretty.str " =", Pretty.brk 1, prt_typ syn rhs]; |
146 |
||
147 |
fun pretty_arity ty (cl, []) = Pretty.block |
|
148 |
[Pretty.str (ty ^ " ::"), Pretty.brk 1, Pretty.str cl] |
|
149 |
| pretty_arity ty (cl, srts) = Pretty.block |
|
150 |
[Pretty.str (ty ^ " ::"), Pretty.brk 1, |
|
151 |
Pretty.list "(" ")" (map pretty_sort srts), |
|
152 |
Pretty.brk 1, Pretty.str cl]; |
|
153 |
||
963 | 154 |
fun pretty_arities (ty, ars) = map (pretty_arity ty) ars; |
251 | 155 |
|
156 |
fun pretty_const syn (c, ty) = Pretty.block |
|
266 | 157 |
[Pretty.str (quote c ^ " ::"), Pretty.brk 1, prt_typ syn ty]; |
251 | 158 |
|
159 |
||
160 |
val Sg {tsig, const_tab, syn, stamps} = sg; |
|
963 | 161 |
val {classes, subclass, default, tycons, abbrs, arities} = rep_tsig tsig; |
251 | 162 |
in |
386 | 163 |
Pretty.writeln (Pretty.strs ("stamps:" :: stamp_names stamps)); |
251 | 164 |
Pretty.writeln (Pretty.strs ("classes:" :: classes)); |
963 | 165 |
Pretty.writeln (Pretty.big_list "subclass:" |
166 |
(map pretty_subclass subclass)); |
|
251 | 167 |
Pretty.writeln (pretty_default default); |
963 | 168 |
Pretty.writeln (Pretty.big_list "types:" (map pretty_ty tycons)); |
251 | 169 |
Pretty.writeln (Pretty.big_list "abbrs:" (map (pretty_abbr syn) abbrs)); |
963 | 170 |
Pretty.writeln (Pretty.big_list "arities:" |
171 |
(flat (map pretty_arities arities))); |
|
251 | 172 |
Pretty.writeln (Pretty.big_list "consts:" |
963 | 173 |
(map (pretty_const syn) (Symtab.dest const_tab))) |
251 | 174 |
end; |
175 |
||
176 |
||
562 | 177 |
fun pretty_sg (Sg {stamps, ...}) = |
178 |
Pretty.str_list "{" "}" (stamp_names stamps); |
|
179 |
||
180 |
val pprint_sg = Pretty.pprint o pretty_sg; |
|
251 | 181 |
|
182 |
||
183 |
||
184 |
(** pretty printing of terms and types **) |
|
185 |
||
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
186 |
fun pretty_term (Sg {syn, stamps, ...}) = |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
187 |
let val curried = "CPure" mem (map ! stamps); |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
188 |
in Syntax.pretty_term curried syn end; |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
189 |
|
251 | 190 |
fun pretty_typ (Sg {syn, ...}) = Syntax.pretty_typ syn; |
191 |
||
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
192 |
fun string_of_term (Sg {syn, stamps, ...}) = |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
193 |
let val curried = "CPure" mem (map ! stamps); |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
194 |
in Syntax.string_of_term curried syn end; |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
195 |
|
251 | 196 |
fun string_of_typ (Sg {syn, ...}) = Syntax.string_of_typ syn; |
197 |
||
198 |
fun pprint_term sg = Pretty.pprint o Pretty.quote o (pretty_term sg); |
|
199 |
fun pprint_typ sg = Pretty.pprint o Pretty.quote o (pretty_typ sg); |
|
0 | 200 |
|
201 |
||
169 | 202 |
|
562 | 203 |
(** read types **) (*exception ERROR*) |
204 |
||
205 |
fun err_in_type s = |
|
206 |
error ("The error(s) above occurred in type " ^ quote s); |
|
207 |
||
208 |
fun read_raw_typ syn tsig sort_of str = |
|
209 |
Syntax.read_typ syn (fn x => if_none (sort_of x) (defaultS tsig)) str |
|
210 |
handle ERROR => err_in_type str; |
|
211 |
||
212 |
(*read and certify typ wrt a signature*) |
|
213 |
fun read_typ (Sg {tsig, syn, ...}, sort_of) str = |
|
214 |
cert_typ tsig (read_raw_typ syn tsig sort_of str) |
|
215 |
handle TYPE (msg, _, _) => (writeln msg; err_in_type str); |
|
216 |
||
217 |
||
218 |
||
386 | 219 |
(** certify types and terms **) (*exception TYPE*) |
251 | 220 |
|
221 |
fun certify_typ (Sg {tsig, ...}) ty = cert_typ tsig ty; |
|
222 |
||
223 |
||
386 | 224 |
fun mapfilt_atoms f (Abs (_, _, t)) = mapfilt_atoms f t |
225 |
| mapfilt_atoms f (t $ u) = mapfilt_atoms f t @ mapfilt_atoms f u |
|
226 |
| mapfilt_atoms f a = (case f a of Some y => [y] | None => []); |
|
0 | 227 |
|
386 | 228 |
fun certify_term (sg as Sg {tsig, ...}) tm = |
251 | 229 |
let |
230 |
fun valid_const a T = |
|
386 | 231 |
(case const_type sg a of |
251 | 232 |
Some U => typ_instance (tsig, T, U) |
233 |
| _ => false); |
|
169 | 234 |
|
251 | 235 |
fun atom_err (Const (a, T)) = |
236 |
if valid_const a T then None |
|
237 |
else Some ("Illegal type for constant " ^ quote a ^ " :: " ^ |
|
238 |
quote (string_of_typ sg T)) |
|
239 |
| atom_err (Var ((x, i), _)) = |
|
240 |
if i < 0 then Some ("Negative index for Var " ^ quote x) else None |
|
241 |
| atom_err _ = None; |
|
242 |
||
243 |
val norm_tm = |
|
244 |
(case it_term_types (typ_errors tsig) (tm, []) of |
|
245 |
[] => map_term_types (norm_typ tsig) tm |
|
246 |
| errs => raise_type (cat_lines errs) [] [tm]); |
|
247 |
in |
|
386 | 248 |
(case mapfilt_atoms atom_err norm_tm of |
251 | 249 |
[] => (norm_tm, type_of norm_tm, maxidx_of_term norm_tm) |
250 |
| errs => raise_type (cat_lines errs) [] [norm_tm]) |
|
251 |
end; |
|
252 |
||
253 |
||
254 |
||
583 | 255 |
(** infer_types **) (*exception ERROR*) |
251 | 256 |
|
959 | 257 |
fun infer_types sg types sorts used freeze (ts, T) = |
251 | 258 |
let |
562 | 259 |
val Sg {tsig, ...} = sg; |
260 |
val show_typ = string_of_typ sg; |
|
261 |
val show_term = string_of_term sg; |
|
169 | 262 |
|
562 | 263 |
fun term_err [] = "" |
264 |
| term_err [t] = "\nInvolving this term:\n" ^ show_term t |
|
623 | 265 |
| term_err ts = |
266 |
"\nInvolving these terms:\n" ^ cat_lines (map show_term ts); |
|
562 | 267 |
|
679
a682bbf70dc6
Pure/sign/infer_types/exn_type_msg: new, for more uniform handling of
lcp
parents:
623
diff
changeset
|
268 |
fun exn_type_msg (msg, Ts, ts) = |
926 | 269 |
"\nType checking error: " ^ msg ^ "\n" ^ |
679
a682bbf70dc6
Pure/sign/infer_types/exn_type_msg: new, for more uniform handling of
lcp
parents:
623
diff
changeset
|
270 |
cat_lines (map show_typ Ts) ^ term_err ts ^ "\n"; |
a682bbf70dc6
Pure/sign/infer_types/exn_type_msg: new, for more uniform handling of
lcp
parents:
623
diff
changeset
|
271 |
|
562 | 272 |
val T' = certify_typ sg T |
949
83c588d6fee9
Changed treatment of during type inference internally generated type
nipkow
parents:
926
diff
changeset
|
273 |
handle TYPE arg => error (exn_type_msg arg); |
623 | 274 |
|
275 |
val ct = const_type sg; |
|
276 |
||
952
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
277 |
fun warn() = |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
278 |
if length ts > 1 andalso length ts <= !Syntax.ambiguity_level |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
279 |
then (*no warning shown yet*) |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
280 |
writeln "Warning: Currently parsed input \ |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
281 |
\produces more than one parse tree.\n\ |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
282 |
\For more information lower Syntax.ambiguity_level." |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
283 |
else (); |
623 | 284 |
|
952
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
285 |
datatype result = One of int * term * (indexname * typ) list |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
286 |
| Errs of (string * typ list * term list)list |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
287 |
| Ambigs of term list; |
623 | 288 |
|
952
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
289 |
fun process_term(res,(t,i)) = |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
290 |
let val (u,tye) = Type.infer_types(tsig,ct,types,sorts,used,freeze,T',t) |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
291 |
in case res of |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
292 |
One(_,t0,_) => Ambigs([u,t0]) |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
293 |
| Errs _ => One(i,u,tye) |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
294 |
| Ambigs(us) => Ambigs(u::us) |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
295 |
end |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
296 |
handle TYPE arg => (case res of Errs(errs) => Errs(arg::errs) |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
297 |
| _ => res); |
623 | 298 |
|
952
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
299 |
in case foldl process_term (Errs[], ts ~~ (0 upto (length ts - 1))) of |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
300 |
One(res) => |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
301 |
(if length ts > !Syntax.ambiguity_level |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
302 |
then writeln "Fortunately, only one parse tree is type correct.\n\ |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
303 |
\It helps (speed!) if you disambiguate your grammar or your input." |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
304 |
else (); |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
305 |
res) |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
306 |
| Errs(errs) => (warn(); error(cat_lines(map exn_type_msg errs))) |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
307 |
| Ambigs(us) => |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
308 |
(warn(); |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
309 |
let val old_show_brackets = !show_brackets |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
310 |
val _ = show_brackets := true; |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
311 |
val errs = cat_lines(map show_term us) |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
312 |
in show_brackets := old_show_brackets; |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
313 |
error("Error: More than one term is type correct:\n" ^ errs) |
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents:
949
diff
changeset
|
314 |
end) |
623 | 315 |
end; |
251 | 316 |
|
317 |
||
623 | 318 |
(** extend signature **) (*exception ERROR*) |
319 |
||
620 | 320 |
(** signature extension functions **) (*exception ERROR*) |
386 | 321 |
|
322 |
fun decls_of name_of mfixs = |
|
323 |
map (fn (x, y, mx) => (name_of x mx, y)) mfixs; |
|
324 |
||
325 |
||
326 |
(* add default sort *) |
|
327 |
||
328 |
fun ext_defsort (syn, tsig, ctab) defsort = |
|
329 |
(syn, ext_tsig_defsort tsig defsort, ctab); |
|
330 |
||
331 |
||
332 |
(* add type constructors *) |
|
333 |
||
334 |
fun ext_types (syn, tsig, ctab) types = |
|
335 |
(Syntax.extend_type_gram syn types, |
|
562 | 336 |
ext_tsig_types tsig (decls_of Syntax.type_name types), |
386 | 337 |
ctab); |
338 |
||
339 |
||
340 |
(* add type abbreviations *) |
|
341 |
||
342 |
fun read_abbr syn tsig (t, vs, rhs_src) = |
|
343 |
(t, vs, read_raw_typ syn tsig (K None) rhs_src) |
|
344 |
handle ERROR => error ("in type abbreviation " ^ t); |
|
345 |
||
346 |
fun ext_abbrs rd_abbr (syn, tsig, ctab) abbrs = |
|
347 |
let |
|
348 |
fun mfix_of (t, vs, _, mx) = (t, length vs, mx); |
|
349 |
val syn1 = Syntax.extend_type_gram syn (map mfix_of abbrs); |
|
350 |
||
351 |
fun decl_of (t, vs, rhs, mx) = |
|
562 | 352 |
rd_abbr syn1 tsig (Syntax.type_name t mx, vs, rhs); |
386 | 353 |
in |
354 |
(syn1, ext_tsig_abbrs tsig (map decl_of abbrs), ctab) |
|
355 |
end; |
|
356 |
||
357 |
val ext_tyabbrs_i = ext_abbrs (K (K I)); |
|
358 |
val ext_tyabbrs = ext_abbrs read_abbr; |
|
359 |
||
360 |
||
361 |
(* add type arities *) |
|
362 |
||
363 |
fun ext_arities (syn, tsig, ctab) arities = |
|
364 |
let |
|
365 |
val tsig1 = ext_tsig_arities tsig arities; |
|
366 |
val log_types = logical_types tsig1; |
|
367 |
in |
|
368 |
(Syntax.extend_log_types syn log_types, tsig1, ctab) |
|
369 |
end; |
|
370 |
||
371 |
||
372 |
(* add term constants and syntax *) |
|
373 |
||
374 |
fun err_in_const c = |
|
375 |
error ("in declaration of constant " ^ quote c); |
|
376 |
||
377 |
fun err_dup_consts cs = |
|
378 |
error ("Duplicate declaration of constant(s) " ^ commas_quote cs); |
|
379 |
||
251 | 380 |
|
386 | 381 |
fun read_const syn tsig (c, ty_src, mx) = |
382 |
(c, read_raw_typ syn tsig (K None) ty_src, mx) |
|
562 | 383 |
handle ERROR => err_in_const (Syntax.const_name c mx); |
386 | 384 |
|
385 |
fun ext_cnsts rd_const syn_only (syn, tsig, ctab) raw_consts = |
|
386 |
let |
|
620 | 387 |
fun prep_const (c, ty, mx) = (c, varifyT (cert_typ tsig (no_tvars ty)), mx) |
562 | 388 |
handle TYPE (msg, _, _) => (writeln msg; err_in_const (Syntax.const_name c mx)); |
386 | 389 |
|
390 |
val consts = map (prep_const o rd_const syn tsig) raw_consts; |
|
391 |
val decls = |
|
392 |
if syn_only then [] |
|
562 | 393 |
else filter_out (equal "" o fst) (decls_of Syntax.const_name consts); |
386 | 394 |
in |
395 |
(Syntax.extend_const_gram syn consts, tsig, |
|
396 |
Symtab.extend_new (ctab, decls) |
|
397 |
handle Symtab.DUPS cs => err_dup_consts cs) |
|
398 |
end; |
|
399 |
||
400 |
val ext_consts_i = ext_cnsts (K (K I)) false; |
|
401 |
val ext_consts = ext_cnsts read_const false; |
|
402 |
val ext_syntax_i = ext_cnsts (K (K I)) true; |
|
403 |
val ext_syntax = ext_cnsts read_const true; |
|
404 |
||
405 |
||
406 |
(* add type classes *) |
|
407 |
||
408 |
fun const_of_class c = c ^ "_class"; |
|
409 |
||
410 |
fun class_of_const c_class = |
|
411 |
let |
|
412 |
val c = implode (take (size c_class - 6, explode c_class)); |
|
413 |
in |
|
414 |
if const_of_class c = c_class then c |
|
415 |
else raise_term ("class_of_const: bad name " ^ quote c_class) [] |
|
416 |
end; |
|
417 |
||
418 |
||
419 |
fun ext_classes (syn, tsig, ctab) classes = |
|
420 |
let |
|
562 | 421 |
val names = map fst classes; |
386 | 422 |
val consts = |
423 |
map (fn c => (const_of_class c, a_itselfT --> propT, NoSyn)) names; |
|
424 |
in |
|
425 |
ext_consts_i |
|
426 |
(Syntax.extend_consts syn names, ext_tsig_classes tsig classes, ctab) |
|
427 |
consts |
|
428 |
end; |
|
429 |
||
430 |
||
421 | 431 |
(* add to subclass relation *) |
432 |
||
433 |
fun ext_classrel (syn, tsig, ctab) pairs = |
|
434 |
(syn, ext_tsig_subclass tsig pairs, ctab); |
|
435 |
||
436 |
||
386 | 437 |
(* add syntactic translations *) |
438 |
||
439 |
fun ext_trfuns (syn, tsig, ctab) trfuns = |
|
440 |
(Syntax.extend_trfuns syn trfuns, tsig, ctab); |
|
441 |
||
442 |
fun ext_trrules (syn, tsig, ctab) xrules = |
|
865 | 443 |
(Syntax.extend_trrules syn xrules, tsig, ctab); |
386 | 444 |
|
445 |
||
446 |
(* build signature *) |
|
447 |
||
448 |
fun ext_stamps stamps name = |
|
449 |
let |
|
450 |
val stmps = (case stamps of ref "#" :: ss => ss | ss => ss); |
|
451 |
in |
|
452 |
if exists (equal name o !) stmps then |
|
453 |
error ("Theory already contains a " ^ quote name ^ " component") |
|
402 | 454 |
else ref name :: stmps |
386 | 455 |
end; |
456 |
||
457 |
fun make_sign (syn, tsig, ctab) stamps name = |
|
458 |
Sg {tsig = tsig, const_tab = ctab, syn = syn, |
|
459 |
stamps = ext_stamps stamps name}; |
|
460 |
||
461 |
fun extend_sign extfun name decls (Sg {tsig, const_tab, syn, stamps}) = |
|
462 |
make_sign (extfun (syn, tsig, const_tab) decls) stamps name; |
|
463 |
||
464 |
||
465 |
(* the external interfaces *) |
|
466 |
||
467 |
val add_classes = extend_sign ext_classes "#"; |
|
421 | 468 |
val add_classrel = extend_sign ext_classrel "#"; |
386 | 469 |
val add_defsort = extend_sign ext_defsort "#"; |
470 |
val add_types = extend_sign ext_types "#"; |
|
471 |
val add_tyabbrs = extend_sign ext_tyabbrs "#"; |
|
472 |
val add_tyabbrs_i = extend_sign ext_tyabbrs_i "#"; |
|
473 |
val add_arities = extend_sign ext_arities "#"; |
|
474 |
val add_consts = extend_sign ext_consts "#"; |
|
475 |
val add_consts_i = extend_sign ext_consts_i "#"; |
|
476 |
val add_syntax = extend_sign ext_syntax "#"; |
|
477 |
val add_syntax_i = extend_sign ext_syntax_i "#"; |
|
478 |
val add_trfuns = extend_sign ext_trfuns "#"; |
|
479 |
val add_trrules = extend_sign ext_trrules "#"; |
|
480 |
||
481 |
fun add_name name sg = extend_sign K name () sg; |
|
482 |
val make_draft = add_name "#"; |
|
483 |
||
484 |
||
485 |
||
620 | 486 |
(** merge signatures **) (*exception TERM*) (*exception ERROR (* FIXME *)*) |
143 | 487 |
|
251 | 488 |
fun merge (sg1, sg2) = |
489 |
if subsig (sg2, sg1) then sg1 |
|
490 |
else if subsig (sg1, sg2) then sg2 |
|
386 | 491 |
else if is_draft sg1 orelse is_draft sg2 then |
492 |
raise_term "Illegal merge of draft signatures" [] |
|
251 | 493 |
else |
494 |
(*neither is union already; must form union*) |
|
495 |
let |
|
496 |
val Sg {tsig = tsig1, const_tab = const_tab1, syn = syn1, |
|
497 |
stamps = stamps1} = sg1; |
|
498 |
val Sg {tsig = tsig2, const_tab = const_tab2, syn = syn2, |
|
499 |
stamps = stamps2} = sg2; |
|
500 |
||
386 | 501 |
|
402 | 502 |
val stamps = merge_rev_lists stamps1 stamps2; |
503 |
val _ = |
|
504 |
(case duplicates (stamp_names stamps) of |
|
505 |
[] => () |
|
506 |
| dups => raise_term ("Attempt to merge different versions of theories " |
|
507 |
^ commas_quote dups) []); |
|
508 |
||
251 | 509 |
val tsig = merge_tsigs (tsig1, tsig2); |
510 |
||
511 |
val const_tab = Symtab.merge (op =) (const_tab1, const_tab2) |
|
386 | 512 |
handle Symtab.DUPS cs => |
513 |
raise_term ("Incompatible types for constant(s) " ^ commas_quote cs) []; |
|
514 |
||
515 |
val syn = Syntax.merge_syntaxes syn1 syn2; |
|
251 | 516 |
in |
517 |
Sg {tsig = tsig, const_tab = const_tab, syn = syn, stamps = stamps} |
|
518 |
end; |
|
143 | 519 |
|
0 | 520 |
|
521 |
||
251 | 522 |
(** the Pure signature **) |
0 | 523 |
|
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
524 |
val proto_pure = |
763 | 525 |
make_sign (Syntax.pure_syn, tsig0, Symtab.null) [] "#" |
410 | 526 |
|> add_types |
386 | 527 |
(("fun", 2, NoSyn) :: |
528 |
("prop", 0, NoSyn) :: |
|
529 |
("itself", 1, NoSyn) :: |
|
562 | 530 |
Syntax.pure_types) |
531 |
|> add_classes [(logicC, [])] |
|
410 | 532 |
|> add_defsort logicS |
533 |
|> add_arities |
|
386 | 534 |
[("fun", [logicS, logicS], logicS), |
535 |
("prop", [], logicS), |
|
536 |
("itself", [logicS], logicS)] |
|
562 | 537 |
|> add_syntax Syntax.pure_syntax |
410 | 538 |
|> add_trfuns Syntax.pure_trfuns |
539 |
|> add_consts |
|
386 | 540 |
[("==", "['a::{}, 'a] => prop", Mixfix ("(_ ==/ _)", [3, 2], 2)), |
541 |
("=?=", "['a::{}, 'a] => prop", Mixfix ("(_ =?=/ _)", [3, 2], 2)), |
|
542 |
("==>", "[prop, prop] => prop", Mixfix ("(_ ==>/ _)", [2, 1], 1)), |
|
865 | 543 |
("all", "('a => prop) => prop", Binder ("!!", 0, 0)), |
573 | 544 |
("TYPE", "'a itself", NoSyn)] |
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
545 |
|> add_name "ProtoPure"; |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
546 |
|
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
547 |
val pure = proto_pure |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
548 |
|> add_syntax |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
549 |
[("_appl", "[('b => 'a), args] => logic", Mixfix ("(1_/(1'(_')))", |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
550 |
[max_pri, 0], max_pri)), |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
551 |
("_appl", "[('b => 'a), args] => aprop", Mixfix ("(1_/(1'(_')))", |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
552 |
[max_pri, 0], max_pri))] |
410 | 553 |
|> add_name "Pure"; |
0 | 554 |
|
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
555 |
val cpure = proto_pure |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
556 |
|> add_syntax |
986 | 557 |
[("_applC", "[('b => 'a), 'c] => logic", Mixfix ("_/ _", |
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
558 |
[max_pri-1, max_pri], |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
559 |
max_pri-1)), |
986 | 560 |
("_applC", "[('b => 'a), 'c] => aprop", Mixfix ("_/ _", |
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
561 |
[max_pri-1, max_pri], |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
562 |
max_pri-1))] |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
901
diff
changeset
|
563 |
|> add_name "CPure"; |
0 | 564 |
|
565 |
end; |