| author | wenzelm |
| Mon, 01 May 2006 17:05:09 +0200 | |
| changeset 19521 | cfdab6a91332 |
| parent 19502 | 369cde91963d |
| child 19576 | 179ad0076f75 |
| permissions | -rw-r--r-- |
| 18060 | 1 |
(* Title: Pure/consts.ML |
2 |
ID: $Id$ |
|
3 |
Author: Makarius |
|
4 |
||
| 18935 | 5 |
Polymorphic constants: declarations, abbreviations, additional type |
6 |
constraints. |
|
| 18060 | 7 |
*) |
8 |
||
9 |
signature CONSTS = |
|
10 |
sig |
|
11 |
type T |
|
| 19027 | 12 |
val eq_consts: T * T -> bool |
| 19364 | 13 |
val abbrevs_of: T -> string list -> (term * term) list |
| 18935 | 14 |
val dest: T -> |
15 |
{constants: (typ * term option) NameSpace.table,
|
|
16 |
constraints: typ NameSpace.table} |
|
| 19364 | 17 |
val declaration: T -> string -> typ (*exception TYPE*) |
18 |
val monomorphic: T -> string -> bool (*exception TYPE*) |
|
19 |
val constraint: T -> string -> typ (*exception TYPE*) |
|
| 18965 | 20 |
val space_of: T -> NameSpace.T |
| 19364 | 21 |
val intern: T -> xstring -> string |
22 |
val extern: T -> string -> xstring |
|
23 |
val extern_early: T -> string -> xstring |
|
| 18965 | 24 |
val read_const: T -> string -> term |
| 19364 | 25 |
val certify: Pretty.pp -> Type.tsig -> T -> term -> term (*exception TYPE*) |
| 18146 | 26 |
val typargs: T -> string * typ -> typ list |
| 18163 | 27 |
val instance: T -> string * typ list -> typ |
| 19364 | 28 |
val declare: NameSpace.naming -> (bstring * typ) * bool -> T -> T |
|
19098
fc736dbbe333
constrain: assert const declaration, optional type (i.e. may delete constraints);
wenzelm
parents:
19027
diff
changeset
|
29 |
val constrain: string * typ option -> T -> T |
| 19364 | 30 |
val expand_abbrevs: bool -> T -> T |
31 |
val abbreviate: Pretty.pp -> Type.tsig -> NameSpace.naming -> string -> |
|
32 |
(bstring * term) * bool -> T -> T |
|
| 18060 | 33 |
val hide: bool -> string -> T -> T |
34 |
val empty: T |
|
35 |
val merge: T * T -> T |
|
| 19364 | 36 |
end; |
| 18060 | 37 |
|
38 |
structure Consts: CONSTS = |
|
39 |
struct |
|
40 |
||
| 19364 | 41 |
|
42 |
(** consts type **) |
|
43 |
||
44 |
(* datatype T *) |
|
| 18060 | 45 |
|
| 19364 | 46 |
datatype kind = |
47 |
LogicalConst of int list list | |
|
48 |
Abbreviation of term; |
|
49 |
||
50 |
type decl = |
|
51 |
(typ * kind) * |
|
52 |
bool; (*early externing*) |
|
| 18935 | 53 |
|
| 18060 | 54 |
datatype T = Consts of |
| 19364 | 55 |
{decls: (decl * stamp) NameSpace.table,
|
56 |
constraints: typ Symtab.table, |
|
57 |
rev_abbrevs: (term * term) list Symtab.table, |
|
58 |
expand_abbrevs: bool} * stamp; |
|
| 19027 | 59 |
|
60 |
fun eq_consts (Consts (_, s1), Consts (_, s2)) = s1 = s2; |
|
| 18060 | 61 |
|
| 19364 | 62 |
fun make_consts (decls, constraints, rev_abbrevs, expand_abbrevs) = |
63 |
Consts ({decls = decls, constraints = constraints,
|
|
64 |
expand_abbrevs = expand_abbrevs, rev_abbrevs = rev_abbrevs}, stamp ()); |
|
| 18060 | 65 |
|
| 19364 | 66 |
fun map_consts f (Consts ({decls, constraints, rev_abbrevs, expand_abbrevs}, _)) =
|
67 |
make_consts (f (decls, constraints, rev_abbrevs, expand_abbrevs)); |
|
68 |
||
69 |
fun abbrevs_of (Consts ({rev_abbrevs, ...}, _)) modes =
|
|
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19433
diff
changeset
|
70 |
maps (Symtab.lookup_list rev_abbrevs) modes; |
| 19364 | 71 |
|
| 18965 | 72 |
|
| 19364 | 73 |
(* dest consts *) |
74 |
||
75 |
fun dest_kind (LogicalConst _) = NONE |
|
76 |
| dest_kind (Abbreviation t) = SOME t; |
|
77 |
||
78 |
fun dest (Consts ({decls = (space, decls), constraints, ...}, _)) =
|
|
79 |
{constants = (space,
|
|
80 |
Symtab.fold (fn (c, (((T, kind), _), _)) => |
|
81 |
Symtab.update (c, (T, dest_kind kind))) decls Symtab.empty), |
|
| 18060 | 82 |
constraints = (space, constraints)}; |
83 |
||
84 |
||
85 |
(* lookup consts *) |
|
86 |
||
| 19027 | 87 |
fun the_const (Consts ({decls = (_, tab), ...}, _)) c =
|
| 19364 | 88 |
(case Symtab.lookup tab c of |
89 |
SOME (decl, _) => decl |
|
90 |
| NONE => raise TYPE ("Undeclared constant: " ^ quote c, [], []));
|
|
| 18935 | 91 |
|
92 |
fun logical_const consts c = |
|
| 19364 | 93 |
(case the_const consts c of |
94 |
((T, LogicalConst ps), _) => (T, ps) |
|
95 |
| _ => raise TYPE ("Illegal abbreviation: " ^ quote c, [], []));
|
|
| 18060 | 96 |
|
| 19364 | 97 |
val declaration = #1 oo logical_const; |
98 |
val type_arguments = #2 oo logical_const; |
|
99 |
val monomorphic = null oo type_arguments; |
|
| 18935 | 100 |
|
| 19027 | 101 |
fun constraint (consts as Consts ({constraints, ...}, _)) c =
|
| 18060 | 102 |
(case Symtab.lookup constraints c of |
103 |
SOME T => T |
|
| 19364 | 104 |
| NONE => #1 (#1 (the_const consts c))); |
| 18935 | 105 |
|
106 |
||
| 19364 | 107 |
(* name space *) |
108 |
||
109 |
fun space_of (Consts ({decls = (space, _), ...}, _)) = space;
|
|
110 |
||
111 |
val intern = NameSpace.intern o space_of; |
|
112 |
val extern = NameSpace.extern o space_of; |
|
113 |
||
114 |
fun extern_early consts c = |
|
115 |
(case try (the_const consts) c of |
|
116 |
SOME (_, true) => extern consts c |
|
117 |
| _ => Syntax.constN ^ c); |
|
118 |
||
| 18060 | 119 |
|
| 19364 | 120 |
(* read_const *) |
121 |
||
122 |
fun read_const consts raw_c = |
|
123 |
let |
|
124 |
val c = intern consts raw_c; |
|
125 |
val _ = the_const consts c handle TYPE (msg, _, _) => error msg; |
|
126 |
in Const (c, dummyT) end; |
|
127 |
||
128 |
||
129 |
(* certify *) |
|
130 |
||
131 |
fun certify pp tsig (consts as Consts ({expand_abbrevs, ...}, _)) =
|
|
| 18965 | 132 |
let |
133 |
fun err msg (c, T) = |
|
134 |
raise TYPE (msg ^ " " ^ quote c ^ " :: " ^ Pretty.string_of_typ pp T, [], []); |
|
135 |
fun cert tm = |
|
136 |
let |
|
137 |
val (head, args) = Term.strip_comb tm; |
|
| 19364 | 138 |
fun comb h = Term.list_comb (h, map cert args); |
| 18965 | 139 |
in |
140 |
(case head of |
|
| 19364 | 141 |
Abs (x, T, t) => comb (Abs (x, T, cert t)) |
142 |
| Const (c, T) => |
|
143 |
let |
|
144 |
val T' = Type.cert_typ tsig T; |
|
145 |
val ((U, kind), _) = the_const consts c; |
|
146 |
in |
|
|
19433
c7a2b7a8c4cb
certify: ignore sort constraints of declarations (MAJOR CHANGE);
wenzelm
parents:
19364
diff
changeset
|
147 |
if not (Type.raw_instance (T', U)) then |
| 19364 | 148 |
err "Illegal type for constant" (c, T) |
149 |
else |
|
150 |
(case (kind, expand_abbrevs) of |
|
151 |
(Abbreviation u, true) => |
|
|
19433
c7a2b7a8c4cb
certify: ignore sort constraints of declarations (MAJOR CHANGE);
wenzelm
parents:
19364
diff
changeset
|
152 |
Term.betapplys (Envir.expand_atom T' (U, u) handle TYPE _ => |
| 19364 | 153 |
err "Illegal type for abbreviation" (c, T), map cert args) |
154 |
| _ => comb head) |
|
155 |
end |
|
156 |
| _ => comb head) |
|
| 18965 | 157 |
end; |
158 |
in cert end; |
|
159 |
||
160 |
||
| 18060 | 161 |
(* typargs -- view actual const type as instance of declaration *) |
162 |
||
| 19364 | 163 |
fun subscript (Type (_, Ts)) (i :: is) = subscript (nth Ts i) is |
164 |
| subscript T [] = T |
|
165 |
| subscript T _ = raise Subscript; |
|
| 18060 | 166 |
|
| 19364 | 167 |
fun typargs consts (c, T) = map (subscript T) (type_arguments consts c); |
| 18060 | 168 |
|
| 18163 | 169 |
fun instance consts (c, Ts) = |
170 |
let |
|
171 |
val declT = declaration consts c; |
|
172 |
val vars = map Term.dest_TVar (typargs consts (c, declT)); |
|
173 |
in declT |> Term.instantiateT (vars ~~ Ts) end; |
|
174 |
||
| 18060 | 175 |
|
176 |
||
| 19364 | 177 |
(** build consts **) |
| 18060 | 178 |
|
179 |
fun err_dup_consts cs = |
|
180 |
error ("Duplicate declaration of constant(s) " ^ commas_quote cs);
|
|
181 |
||
182 |
fun err_inconsistent_constraints cs = |
|
183 |
error ("Inconsistent type constraints for constant(s) " ^ commas_quote cs);
|
|
184 |
||
| 18935 | 185 |
fun extend_decls naming decl tab = NameSpace.extend_table naming (tab, [decl]) |
186 |
handle Symtab.DUPS cs => err_dup_consts cs; |
|
| 18060 | 187 |
|
| 18935 | 188 |
|
189 |
(* name space *) |
|
| 18060 | 190 |
|
| 19364 | 191 |
fun hide fully c = map_consts (fn (decls, constraints, rev_abbrevs, expand_abbrevs) => |
192 |
(apfst (NameSpace.hide fully c) decls, constraints, rev_abbrevs, expand_abbrevs)); |
|
| 18935 | 193 |
|
194 |
||
195 |
(* declarations *) |
|
196 |
||
| 19364 | 197 |
fun declare naming ((c, declT), early) = |
198 |
map_consts (fn (decls, constraints, rev_abbrevs, expand_abbrevs) => |
|
| 18060 | 199 |
let |
| 18935 | 200 |
fun args_of (Type (_, Ts)) pos = args_of_list Ts 0 pos |
201 |
| args_of (TVar v) pos = insert (eq_fst op =) (v, rev pos) |
|
202 |
| args_of (TFree _) _ = I |
|
203 |
and args_of_list (T :: Ts) i is = args_of T (i :: is) #> args_of_list Ts (i + 1) is |
|
204 |
| args_of_list [] _ _ = I; |
|
| 19364 | 205 |
val decl = (((declT, LogicalConst (map #2 (rev (args_of declT [] [])))), early), stamp ()); |
206 |
in (extend_decls naming (c, decl) decls, constraints, rev_abbrevs, expand_abbrevs) end); |
|
207 |
||
208 |
||
209 |
(* constraints *) |
|
210 |
||
211 |
fun constrain (c, C) consts = |
|
212 |
consts |> map_consts (fn (decls, constraints, rev_abbrevs, expand_abbrevs) => |
|
213 |
(the_const consts c handle TYPE (msg, _, _) => error msg; |
|
214 |
(decls, |
|
215 |
constraints |> (case C of SOME T => Symtab.update (c, T) | NONE => Symtab.delete_safe c), |
|
216 |
rev_abbrevs, expand_abbrevs))); |
|
| 18060 | 217 |
|
| 18935 | 218 |
|
219 |
(* abbreviations *) |
|
220 |
||
| 19364 | 221 |
fun expand_abbrevs b = map_consts (fn (decls, constraints, rev_abbrevs, _) => |
222 |
(decls, constraints, rev_abbrevs, b)); |
|
223 |
||
| 19027 | 224 |
local |
225 |
||
| 19364 | 226 |
fun strip_abss tm = ([], tm) :: |
227 |
(case tm of |
|
228 |
Abs (a, T, t) => |
|
229 |
if Term.loose_bvar1 (t, 0) then |
|
230 |
strip_abss t |> map (fn (xs, b) => ((a, T) :: xs, b)) |
|
231 |
else [] |
|
232 |
| _ => []); |
|
| 19027 | 233 |
|
| 19364 | 234 |
fun rev_abbrev const rhs = |
| 18060 | 235 |
let |
| 19027 | 236 |
fun abbrev (xs, body) = |
237 |
let val vars = fold (fn (x, T) => cons (Var ((x, 0), T))) (Term.rename_wrt_term body xs) [] |
|
238 |
in (Term.subst_bounds (rev vars, body), Term.list_comb (Const const, vars)) end; |
|
| 19502 | 239 |
in map abbrev (strip_abss (Envir.beta_eta_contract rhs)) end; |
| 19027 | 240 |
|
241 |
in |
|
| 18965 | 242 |
|
| 19364 | 243 |
fun abbreviate pp tsig naming mode ((c, raw_rhs), early) consts = |
| 18965 | 244 |
let |
| 19364 | 245 |
val rhs = raw_rhs |
246 |
|> Term.map_term_types (Type.cert_typ tsig) |
|
247 |
|> certify pp tsig (consts |> expand_abbrevs false); |
|
248 |
val rhs' = rhs |
|
249 |
|> certify pp tsig (consts |> expand_abbrevs true); |
|
| 18965 | 250 |
val T = Term.fastype_of rhs; |
| 19364 | 251 |
in |
252 |
consts |> map_consts (fn (decls, constraints, rev_abbrevs, expand_abbrevs) => |
|
253 |
let |
|
254 |
val decls' = decls |
|
255 |
|> extend_decls naming (c, (((T, Abbreviation rhs'), early), stamp ())); |
|
256 |
val rev_abbrevs' = rev_abbrevs |
|
| 19502 | 257 |
|> fold (curry Symtab.update_list mode) (rev_abbrev (NameSpace.full naming c, T) rhs); |
| 19364 | 258 |
in (decls', constraints, rev_abbrevs', expand_abbrevs) end) |
259 |
end; |
|
| 18935 | 260 |
|
| 19027 | 261 |
end; |
262 |
||
| 18060 | 263 |
|
264 |
(* empty and merge *) |
|
265 |
||
| 19364 | 266 |
val empty = make_consts (NameSpace.empty_table, Symtab.empty, Symtab.empty, true); |
| 18060 | 267 |
|
268 |
fun merge |
|
| 19364 | 269 |
(Consts ({decls = decls1, constraints = constraints1,
|
270 |
rev_abbrevs = rev_abbrevs1, expand_abbrevs = expand_abbrevs1}, _), |
|
271 |
Consts ({decls = decls2, constraints = constraints2,
|
|
272 |
rev_abbrevs = rev_abbrevs2, expand_abbrevs = expand_abbrevs2}, _)) = |
|
| 18060 | 273 |
let |
| 18935 | 274 |
val decls' = NameSpace.merge_tables (eq_snd (op =)) (decls1, decls2) |
| 18060 | 275 |
handle Symtab.DUPS cs => err_dup_consts cs; |
| 18935 | 276 |
val constraints' = Symtab.merge (op =) (constraints1, constraints2) |
| 18060 | 277 |
handle Symtab.DUPS cs => err_inconsistent_constraints cs; |
| 19364 | 278 |
val rev_abbrevs' = (rev_abbrevs1, rev_abbrevs2) |> Symtab.join |
279 |
(K (Library.merge (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u'))); |
|
280 |
val expand_abbrevs' = expand_abbrevs1 orelse expand_abbrevs2; |
|
281 |
in make_consts (decls', constraints', rev_abbrevs', expand_abbrevs') end; |
|
| 18060 | 282 |
|
283 |
end; |