| author | wenzelm | 
| Tue, 21 Jul 2009 20:37:32 +0200 | |
| changeset 32104 | d1d98a02473e | 
| parent 31977 | e03059ae2d82 | 
| child 32784 | 1a5dde5079ac | 
| permissions | -rw-r--r-- | 
| 18060 | 1 | (* Title: Pure/consts.ML | 
| 2 | Author: Makarius | |
| 3 | ||
| 18935 | 4 | Polymorphic constants: declarations, abbreviations, additional type | 
| 5 | constraints. | |
| 18060 | 6 | *) | 
| 7 | ||
| 8 | signature CONSTS = | |
| 9 | sig | |
| 10 | type T | |
| 30424 
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
 wenzelm parents: 
30364diff
changeset | 11 | val eq_consts: T * T -> bool | 
| 30566 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 12 | val retrieve_abbrevs: T -> string list -> term -> (term * term) list | 
| 18935 | 13 | val dest: T -> | 
| 25404 | 14 |    {constants: (typ * term option) NameSpace.table,
 | 
| 18935 | 15 | constraints: typ NameSpace.table} | 
| 25048 | 16 | val the_type: T -> string -> typ (*exception TYPE*) | 
| 17 | val the_abbreviation: T -> string -> typ * term (*exception TYPE*) | |
| 18 | val type_scheme: T -> string -> typ (*exception TYPE*) | |
| 28017 | 19 | val the_tags: T -> string -> Properties.T (*exception TYPE*) | 
| 25048 | 20 | val is_monomorphic: T -> string -> bool (*exception TYPE*) | 
| 21 | val the_constraint: T -> string -> typ (*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: 
23655diff
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 | 
| 29581 | 32 | val declare: bool -> NameSpace.naming -> Properties.T -> (binding * typ) -> T -> T | 
| 19098 
fc736dbbe333
constrain: assert const declaration, optional type (i.e. may delete constraints);
 wenzelm parents: 
19027diff
changeset | 33 | val constrain: string * typ option -> T -> T | 
| 28017 | 34 | val abbreviate: Pretty.pp -> Type.tsig -> NameSpace.naming -> string -> Properties.T -> | 
| 29581 | 35 | binding * term -> T -> (term * term) * T | 
| 25048 | 36 | val revert_abbrev: string -> string -> T -> T | 
| 18060 | 37 | val hide: bool -> string -> T -> T | 
| 38 | val empty: T | |
| 39 | val merge: T * T -> T | |
| 19364 | 40 | end; | 
| 18060 | 41 | |
| 42 | structure Consts: CONSTS = | |
| 43 | struct | |
| 44 | ||
| 19364 | 45 | (** consts type **) | 
| 46 | ||
| 47 | (* datatype T *) | |
| 18060 | 48 | |
| 28017 | 49 | type decl = {T: typ, typargs: int list list, tags: Properties.T, authentic: bool};
 | 
| 24909 | 50 | type abbrev = {rhs: term, normal_rhs: term, force_expand: bool};
 | 
| 18935 | 51 | |
| 18060 | 52 | datatype T = Consts of | 
| 24909 | 53 |  {decls: ((decl * abbrev option) * serial) NameSpace.table,
 | 
| 19364 | 54 | constraints: typ Symtab.table, | 
| 30566 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 55 | rev_abbrevs: (term * term) Item_Net.T Symtab.table}; | 
| 18060 | 56 | |
| 30424 
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
 wenzelm parents: 
30364diff
changeset | 57 | fun eq_consts | 
| 
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
 wenzelm parents: 
30364diff
changeset | 58 |    (Consts {decls = decls1, constraints = constraints1, rev_abbrevs = rev_abbrevs1},
 | 
| 
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
 wenzelm parents: 
30364diff
changeset | 59 |     Consts {decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}) =
 | 
| 
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
 wenzelm parents: 
30364diff
changeset | 60 | pointer_eq (decls1, decls2) andalso | 
| 
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
 wenzelm parents: 
30364diff
changeset | 61 | pointer_eq (constraints1, constraints2) andalso | 
| 
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
 wenzelm parents: 
30364diff
changeset | 62 | pointer_eq (rev_abbrevs1, rev_abbrevs2); | 
| 
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
 wenzelm parents: 
30364diff
changeset | 63 | |
| 24673 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 64 | fun make_consts (decls, constraints, rev_abbrevs) = | 
| 30284 
907da436f8a9
eliminated Consts.eq_consts tuning -- this is built into tables and name spaces already;
 wenzelm parents: 
30280diff
changeset | 65 |   Consts {decls = decls, constraints = constraints, rev_abbrevs = rev_abbrevs};
 | 
| 18060 | 66 | |
| 30284 
907da436f8a9
eliminated Consts.eq_consts tuning -- this is built into tables and name spaces already;
 wenzelm parents: 
30280diff
changeset | 67 | fun map_consts f (Consts {decls, constraints, rev_abbrevs}) =
 | 
| 24673 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 68 | make_consts (f (decls, constraints, rev_abbrevs)); | 
| 19364 | 69 | |
| 30566 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 70 | |
| 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 71 | (* reverted abbrevs *) | 
| 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 72 | |
| 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 73 | val empty_abbrevs = Item_Net.init (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u') #1; | 
| 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 74 | |
| 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 75 | fun insert_abbrevs mode abbrs = | 
| 30568 
e6a55291102e
strip_abss: always strip abstractions as far as possible, without keeping alternatives (which appear to be redundant anyway, but cause significant slowdown since discrimination nets collapse abstractions);
 wenzelm parents: 
30566diff
changeset | 76 | Symtab.map_default (mode, empty_abbrevs) (Item_Net.insert abbrs); | 
| 30566 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 77 | |
| 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 78 | fun retrieve_abbrevs (Consts {rev_abbrevs, ...}) modes =
 | 
| 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 79 | let val nets = map_filter (Symtab.lookup rev_abbrevs) modes | 
| 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 80 | in fn t => maps (fn net => Item_Net.retrieve net t) nets end; | 
| 19364 | 81 | |
| 18965 | 82 | |
| 19364 | 83 | (* dest consts *) | 
| 84 | ||
| 30284 
907da436f8a9
eliminated Consts.eq_consts tuning -- this is built into tables and name spaces already;
 wenzelm parents: 
30280diff
changeset | 85 | fun dest (Consts {decls = (space, decls), constraints, ...}) =
 | 
| 19364 | 86 |  {constants = (space,
 | 
| 24909 | 87 |     Symtab.fold (fn (c, (({T, ...}, abbr), _)) =>
 | 
| 25404 | 88 | Symtab.update (c, (T, Option.map #rhs abbr))) decls Symtab.empty), | 
| 18060 | 89 | constraints = (space, constraints)}; | 
| 90 | ||
| 91 | ||
| 92 | (* lookup consts *) | |
| 93 | ||
| 30284 
907da436f8a9
eliminated Consts.eq_consts tuning -- this is built into tables and name spaces already;
 wenzelm parents: 
30280diff
changeset | 94 | fun the_const (Consts {decls = (_, tab), ...}) c =
 | 
| 19364 | 95 | (case Symtab.lookup tab c of | 
| 24772 | 96 | SOME (decl, _) => decl | 
| 25041 | 97 |   | NONE => raise TYPE ("Unknown constant: " ^ quote c, [], []));
 | 
| 18935 | 98 | |
| 25041 | 99 | fun the_type consts c = | 
| 24772 | 100 | (case the_const consts c of | 
| 24909 | 101 |     ({T, ...}, NONE) => T
 | 
| 21720 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 wenzelm parents: 
21694diff
changeset | 102 |   | _ => raise TYPE ("Not a logical constant: " ^ quote c, [], []));
 | 
| 18060 | 103 | |
| 21720 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 wenzelm parents: 
21694diff
changeset | 104 | fun the_abbreviation consts c = | 
| 24772 | 105 | (case the_const consts c of | 
| 25048 | 106 |     ({T, ...}, SOME {rhs, ...}) => (T, rhs)
 | 
| 21720 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 wenzelm parents: 
21694diff
changeset | 107 |   | _ => raise TYPE ("Not an abbreviated constant: " ^ quote c, [], []));
 | 
| 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 wenzelm parents: 
21694diff
changeset | 108 | |
| 25041 | 109 | val the_decl = #1 oo the_const; | 
| 110 | val type_scheme = #T oo the_decl; | |
| 111 | val type_arguments = #typargs oo the_decl; | |
| 112 | val the_tags = #tags oo the_decl; | |
| 113 | ||
| 21720 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 wenzelm parents: 
21694diff
changeset | 114 | val is_monomorphic = null oo type_arguments; | 
| 18935 | 115 | |
| 30284 
907da436f8a9
eliminated Consts.eq_consts tuning -- this is built into tables and name spaces already;
 wenzelm parents: 
30280diff
changeset | 116 | fun the_constraint (consts as Consts {constraints, ...}) c =
 | 
| 18060 | 117 | (case Symtab.lookup constraints c of | 
| 118 | SOME T => T | |
| 25041 | 119 | | NONE => type_scheme consts c); | 
| 18935 | 120 | |
| 121 | ||
| 19657 | 122 | (* name space and syntax *) | 
| 19364 | 123 | |
| 30284 
907da436f8a9
eliminated Consts.eq_consts tuning -- this is built into tables and name spaces already;
 wenzelm parents: 
30280diff
changeset | 124 | fun space_of (Consts {decls = (space, _), ...}) = space;
 | 
| 19364 | 125 | |
| 126 | val intern = NameSpace.intern o space_of; | |
| 127 | val extern = NameSpace.extern o space_of; | |
| 128 | ||
| 129 | fun extern_early consts c = | |
| 130 | (case try (the_const consts) c of | |
| 24909 | 131 |     SOME ({authentic = true, ...}, _) => Syntax.constN ^ c
 | 
| 19576 
179ad0076f75
extern_early: improved handling of undeclared constants;
 wenzelm parents: 
19502diff
changeset | 132 | | _ => extern consts c); | 
| 19364 | 133 | |
| 19657 | 134 | fun syntax consts (c, mx) = | 
| 135 | let | |
| 24909 | 136 |     val ({T, authentic, ...}, _) = the_const consts c handle TYPE (msg, _, _) => error msg;
 | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30286diff
changeset | 137 | val c' = if authentic then Syntax.constN ^ c else Long_Name.base_name c; | 
| 19657 | 138 | in (c', T, mx) end; | 
| 139 | ||
| 21181 | 140 | fun syntax_name consts c = #1 (syntax consts (c, NoSyn)); | 
| 141 | ||
| 18060 | 142 | |
| 19364 | 143 | (* read_const *) | 
| 144 | ||
| 145 | fun read_const consts raw_c = | |
| 146 | let | |
| 147 | val c = intern consts raw_c; | |
| 25041 | 148 | val T = type_scheme consts c handle TYPE (msg, _, _) => error msg; | 
| 21205 | 149 | in Const (c, T) end; | 
| 19364 | 150 | |
| 151 | ||
| 152 | (* certify *) | |
| 153 | ||
| 24673 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 154 | fun certify pp tsig do_expand consts = | 
| 18965 | 155 | let | 
| 156 | fun err msg (c, T) = | |
| 157 | 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: 
23655diff
changeset | 158 | val certT = Type.cert_typ tsig; | 
| 18965 | 159 | fun cert tm = | 
| 160 | let | |
| 161 | val (head, args) = Term.strip_comb tm; | |
| 21694 | 162 | val args' = map cert args; | 
| 163 | fun comb head' = Term.list_comb (head', args'); | |
| 18965 | 164 | in | 
| 165 | (case head of | |
| 24673 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 166 | Abs (x, T, t) => comb (Abs (x, certT T, cert t)) | 
| 19364 | 167 | | Const (c, T) => | 
| 168 | let | |
| 24673 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 169 | val T' = certT T; | 
| 24909 | 170 |               val ({T = U, ...}, abbr) = the_const consts c;
 | 
| 25048 | 171 | fun expand u = | 
| 172 | Term.betapplys (Envir.expand_atom T' (U, u) handle TYPE _ => | |
| 173 | err "Illegal type for abbreviation" (c, T), args'); | |
| 19364 | 174 | in | 
| 19433 
c7a2b7a8c4cb
certify: ignore sort constraints of declarations (MAJOR CHANGE);
 wenzelm parents: 
19364diff
changeset | 175 | if not (Type.raw_instance (T', U)) then | 
| 19364 | 176 | err "Illegal type for constant" (c, T) | 
| 177 | else | |
| 24909 | 178 | (case abbr of | 
| 25048 | 179 |                   SOME {rhs, normal_rhs, force_expand} =>
 | 
| 180 | if do_expand then expand normal_rhs | |
| 181 | else if force_expand then expand rhs | |
| 21720 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 wenzelm parents: 
21694diff
changeset | 182 | else comb head | 
| 19364 | 183 | | _ => comb head) | 
| 184 | end | |
| 185 | | _ => comb head) | |
| 18965 | 186 | end; | 
| 187 | in cert end; | |
| 188 | ||
| 189 | ||
| 18060 | 190 | (* typargs -- view actual const type as instance of declaration *) | 
| 191 | ||
| 24909 | 192 | local | 
| 193 | ||
| 194 | fun args_of (Type (_, Ts)) pos = args_of_list Ts 0 pos | |
| 195 | | args_of (TVar v) pos = insert (eq_fst op =) (v, rev pos) | |
| 196 | | args_of (TFree _) _ = I | |
| 197 | and args_of_list (T :: Ts) i is = args_of T (i :: is) #> args_of_list Ts (i + 1) is | |
| 198 | | args_of_list [] _ _ = I; | |
| 199 | ||
| 19364 | 200 | fun subscript (Type (_, Ts)) (i :: is) = subscript (nth Ts i) is | 
| 201 | | subscript T [] = T | |
| 202 | | subscript T _ = raise Subscript; | |
| 18060 | 203 | |
| 24909 | 204 | in | 
| 205 | ||
| 206 | fun typargs_of T = map #2 (rev (args_of T [] [])); | |
| 207 | ||
| 19364 | 208 | fun typargs consts (c, T) = map (subscript T) (type_arguments consts c); | 
| 18060 | 209 | |
| 24909 | 210 | end; | 
| 211 | ||
| 18163 | 212 | fun instance consts (c, Ts) = | 
| 213 | let | |
| 25041 | 214 | val declT = type_scheme consts c; | 
| 18163 | 215 | 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: 
23655diff
changeset | 216 | val inst = vars ~~ Ts handle UnequalLengths => | 
| 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 217 |       raise TYPE ("Consts.instance", Ts, [Const (c, dummyT)]);
 | 
| 31977 | 218 | in declT |> Term_Subst.instantiateT inst end; | 
| 18163 | 219 | |
| 18060 | 220 | |
| 221 | ||
| 19364 | 222 | (** build consts **) | 
| 18060 | 223 | |
| 23655 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
23086diff
changeset | 224 | fun err_dup_const c = | 
| 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
23086diff
changeset | 225 |   error ("Duplicate declaration of constant " ^ quote c);
 | 
| 18060 | 226 | |
| 30466 | 227 | fun extend_decls naming decl tab = NameSpace.define naming decl tab | 
| 23655 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
23086diff
changeset | 228 | handle Symtab.DUP c => err_dup_const c; | 
| 18060 | 229 | |
| 18935 | 230 | |
| 231 | (* name space *) | |
| 18060 | 232 | |
| 24673 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 233 | 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: 
23655diff
changeset | 234 | (apfst (NameSpace.hide fully c) decls, constraints, rev_abbrevs)); | 
| 18935 | 235 | |
| 236 | ||
| 237 | (* declarations *) | |
| 238 | ||
| 28861 | 239 | fun declare authentic naming tags (b, declT) = map_consts (fn (decls, constraints, rev_abbrevs) => | 
| 18060 | 240 | let | 
| 26050 | 241 | val tags' = tags |> Position.default_properties (Position.thread_data ()); | 
| 242 |     val decl = {T = declT, typargs = typargs_of declT, tags = tags', authentic = authentic};
 | |
| 28861 | 243 | val (_, decls') = decls |> extend_decls naming (b, ((decl, NONE), serial ())); | 
| 24772 | 244 | in (decls', constraints, rev_abbrevs) end); | 
| 19364 | 245 | |
| 246 | ||
| 247 | (* constraints *) | |
| 248 | ||
| 249 | fun constrain (c, C) consts = | |
| 24673 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 250 | consts |> map_consts (fn (decls, constraints, rev_abbrevs) => | 
| 19364 | 251 | (the_const consts c handle TYPE (msg, _, _) => error msg; | 
| 252 | (decls, | |
| 253 | 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: 
23655diff
changeset | 254 | rev_abbrevs))); | 
| 18060 | 255 | |
| 18935 | 256 | |
| 257 | (* abbreviations *) | |
| 258 | ||
| 19027 | 259 | local | 
| 260 | ||
| 30568 
e6a55291102e
strip_abss: always strip abstractions as far as possible, without keeping alternatives (which appear to be redundant anyway, but cause significant slowdown since discrimination nets collapse abstractions);
 wenzelm parents: 
30566diff
changeset | 261 | fun strip_abss (t as Abs (x, T, b)) = | 
| 
e6a55291102e
strip_abss: always strip abstractions as far as possible, without keeping alternatives (which appear to be redundant anyway, but cause significant slowdown since discrimination nets collapse abstractions);
 wenzelm parents: 
30566diff
changeset | 262 | if Term.loose_bvar1 (b, 0) then strip_abss b |>> cons (x, T) | 
| 
e6a55291102e
strip_abss: always strip abstractions as far as possible, without keeping alternatives (which appear to be redundant anyway, but cause significant slowdown since discrimination nets collapse abstractions);
 wenzelm parents: 
30566diff
changeset | 263 | else ([], t) | 
| 
e6a55291102e
strip_abss: always strip abstractions as far as possible, without keeping alternatives (which appear to be redundant anyway, but cause significant slowdown since discrimination nets collapse abstractions);
 wenzelm parents: 
30566diff
changeset | 264 | | strip_abss t = ([], t); | 
| 19027 | 265 | |
| 21720 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 wenzelm parents: 
21694diff
changeset | 266 | fun rev_abbrev lhs rhs = | 
| 18060 | 267 | let | 
| 30568 
e6a55291102e
strip_abss: always strip abstractions as far as possible, without keeping alternatives (which appear to be redundant anyway, but cause significant slowdown since discrimination nets collapse abstractions);
 wenzelm parents: 
30566diff
changeset | 268 | val (xs, body) = strip_abss (Envir.beta_eta_contract rhs); | 
| 
e6a55291102e
strip_abss: always strip abstractions as far as possible, without keeping alternatives (which appear to be redundant anyway, but cause significant slowdown since discrimination nets collapse abstractions);
 wenzelm parents: 
30566diff
changeset | 269 | val vars = fold (fn (x, T) => cons (Var ((x, 0), T))) (Term.rename_wrt_term body xs) []; | 
| 
e6a55291102e
strip_abss: always strip abstractions as far as possible, without keeping alternatives (which appear to be redundant anyway, but cause significant slowdown since discrimination nets collapse abstractions);
 wenzelm parents: 
30566diff
changeset | 270 | in (Term.subst_bounds (rev vars, body), Term.list_comb (lhs, vars)) end; | 
| 19027 | 271 | |
| 272 | in | |
| 18965 | 273 | |
| 28861 | 274 | fun abbreviate pp tsig naming mode tags (b, raw_rhs) consts = | 
| 18965 | 275 | let | 
| 24673 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 276 | val cert_term = certify pp tsig false consts; | 
| 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 277 | val expand_term = certify pp tsig true consts; | 
| 25116 | 278 | val force_expand = mode = PrintMode.internal; | 
| 21720 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 wenzelm parents: 
21694diff
changeset | 279 | |
| 30286 
cf89a03ee308
Consts.abbreviate: reject schematic term variables, prevent schematic type variables (hidden polymorphism) via Term.close_schematic_term -- see also 8f84a608883d;
 wenzelm parents: 
30284diff
changeset | 280 | val _ = Term.exists_subterm Term.is_Var raw_rhs andalso | 
| 
cf89a03ee308
Consts.abbreviate: reject schematic term variables, prevent schematic type variables (hidden polymorphism) via Term.close_schematic_term -- see also 8f84a608883d;
 wenzelm parents: 
30284diff
changeset | 281 |       error ("Illegal schematic variables on rhs of abbreviation: " ^ Binding.str_of b);
 | 
| 
cf89a03ee308
Consts.abbreviate: reject schematic term variables, prevent schematic type variables (hidden polymorphism) via Term.close_schematic_term -- see also 8f84a608883d;
 wenzelm parents: 
30284diff
changeset | 282 | |
| 19364 | 283 | val rhs = raw_rhs | 
| 20548 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 wenzelm parents: 
20509diff
changeset | 284 | |> Term.map_types (Type.cert_typ tsig) | 
| 30286 
cf89a03ee308
Consts.abbreviate: reject schematic term variables, prevent schematic type variables (hidden polymorphism) via Term.close_schematic_term -- see also 8f84a608883d;
 wenzelm parents: 
30284diff
changeset | 285 | |> cert_term | 
| 
cf89a03ee308
Consts.abbreviate: reject schematic term variables, prevent schematic type variables (hidden polymorphism) via Term.close_schematic_term -- see also 8f84a608883d;
 wenzelm parents: 
30284diff
changeset | 286 | |> Term.close_schematic_term; | 
| 25404 | 287 | val normal_rhs = expand_term rhs; | 
| 21794 | 288 | val T = Term.fastype_of rhs; | 
| 28965 | 289 | val lhs = Const (NameSpace.full_name naming b, T); | 
| 19364 | 290 | in | 
| 24673 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 291 | consts |> map_consts (fn (decls, constraints, rev_abbrevs) => | 
| 19364 | 292 | let | 
| 26050 | 293 | val tags' = tags |> Position.default_properties (Position.thread_data ()); | 
| 294 |         val decl = {T = T, typargs = typargs_of T, tags = tags', authentic = true};
 | |
| 25404 | 295 |         val abbr = {rhs = rhs, normal_rhs = normal_rhs, force_expand = force_expand};
 | 
| 28861 | 296 | val (_, decls') = decls | 
| 297 | |> extend_decls naming (b, ((decl, SOME abbr), serial ())); | |
| 19364 | 298 | val rev_abbrevs' = rev_abbrevs | 
| 30566 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 299 | |> insert_abbrevs mode (rev_abbrev lhs rhs); | 
| 24673 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 300 | in (decls', constraints, rev_abbrevs') end) | 
| 21794 | 301 | |> pair (lhs, rhs) | 
| 19364 | 302 | end; | 
| 18935 | 303 | |
| 25048 | 304 | fun revert_abbrev mode c consts = consts |> map_consts (fn (decls, constraints, rev_abbrevs) => | 
| 305 | let | |
| 306 | val (T, rhs) = the_abbreviation consts c; | |
| 307 | val rev_abbrevs' = rev_abbrevs | |
| 30566 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 308 | |> insert_abbrevs mode (rev_abbrev (Const (c, T)) rhs); | 
| 25048 | 309 | in (decls, constraints, rev_abbrevs') end); | 
| 310 | ||
| 19027 | 311 | end; | 
| 312 | ||
| 18060 | 313 | |
| 314 | (* empty and merge *) | |
| 315 | ||
| 24673 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 316 | val empty = make_consts (NameSpace.empty_table, Symtab.empty, Symtab.empty); | 
| 18060 | 317 | |
| 318 | fun merge | |
| 30284 
907da436f8a9
eliminated Consts.eq_consts tuning -- this is built into tables and name spaces already;
 wenzelm parents: 
30280diff
changeset | 319 |    (Consts {decls = decls1, constraints = constraints1, rev_abbrevs = rev_abbrevs1},
 | 
| 
907da436f8a9
eliminated Consts.eq_consts tuning -- this is built into tables and name spaces already;
 wenzelm parents: 
30280diff
changeset | 320 |     Consts {decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}) =
 | 
| 18060 | 321 | let | 
| 18935 | 322 | val decls' = NameSpace.merge_tables (eq_snd (op =)) (decls1, decls2) | 
| 23655 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
23086diff
changeset | 323 | handle Symtab.DUP c => err_dup_const c; | 
| 25037 | 324 | val constraints' = Symtab.join (K fst) (constraints1, constraints2); | 
| 30566 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 wenzelm parents: 
30466diff
changeset | 325 | val rev_abbrevs' = Symtab.join (K Item_Net.merge) (rev_abbrevs1, rev_abbrevs2); | 
| 24673 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 wenzelm parents: 
23655diff
changeset | 326 | in make_consts (decls', constraints', rev_abbrevs') end; | 
| 18060 | 327 | |
| 328 | end; |