author | nipkow |
Tue, 14 Dec 1993 14:02:52 +0100 | |
changeset 197 | 7c7179e687b2 |
parent 171 | ab0f93a291b5 |
child 200 | 39a931cc6558 |
permissions | -rw-r--r-- |
19 | 1 |
(* Title: Pure/sign.ML |
0 | 2 |
ID: $Id$ |
19 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
143 | 6 |
The abstract types "sg" (signatures) and "cterm" / "ctyp" (certified terms / |
7 |
typs under a signature). |
|
0 | 8 |
*) |
9 |
||
19 | 10 |
signature SIGN = |
0 | 11 |
sig |
12 |
structure Type: TYPE |
|
13 |
structure Symtab: SYMTAB |
|
14 |
structure Syntax: SYNTAX |
|
143 | 15 |
sharing Symtab = Type.Symtab |
0 | 16 |
type sg |
17 |
type cterm |
|
18 |
type ctyp |
|
19 |
val cfun: (term -> term) -> (cterm -> cterm) |
|
20 |
val cterm_of: sg -> term -> cterm |
|
21 |
val ctyp_of: sg -> typ -> ctyp |
|
22 |
val extend: sg -> string -> |
|
19 | 23 |
(class * class list) list * class list * |
24 |
(string list * int) list * |
|
25 |
(string list * (sort list * class)) list * |
|
26 |
(string list * string)list * Syntax.sext option -> sg |
|
0 | 27 |
val merge: sg * sg -> sg |
28 |
val pure: sg |
|
29 |
val read_cterm: sg -> string * typ -> cterm |
|
30 |
val read_ctyp: sg -> string -> ctyp |
|
31 |
val read_insts: sg -> (indexname -> typ option) * (indexname -> sort option) |
|
19 | 32 |
-> (indexname -> typ option) * (indexname -> sort option) |
33 |
-> (string*string)list |
|
34 |
-> (indexname*ctyp)list * (cterm*cterm)list |
|
0 | 35 |
val read_typ: sg * (indexname -> sort option) -> string -> typ |
36 |
val rep_cterm: cterm -> {T: typ, t: term, sign: sg, maxidx: int} |
|
37 |
val rep_ctyp: ctyp -> {T: typ, sign: sg} |
|
38 |
val rep_sg: sg -> {tsig: Type.type_sig, |
|
19 | 39 |
const_tab: typ Symtab.table, |
40 |
syn: Syntax.syntax, |
|
41 |
stamps: string ref list} |
|
0 | 42 |
val string_of_cterm: cterm -> string |
43 |
val string_of_ctyp: ctyp -> string |
|
44 |
val pprint_cterm: cterm -> pprint_args -> unit |
|
45 |
val pprint_ctyp: ctyp -> pprint_args -> unit |
|
46 |
val string_of_term: sg -> term -> string |
|
47 |
val string_of_typ: sg -> typ -> string |
|
48 |
val pprint_term: sg -> term -> pprint_args -> unit |
|
49 |
val pprint_typ: sg -> typ -> pprint_args -> unit |
|
50 |
val term_of: cterm -> term |
|
51 |
val typ_of: ctyp -> typ |
|
52 |
val pretty_term: sg -> term -> Syntax.Pretty.T |
|
53 |
end; |
|
54 |
||
143 | 55 |
functor SignFun(structure Type: TYPE and Syntax: SYNTAX): SIGN = |
56 |
struct |
|
0 | 57 |
|
58 |
structure Type = Type; |
|
59 |
structure Symtab = Type.Symtab; |
|
60 |
structure Syntax = Syntax; |
|
61 |
structure Pretty = Syntax.Pretty |
|
62 |
||
143 | 63 |
|
64 |
(* Signatures of theories. *) |
|
65 |
||
19 | 66 |
datatype sg = |
143 | 67 |
Sg of { |
68 |
tsig: Type.type_sig, (*order-sorted signature of types*) |
|
69 |
const_tab: typ Symtab.table, (*types of constants*) |
|
70 |
syn: Syntax.syntax, (*syntax for parsing and printing*) |
|
71 |
stamps: string ref list}; (*unique theory indentifier*) |
|
0 | 72 |
|
73 |
||
74 |
fun rep_sg (Sg args) = args; |
|
75 |
||
76 |
fun string_of_typ(Sg{tsig,syn,...}) = Syntax.string_of_typ syn; |
|
77 |
||
78 |
fun pprint_typ(Sg{syn,...}) = Pretty.pprint o Pretty.quote o (Syntax.pretty_typ syn); |
|
79 |
||
80 |
(*Is constant present in table with more generic type?*) |
|
81 |
fun valid_const tsig ctab (a,T) = case Symtab.lookup(ctab, a) of |
|
19 | 82 |
Some U => Type.typ_instance(tsig,T,U) | _ => false; |
0 | 83 |
|
84 |
||
85 |
(*Check a term for errors. Are all constants and types valid in signature? |
|
86 |
Does not check that term is well-typed!*) |
|
19 | 87 |
fun term_errors (sign as Sg{tsig,const_tab,...}) = |
0 | 88 |
let val showtyp = string_of_typ sign; |
89 |
fun terrs (Const (a,T), errs) = |
|
19 | 90 |
if valid_const tsig const_tab (a,T) |
91 |
then Type.type_errors (tsig,showtyp) (T,errs) |
|
92 |
else ("Illegal type for constant: " ^ a ^ ": " ^ showtyp T) :: errs |
|
0 | 93 |
| terrs (Free (_,T), errs) = Type.type_errors (tsig,showtyp) (T,errs) |
94 |
| terrs (Var ((a,i),T), errs) = |
|
19 | 95 |
if i>=0 then Type.type_errors (tsig,showtyp) (T,errs) |
96 |
else ("Negative index for Var: " ^ a) :: errs |
|
0 | 97 |
| terrs (Bound _, errs) = errs (*loose bvars detected by type_of*) |
19 | 98 |
| terrs (Abs (_,T,t), errs) = |
99 |
Type.type_errors(tsig,showtyp)(T,terrs(t,errs)) |
|
0 | 100 |
| terrs (f$t, errs) = terrs(f, terrs (t,errs)) |
101 |
in terrs end; |
|
102 |
||
103 |
||
169 | 104 |
|
0 | 105 |
(** The Extend operation **) |
106 |
||
169 | 107 |
(* Extend a signature: may add classes, types and constants. The "ref" in |
108 |
stamps ensures that no two signatures are identical -- it is impossible to |
|
109 |
forge a signature. *) |
|
110 |
||
111 |
fun extend (Sg {tsig, const_tab, syn, stamps}) name |
|
112 |
(classes, default, types, arities, const_decs, opt_sext) = |
|
113 |
let |
|
114 |
fun err_in_typ s = error ("The error(s) above occurred in type " ^ quote s); |
|
115 |
||
116 |
fun read_typ tsg sy s = |
|
117 |
Syntax.read_typ sy (K (Type.defaultS tsg)) s handle ERROR => err_in_typ s; |
|
118 |
||
119 |
fun check_typ tsg sy ty = |
|
120 |
(case Type.type_errors (tsg, Syntax.string_of_typ sy) (ty, []) of |
|
121 |
[] => ty |
|
122 |
| errs => (prs (cat_lines errs); err_in_typ (Syntax.string_of_typ sy ty))); |
|
123 |
||
124 |
(*reset TVar indices to zero, renaming to preserve distinctness*) |
|
125 |
fun zero_tvar_indices T = |
|
126 |
let |
|
127 |
val inxSs = typ_tvars T; |
|
128 |
val nms' = variantlist (map (#1 o #1) inxSs, []); |
|
129 |
val tye = map (fn ((v, S), a) => (v, TVar ((a, 0), S))) (inxSs ~~ nms'); |
|
130 |
in typ_subst_TVars tye T end; |
|
131 |
||
132 |
(*read and check the type mentioned in a const declaration; zero type var |
|
133 |
indices because type inference requires it*) |
|
134 |
||
135 |
fun read_consts tsg sy (cs, s) = |
|
136 |
let val ty = zero_tvar_indices (Type.varifyT (read_typ tsg sy s)); |
|
137 |
in |
|
138 |
(case Type.type_errors (tsg, Syntax.string_of_typ sy) (ty, []) of |
|
139 |
[] => (cs, ty) |
|
140 |
| errs => error (cat_lines (("Error in type of constants " ^ |
|
141 |
space_implode " " (map quote cs)) :: errs))) |
|
142 |
end; |
|
0 | 143 |
|
143 | 144 |
|
145 |
(* FIXME abbr *) |
|
155 | 146 |
val tsig' = Type.extend (tsig, classes, default, types, arities); |
143 | 147 |
|
169 | 148 |
(* FIXME *) |
149 |
fun expand_typ _ ty = ty; |
|
143 | 150 |
|
169 | 151 |
val read_ty = |
152 |
(expand_typ tsig') o (check_typ tsig' syn) o (read_typ tsig' syn); |
|
153 |
val log_types = Type.logical_types tsig'; |
|
155 | 154 |
val xconsts = map #1 classes @ flat (map #1 types) @ flat (map #1 const_decs); |
169 | 155 |
val sext = case opt_sext of Some sx => sx | None => Syntax.empty_sext; |
156 |
||
157 |
val syn' = Syntax.extend syn read_ty (log_types, xconsts, sext); |
|
158 |
||
159 |
val const_decs' = |
|
160 |
map (read_consts tsig' syn') (Syntax.constants sext @ const_decs); |
|
143 | 161 |
in |
162 |
Sg { |
|
163 |
tsig = tsig', |
|
164 |
const_tab = Symtab.st_of_declist (const_decs', const_tab) |
|
165 |
handle Symtab.DUPLICATE a => error ("Constant " ^ quote a ^ " declared twice"), |
|
166 |
syn = syn', |
|
169 | 167 |
stamps = ref name :: stamps} |
143 | 168 |
end; |
0 | 169 |
|
170 |
||
171 |
(* The empty signature *) |
|
143 | 172 |
|
169 | 173 |
val sg0 = Sg {tsig = Type.tsig0, const_tab = Symtab.null, |
174 |
syn = Syntax.type_syn, stamps = []}; |
|
0 | 175 |
|
143 | 176 |
|
177 |
(* The pure signature *) |
|
178 |
||
179 |
val pure = extend sg0 "Pure" |
|
0 | 180 |
([("logic", [])], |
181 |
["logic"], |
|
143 | 182 |
[(["fun"], 2), |
183 |
(["prop"], 0), |
|
184 |
(Syntax.syntax_types, 0)], |
|
0 | 185 |
[(["fun"], ([["logic"], ["logic"]], "logic")), |
186 |
(["prop"], ([], "logic"))], |
|
171 | 187 |
[([Syntax.constrainC], "'a::logic => 'a")], |
143 | 188 |
Some Syntax.pure_sext); |
189 |
||
0 | 190 |
|
191 |
||
192 |
(** The Merge operation **) |
|
193 |
||
194 |
(*Update table with (a,x) providing any existing asgt to "a" equals x. *) |
|
195 |
fun update_eq ((a,x),tab) = |
|
196 |
case Symtab.lookup(tab,a) of |
|
19 | 197 |
None => Symtab.update((a,x), tab) |
198 |
| Some y => if x=y then tab |
|
199 |
else raise TERM ("Incompatible types for constant: "^a, []); |
|
0 | 200 |
|
201 |
(*Combine tables, updating tab2 by tab1 and checking.*) |
|
19 | 202 |
fun merge_tabs (tab1,tab2) = |
0 | 203 |
Symtab.balance (foldr update_eq (Symtab.alist_of tab1, tab2)); |
204 |
||
205 |
(*Combine tables, overwriting tab2 with tab1.*) |
|
19 | 206 |
fun smash_tabs (tab1,tab2) = |
0 | 207 |
Symtab.balance (foldr Symtab.update (Symtab.alist_of tab1, tab2)); |
208 |
||
209 |
(*Combine stamps, checking that theory names are disjoint. *) |
|
19 | 210 |
fun merge_stamps (stamps1,stamps2) = |
0 | 211 |
let val stamps = stamps1 union stamps2 in |
212 |
case findrep (map ! stamps) of |
|
213 |
a::_ => error ("Attempt to merge different versions of theory: " ^ a) |
|
214 |
| [] => stamps |
|
215 |
end; |
|
216 |
||
217 |
(*Merge two signatures. Forms unions of tables. Prefers sign1. *) |
|
169 | 218 |
fun merge |
219 |
(sign1 as Sg {tsig = tsig1, const_tab = ctab1, stamps = stamps1, syn = syn1}, |
|
220 |
sign2 as Sg {tsig = tsig2, const_tab = ctab2, stamps = stamps2, syn = syn2}) = |
|
0 | 221 |
if stamps2 subset stamps1 then sign1 |
222 |
else if stamps1 subset stamps2 then sign2 |
|
169 | 223 |
else (*neither is union already; must form union*) |
224 |
let val tsig = Type.merge (tsig1, tsig2); |
|
225 |
in |
|
226 |
Sg {tsig = tsig, const_tab = merge_tabs (ctab1, ctab2), |
|
227 |
stamps = merge_stamps (stamps1, stamps2), |
|
228 |
syn = Syntax.merge (Type.logical_types tsig) syn1 syn2} |
|
229 |
end; |
|
230 |
||
0 | 231 |
|
232 |
||
233 |
(**** CERTIFIED TYPES ****) |
|
234 |
||
235 |
||
236 |
(*Certified typs under a signature*) |
|
237 |
datatype ctyp = Ctyp of {sign: sg, T: typ}; |
|
238 |
||
239 |
fun rep_ctyp(Ctyp ctyp) = ctyp; |
|
240 |
||
241 |
fun typ_of (Ctyp{sign,T}) = T; |
|
242 |
||
243 |
fun ctyp_of (sign as Sg{tsig,...}) T = |
|
19 | 244 |
case Type.type_errors (tsig,string_of_typ sign) (T,[]) of |
245 |
[] => Ctyp{sign= sign,T= T} |
|
246 |
| errs => error (cat_lines ("Error in type:" :: errs)); |
|
0 | 247 |
|
248 |
(*The only use is a horrible hack in the simplifier!*) |
|
249 |
fun read_typ(Sg{tsig,syn,...}, defS) s = |
|
250 |
let val term = Syntax.read syn Syntax.typeT s; |
|
19 | 251 |
val S0 = Type.defaultS tsig; |
252 |
fun defS0 s = case defS s of Some S => S | None => S0; |
|
0 | 253 |
in Syntax.typ_of_term defS0 term end; |
254 |
||
255 |
fun read_ctyp sign = ctyp_of sign o read_typ(sign, K None); |
|
256 |
||
257 |
fun string_of_ctyp (Ctyp{sign,T}) = string_of_typ sign T; |
|
258 |
||
259 |
fun pprint_ctyp (Ctyp{sign,T}) = pprint_typ sign T; |
|
260 |
||
261 |
||
262 |
(**** CERTIFIED TERMS ****) |
|
263 |
||
264 |
(*Certified terms under a signature, with checked typ and maxidx of Vars*) |
|
265 |
datatype cterm = Cterm of {sign: sg, t: term, T: typ, maxidx: int}; |
|
266 |
||
267 |
fun rep_cterm (Cterm args) = args; |
|
268 |
||
269 |
(*Return the underlying term*) |
|
270 |
fun term_of (Cterm{sign,t,T,maxidx}) = t; |
|
271 |
||
272 |
(** pretty printing of terms **) |
|
273 |
||
274 |
fun pretty_term (Sg{tsig,syn,...}) = Syntax.pretty_term syn; |
|
275 |
||
276 |
fun string_of_term sign t = Pretty.string_of (pretty_term sign t); |
|
277 |
||
278 |
fun pprint_term sign = Pretty.pprint o Pretty.quote o (pretty_term sign); |
|
279 |
||
280 |
fun string_of_cterm (Cterm{sign,t,...}) = string_of_term sign t; |
|
281 |
||
282 |
fun pprint_cterm (Cterm{sign,t,...}) = pprint_term sign t; |
|
283 |
||
284 |
(*Create a cterm by checking a "raw" term with respect to a signature*) |
|
285 |
fun cterm_of sign t = |
|
286 |
case term_errors sign (t,[]) of |
|
287 |
[] => Cterm{sign=sign, t=t, T= type_of t, maxidx= maxidx_of_term t} |
|
288 |
| errs => raise TERM(cat_lines("Term not in signature"::errs), [t]); |
|
289 |
||
290 |
fun cfun f = fn Cterm{sign,t,...} => cterm_of sign (f t); |
|
291 |
||
292 |
(*Lexing, parsing, polymorphic typechecking of a term.*) |
|
293 |
fun read_def_cterm (sign as Sg{tsig, const_tab, syn,...}, types, sorts) |
|
19 | 294 |
(a,T) = |
0 | 295 |
let val showtyp = string_of_typ sign |
296 |
and showterm = string_of_term sign |
|
297 |
fun termerr [] = "" |
|
19 | 298 |
| termerr [t] = "\nInvolving this term:\n" ^ showterm t ^ "\n" |
299 |
| termerr ts = "\nInvolving these terms:\n" ^ |
|
300 |
cat_lines (map showterm ts) |
|
0 | 301 |
val t = Syntax.read syn T a; |
302 |
val (t',tye) = Type.infer_types (tsig, const_tab, types, |
|
19 | 303 |
sorts, showtyp, T, t) |
304 |
handle TYPE (msg, Ts, ts) => |
|
305 |
error ("Type checking error: " ^ msg ^ "\n" ^ |
|
306 |
cat_lines (map showtyp Ts) ^ termerr ts) |
|
0 | 307 |
in (cterm_of sign t', tye) |
308 |
end |
|
309 |
handle TERM (msg, _) => error ("Error: " ^ msg); |
|
310 |
||
311 |
||
312 |
fun read_cterm sign = #1 o (read_def_cterm (sign, K None, K None)); |
|
313 |
||
314 |
(** reading of instantiations **) |
|
315 |
||
316 |
fun indexname cs = case Syntax.scan_varname cs of (v,[]) => v |
|
19 | 317 |
| _ => error("Lexical error in variable name " ^ quote (implode cs)); |
0 | 318 |
|
319 |
fun absent ixn = |
|
320 |
error("No such variable in term: " ^ Syntax.string_of_vname ixn); |
|
321 |
||
322 |
fun inst_failure ixn = |
|
323 |
error("Instantiation of " ^ Syntax.string_of_vname ixn ^ " fails"); |
|
324 |
||
325 |
fun read_insts (sign as Sg{tsig,...}) (rtypes,rsorts) (types,sorts) insts = |
|
326 |
let fun split([],tvs,vs) = (tvs,vs) |
|
327 |
| split((sv,st)::l,tvs,vs) = (case explode sv of |
|
19 | 328 |
"'"::cs => split(l,(indexname cs,st)::tvs,vs) |
329 |
| cs => split(l,tvs,(indexname cs,st)::vs)); |
|
0 | 330 |
val (tvs,vs) = split(insts,[],[]); |
331 |
fun readT((a,i),st) = |
|
19 | 332 |
let val ixn = ("'" ^ a,i); |
333 |
val S = case rsorts ixn of Some S => S | None => absent ixn; |
|
334 |
val T = read_typ (sign,sorts) st; |
|
335 |
in if Type.typ_instance(tsig,T,TVar(ixn,S)) then (ixn,T) |
|
336 |
else inst_failure ixn |
|
337 |
end |
|
0 | 338 |
val tye = map readT tvs; |
339 |
fun add_cterm ((cts,tye), (ixn,st)) = |
|
19 | 340 |
let val T = case rtypes ixn of |
341 |
Some T => typ_subst_TVars tye T |
|
342 |
| None => absent ixn; |
|
343 |
val (ct,tye2) = read_def_cterm (sign,types,sorts) (st,T); |
|
197
7c7179e687b2
Updated read_insts to approximate simultaneous type checking of substitution
nipkow
parents:
171
diff
changeset
|
344 |
in ((ixn,T,ct)::cts,tye2 @ tye) end |
0 | 345 |
val (cterms,tye') = foldl add_cterm (([],tye), vs); |
197
7c7179e687b2
Updated read_insts to approximate simultaneous type checking of substitution
nipkow
parents:
171
diff
changeset
|
346 |
in (map (fn (ixn,T) => (ixn,ctyp_of sign T)) tye', |
7c7179e687b2
Updated read_insts to approximate simultaneous type checking of substitution
nipkow
parents:
171
diff
changeset
|
347 |
map (fn (ixn,T,ct)=> (cterm_of sign (Var(ixn,typ_subst_TVars tye' T)), ct)) |
7c7179e687b2
Updated read_insts to approximate simultaneous type checking of substitution
nipkow
parents:
171
diff
changeset
|
348 |
cterms) |
7c7179e687b2
Updated read_insts to approximate simultaneous type checking of substitution
nipkow
parents:
171
diff
changeset
|
349 |
end; |
0 | 350 |
|
351 |
end; |
|
143 | 352 |