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