author | haftmann |
Mon, 08 Oct 2007 08:04:28 +0200 | |
changeset 24901 | d3cbf79769b9 |
parent 24772 | 2b838dfeca1e |
child 24909 | afae05eb1f1c |
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 -> |
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
15 |
{constants: (typ * (term * term) option) NameSpace.table, |
18935 | 16 |
constraints: typ NameSpace.table} |
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
17 |
val the_abbreviation: T -> string -> typ * (term * term) (*exception TYPE*) |
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
18 |
val the_declaration: T -> string -> typ (*exception TYPE*) |
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
19 |
val is_monomorphic: T -> string -> bool (*exception TYPE*) |
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
20 |
val the_constraint: T -> string -> typ (*exception TYPE*) |
24772 | 21 |
val the_tags: T -> string -> Markup.property list (*exception TYPE*) |
18965 | 22 |
val space_of: T -> NameSpace.T |
19364 | 23 |
val intern: T -> xstring -> string |
24 |
val extern: T -> string -> xstring |
|
25 |
val extern_early: T -> string -> xstring |
|
19657 | 26 |
val syntax: T -> string * mixfix -> string * typ * mixfix |
21181 | 27 |
val syntax_name: T -> string -> string |
18965 | 28 |
val read_const: T -> string -> term |
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
29 |
val certify: Pretty.pp -> Type.tsig -> bool -> T -> term -> term (*exception TYPE*) |
18146 | 30 |
val typargs: T -> string * typ -> typ list |
18163 | 31 |
val instance: T -> string * typ list -> typ |
24772 | 32 |
val declare: bool -> NameSpace.naming -> Markup.property list -> (bstring * typ) -> T -> T |
19098
fc736dbbe333
constrain: assert const declaration, optional type (i.e. may delete constraints);
wenzelm
parents:
19027
diff
changeset
|
33 |
val constrain: string * typ option -> T -> T |
24772 | 34 |
val abbreviate: Pretty.pp -> Type.tsig -> NameSpace.naming -> string -> Markup.property list -> |
21794 | 35 |
bstring * term -> T -> (term * term) * T |
18060 | 36 |
val hide: bool -> string -> T -> T |
37 |
val empty: T |
|
38 |
val merge: T * T -> T |
|
19364 | 39 |
end; |
18060 | 40 |
|
41 |
structure Consts: CONSTS = |
|
42 |
struct |
|
43 |
||
19364 | 44 |
|
45 |
(** consts type **) |
|
46 |
||
47 |
(* datatype T *) |
|
18060 | 48 |
|
19364 | 49 |
datatype kind = |
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
50 |
LogicalConst of int list list | (*typargs positions*) |
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
51 |
Abbreviation of term * term * bool (*rhs, normal rhs, force_expand*); |
19364 | 52 |
|
24772 | 53 |
type decl = {T: typ, kind: kind, tags: Markup.property list, authentic: bool}; |
18935 | 54 |
|
18060 | 55 |
datatype T = Consts of |
20667 | 56 |
{decls: (decl * serial) NameSpace.table, |
19364 | 57 |
constraints: typ Symtab.table, |
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
58 |
rev_abbrevs: (term * term) list Symtab.table} * stamp; |
19027 | 59 |
|
60 |
fun eq_consts (Consts (_, s1), Consts (_, s2)) = s1 = s2; |
|
18060 | 61 |
|
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
62 |
fun make_consts (decls, constraints, rev_abbrevs) = |
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
63 |
Consts ({decls = decls, constraints = constraints, rev_abbrevs = rev_abbrevs}, stamp ()); |
18060 | 64 |
|
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
65 |
fun map_consts f (Consts ({decls, constraints, rev_abbrevs}, _)) = |
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
66 |
make_consts (f (decls, constraints, rev_abbrevs)); |
19364 | 67 |
|
68 |
fun abbrevs_of (Consts ({rev_abbrevs, ...}, _)) modes = |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19433
diff
changeset
|
69 |
maps (Symtab.lookup_list rev_abbrevs) modes; |
19364 | 70 |
|
18965 | 71 |
|
19364 | 72 |
(* dest consts *) |
73 |
||
74 |
fun dest_kind (LogicalConst _) = NONE |
|
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
75 |
| dest_kind (Abbreviation (t, t', _)) = SOME (t, t'); |
19364 | 76 |
|
77 |
fun dest (Consts ({decls = (space, decls), constraints, ...}, _)) = |
|
78 |
{constants = (space, |
|
24772 | 79 |
Symtab.fold (fn (c, ({T, kind, ...}, _)) => |
19364 | 80 |
Symtab.update (c, (T, dest_kind kind))) decls Symtab.empty), |
18060 | 81 |
constraints = (space, constraints)}; |
82 |
||
83 |
||
84 |
(* lookup consts *) |
|
85 |
||
19027 | 86 |
fun the_const (Consts ({decls = (_, tab), ...}, _)) c = |
19364 | 87 |
(case Symtab.lookup tab c of |
24772 | 88 |
SOME (decl, _) => decl |
19364 | 89 |
| NONE => raise TYPE ("Undeclared constant: " ^ quote c, [], [])); |
18935 | 90 |
|
91 |
fun logical_const consts c = |
|
24772 | 92 |
(case the_const consts c of |
93 |
{T, kind = LogicalConst ps, ...} => (T, ps) |
|
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
94 |
| _ => raise TYPE ("Not a logical constant: " ^ quote c, [], [])); |
18060 | 95 |
|
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
96 |
fun the_abbreviation consts c = |
24772 | 97 |
(case the_const consts c of |
98 |
{T, kind = Abbreviation (t, t', _), ...} => (T, (t, t')) |
|
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
99 |
| _ => raise TYPE ("Not an abbreviated constant: " ^ quote c, [], [])); |
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
100 |
|
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
101 |
val the_declaration = #1 oo logical_const; |
19364 | 102 |
val type_arguments = #2 oo logical_const; |
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
103 |
val is_monomorphic = null oo type_arguments; |
18935 | 104 |
|
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
105 |
fun the_constraint (consts as Consts ({constraints, ...}, _)) c = |
18060 | 106 |
(case Symtab.lookup constraints c of |
107 |
SOME T => T |
|
24772 | 108 |
| NONE => #T (the_const consts c)); |
109 |
||
110 |
val the_tags = #tags oo the_const; |
|
18935 | 111 |
|
112 |
||
19657 | 113 |
(* name space and syntax *) |
19364 | 114 |
|
115 |
fun space_of (Consts ({decls = (space, _), ...}, _)) = space; |
|
116 |
||
117 |
val intern = NameSpace.intern o space_of; |
|
118 |
val extern = NameSpace.extern o space_of; |
|
119 |
||
120 |
fun extern_early consts c = |
|
121 |
(case try (the_const consts) c of |
|
24772 | 122 |
SOME {authentic = true, ...} => Syntax.constN ^ c |
19576
179ad0076f75
extern_early: improved handling of undeclared constants;
wenzelm
parents:
19502
diff
changeset
|
123 |
| _ => extern consts c); |
19364 | 124 |
|
19657 | 125 |
fun syntax consts (c, mx) = |
126 |
let |
|
24772 | 127 |
val {T, authentic, ...} = the_const consts c handle TYPE (msg, _, _) => error msg; |
19677 | 128 |
val c' = if authentic then Syntax.constN ^ c else NameSpace.base c; |
19657 | 129 |
in (c', T, mx) end; |
130 |
||
21181 | 131 |
fun syntax_name consts c = #1 (syntax consts (c, NoSyn)); |
132 |
||
18060 | 133 |
|
19364 | 134 |
(* read_const *) |
135 |
||
136 |
fun read_const consts raw_c = |
|
137 |
let |
|
138 |
val c = intern consts raw_c; |
|
24772 | 139 |
val {T, ...} = the_const consts c handle TYPE (msg, _, _) => error msg; |
21205 | 140 |
in Const (c, T) end; |
19364 | 141 |
|
142 |
||
143 |
(* certify *) |
|
144 |
||
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
145 |
fun certify pp tsig do_expand consts = |
18965 | 146 |
let |
147 |
fun err msg (c, T) = |
|
148 |
raise TYPE (msg ^ " " ^ quote c ^ " :: " ^ Pretty.string_of_typ pp T, [], []); |
|
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
149 |
val certT = Type.cert_typ tsig; |
18965 | 150 |
fun cert tm = |
151 |
let |
|
152 |
val (head, args) = Term.strip_comb tm; |
|
21694 | 153 |
val args' = map cert args; |
154 |
fun comb head' = Term.list_comb (head', args'); |
|
18965 | 155 |
in |
156 |
(case head of |
|
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
157 |
Abs (x, T, t) => comb (Abs (x, certT T, cert t)) |
19364 | 158 |
| Const (c, T) => |
159 |
let |
|
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
160 |
val T' = certT T; |
24772 | 161 |
val {T = U, kind, ...} = the_const consts c; |
19364 | 162 |
in |
19433
c7a2b7a8c4cb
certify: ignore sort constraints of declarations (MAJOR CHANGE);
wenzelm
parents:
19364
diff
changeset
|
163 |
if not (Type.raw_instance (T', U)) then |
19364 | 164 |
err "Illegal type for constant" (c, T) |
165 |
else |
|
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
166 |
(case kind of |
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
167 |
Abbreviation (_, u, force_expand) => |
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
168 |
if do_expand orelse force_expand then |
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
169 |
Term.betapplys (Envir.expand_atom T' (U, u) handle TYPE _ => |
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
170 |
err "Illegal type for abbreviation" (c, T), args') |
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
171 |
else comb head |
19364 | 172 |
| _ => comb head) |
173 |
end |
|
174 |
| _ => comb head) |
|
18965 | 175 |
end; |
176 |
in cert end; |
|
177 |
||
178 |
||
18060 | 179 |
(* typargs -- view actual const type as instance of declaration *) |
180 |
||
19364 | 181 |
fun subscript (Type (_, Ts)) (i :: is) = subscript (nth Ts i) is |
182 |
| subscript T [] = T |
|
183 |
| subscript T _ = raise Subscript; |
|
18060 | 184 |
|
19364 | 185 |
fun typargs consts (c, T) = map (subscript T) (type_arguments consts c); |
18060 | 186 |
|
18163 | 187 |
fun instance consts (c, Ts) = |
188 |
let |
|
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
189 |
val declT = the_declaration consts c; |
18163 | 190 |
val vars = map Term.dest_TVar (typargs consts (c, declT)); |
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
191 |
val inst = vars ~~ Ts handle UnequalLengths => |
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
192 |
raise TYPE ("Consts.instance", Ts, [Const (c, dummyT)]); |
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
193 |
in declT |> TermSubst.instantiateT inst end; |
18163 | 194 |
|
18060 | 195 |
|
196 |
||
19364 | 197 |
(** build consts **) |
18060 | 198 |
|
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23086
diff
changeset
|
199 |
fun err_dup_const c = |
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23086
diff
changeset
|
200 |
error ("Duplicate declaration of constant " ^ quote c); |
18060 | 201 |
|
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23086
diff
changeset
|
202 |
fun err_inconsistent_constraints c = |
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23086
diff
changeset
|
203 |
error ("Inconsistent type constraints for constant " ^ quote c); |
18060 | 204 |
|
23086 | 205 |
fun extend_decls naming decl tab = NameSpace.extend_table naming [decl] tab |
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23086
diff
changeset
|
206 |
handle Symtab.DUP c => err_dup_const c; |
18060 | 207 |
|
18935 | 208 |
|
209 |
(* name space *) |
|
18060 | 210 |
|
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
211 |
fun hide fully c = map_consts (fn (decls, constraints, rev_abbrevs) => |
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
212 |
(apfst (NameSpace.hide fully c) decls, constraints, rev_abbrevs)); |
18935 | 213 |
|
214 |
||
215 |
(* declarations *) |
|
216 |
||
24772 | 217 |
fun declare authentic naming tags (c, declT) = map_consts (fn (decls, constraints, rev_abbrevs) => |
18060 | 218 |
let |
18935 | 219 |
fun args_of (Type (_, Ts)) pos = args_of_list Ts 0 pos |
220 |
| args_of (TVar v) pos = insert (eq_fst op =) (v, rev pos) |
|
221 |
| args_of (TFree _) _ = I |
|
222 |
and args_of_list (T :: Ts) i is = args_of T (i :: is) #> args_of_list Ts (i + 1) is |
|
223 |
| args_of_list [] _ _ = I; |
|
20667 | 224 |
val decl = |
24772 | 225 |
{T = declT, kind = LogicalConst (map #2 (rev (args_of declT [] []))), |
226 |
tags = tags, authentic = authentic}; |
|
227 |
val decls' = decls |> extend_decls naming (c, (decl, serial ())); |
|
228 |
in (decls', constraints, rev_abbrevs) end); |
|
19364 | 229 |
|
230 |
||
231 |
(* constraints *) |
|
232 |
||
233 |
fun constrain (c, C) consts = |
|
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
234 |
consts |> map_consts (fn (decls, constraints, rev_abbrevs) => |
19364 | 235 |
(the_const consts c handle TYPE (msg, _, _) => error msg; |
236 |
(decls, |
|
237 |
constraints |> (case C of SOME T => Symtab.update (c, T) | NONE => Symtab.delete_safe c), |
|
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
238 |
rev_abbrevs))); |
18060 | 239 |
|
18935 | 240 |
|
241 |
(* abbreviations *) |
|
242 |
||
19027 | 243 |
local |
244 |
||
19364 | 245 |
fun strip_abss tm = ([], tm) :: |
246 |
(case tm of |
|
247 |
Abs (a, T, t) => |
|
248 |
if Term.loose_bvar1 (t, 0) then |
|
249 |
strip_abss t |> map (fn (xs, b) => ((a, T) :: xs, b)) |
|
250 |
else [] |
|
251 |
| _ => []); |
|
19027 | 252 |
|
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
253 |
fun rev_abbrev lhs rhs = |
18060 | 254 |
let |
19027 | 255 |
fun abbrev (xs, body) = |
256 |
let val vars = fold (fn (x, T) => cons (Var ((x, 0), T))) (Term.rename_wrt_term body xs) [] |
|
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
257 |
in (Term.subst_bounds (rev vars, body), Term.list_comb (lhs, vars)) end; |
19502 | 258 |
in map abbrev (strip_abss (Envir.beta_eta_contract rhs)) end; |
19027 | 259 |
|
260 |
in |
|
18965 | 261 |
|
24772 | 262 |
fun abbreviate pp tsig naming mode tags (c, raw_rhs) consts = |
18965 | 263 |
let |
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
264 |
val cert_term = certify pp tsig false consts; |
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
265 |
val expand_term = certify pp tsig true consts; |
21822 | 266 |
val force_expand = mode = Syntax.internalM; |
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
267 |
|
19364 | 268 |
val rhs = raw_rhs |
20548
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
wenzelm
parents:
20509
diff
changeset
|
269 |
|> Term.map_types (Type.cert_typ tsig) |
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
270 |
|> cert_term; |
21794 | 271 |
val T = Term.fastype_of rhs; |
272 |
val lhs = Const (NameSpace.full naming c, T); |
|
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
273 |
val rhs' = expand_term rhs; |
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
274 |
|
21680
5d2230ad1261
abbreviate: improved error handling, return result;
wenzelm
parents:
21205
diff
changeset
|
275 |
fun err msg = error (msg ^ " on rhs of abbreviation:\n" ^ |
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
276 |
Pretty.string_of_term pp (Logic.mk_equals (lhs, rhs))); |
21680
5d2230ad1261
abbreviate: improved error handling, return result;
wenzelm
parents:
21205
diff
changeset
|
277 |
val _ = Term.exists_subterm Term.is_Var rhs andalso err "Illegal schematic variables" |
5d2230ad1261
abbreviate: improved error handling, return result;
wenzelm
parents:
21205
diff
changeset
|
278 |
val _ = null (Term.hidden_polymorphism rhs T) orelse err "Extra type variables"; |
19364 | 279 |
in |
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
280 |
consts |> map_consts (fn (decls, constraints, rev_abbrevs) => |
19364 | 281 |
let |
24772 | 282 |
val decl = |
283 |
{T = T, kind = Abbreviation (rhs, rhs', force_expand), tags = tags, authentic = true}; |
|
284 |
val decls' = decls |
|
285 |
|> extend_decls naming (c, (decl, serial ())); |
|
19364 | 286 |
val rev_abbrevs' = rev_abbrevs |
21720
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
wenzelm
parents:
21694
diff
changeset
|
287 |
|> fold (curry Symtab.update_list mode) (rev_abbrev lhs rhs); |
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
288 |
in (decls', constraints, rev_abbrevs') end) |
21794 | 289 |
|> pair (lhs, rhs) |
19364 | 290 |
end; |
18935 | 291 |
|
19027 | 292 |
end; |
293 |
||
18060 | 294 |
|
295 |
(* empty and merge *) |
|
296 |
||
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
297 |
val empty = make_consts (NameSpace.empty_table, Symtab.empty, Symtab.empty); |
18060 | 298 |
|
299 |
fun merge |
|
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
300 |
(Consts ({decls = decls1, constraints = constraints1, rev_abbrevs = rev_abbrevs1}, _), |
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
301 |
Consts ({decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}, _)) = |
18060 | 302 |
let |
18935 | 303 |
val decls' = NameSpace.merge_tables (eq_snd (op =)) (decls1, decls2) |
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23086
diff
changeset
|
304 |
handle Symtab.DUP c => err_dup_const c; |
18935 | 305 |
val constraints' = Symtab.merge (op =) (constraints1, constraints2) |
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23086
diff
changeset
|
306 |
handle Symtab.DUP c => err_inconsistent_constraints c; |
19364 | 307 |
val rev_abbrevs' = (rev_abbrevs1, rev_abbrevs2) |> Symtab.join |
308 |
(K (Library.merge (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u'))); |
|
24673
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
wenzelm
parents:
23655
diff
changeset
|
309 |
in make_consts (decls', constraints', rev_abbrevs') end; |
18060 | 310 |
|
311 |
end; |