| author | haftmann | 
| Mon, 22 Nov 2010 17:46:51 +0100 | |
| changeset 40672 | abd4e7358847 | 
| parent 37146 | f652333bbf8e | 
| child 40722 | 441260986b63 | 
| 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: 
30364 
diff
changeset
 | 
11  | 
val eq_consts: T * T -> bool  | 
| 
30566
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
12  | 
val retrieve_abbrevs: T -> string list -> term -> (term * term) list  | 
| 18935 | 13  | 
val dest: T ->  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
14  | 
   {constants: (typ * term option) Name_Space.table,
 | 
| 
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
15  | 
constraints: typ Name_Space.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*)  | 
|
19  | 
val is_monomorphic: T -> string -> bool (*exception TYPE*)  | 
|
20  | 
val the_constraint: T -> string -> typ (*exception TYPE*)  | 
|
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
21  | 
val space_of: T -> Name_Space.T  | 
| 35680 | 22  | 
val alias: Name_Space.naming -> binding -> string -> T -> T  | 
| 33158 | 23  | 
val is_concealed: T -> string -> bool  | 
| 19364 | 24  | 
val intern: T -> xstring -> string  | 
| 35554 | 25  | 
val extern: T -> string -> xstring  | 
| 35255 | 26  | 
val intern_syntax: T -> xstring -> string  | 
| 35554 | 27  | 
val extern_syntax: T -> string -> xstring  | 
| 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  | 
| 35255 | 32  | 
val declare: Name_Space.naming -> binding * 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  | 
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33158 
diff
changeset
 | 
34  | 
val abbreviate: Pretty.pp -> Type.tsig -> Name_Space.naming -> string ->  | 
| 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  | 
|
| 35255 | 49  | 
type decl = {T: typ, typargs: int list list};
 | 
| 24909 | 50  | 
type abbrev = {rhs: term, normal_rhs: term, force_expand: bool};
 | 
| 18935 | 51  | 
|
| 18060 | 52  | 
datatype T = Consts of  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
53  | 
 {decls: (decl * abbrev option) Name_Space.table,
 | 
| 19364 | 54  | 
constraints: typ Symtab.table,  | 
| 
30566
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
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: 
30364 
diff
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: 
30364 
diff
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: 
30364 
diff
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: 
30364 
diff
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: 
30364 
diff
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: 
30364 
diff
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: 
30364 
diff
changeset
 | 
63  | 
|
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
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: 
30280 
diff
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: 
30280 
diff
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: 
23655 
diff
changeset
 | 
68  | 
make_consts (f (decls, constraints, rev_abbrevs));  | 
| 19364 | 69  | 
|
| 
30566
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
70  | 
|
| 
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
71  | 
(* reverted abbrevs *)  | 
| 
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
72  | 
|
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33158 
diff
changeset
 | 
73  | 
val empty_abbrevs =  | 
| 33373 | 74  | 
Item_Net.init (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u') (single o #1);  | 
| 
30566
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
75  | 
|
| 33373 | 76  | 
fun update_abbrevs mode abbrs =  | 
77  | 
Symtab.map_default (mode, empty_abbrevs) (Item_Net.update abbrs);  | 
|
| 
30566
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
78  | 
|
| 
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
79  | 
fun retrieve_abbrevs (Consts {rev_abbrevs, ...}) modes =
 | 
| 
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
80  | 
let val nets = map_filter (Symtab.lookup rev_abbrevs) modes  | 
| 
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
81  | 
in fn t => maps (fn net => Item_Net.retrieve net t) nets end;  | 
| 19364 | 82  | 
|
| 18965 | 83  | 
|
| 19364 | 84  | 
(* dest consts *)  | 
85  | 
||
| 
30284
 
907da436f8a9
eliminated Consts.eq_consts tuning -- this is built into tables and name spaces already;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
86  | 
fun dest (Consts {decls = (space, decls), constraints, ...}) =
 | 
| 19364 | 87  | 
 {constants = (space,
 | 
| 
33092
 
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
88  | 
    Symtab.fold (fn (c, ({T, ...}, abbr)) =>
 | 
| 25404 | 89  | 
Symtab.update (c, (T, Option.map #rhs abbr))) decls Symtab.empty),  | 
| 18060 | 90  | 
constraints = (space, constraints)};  | 
91  | 
||
92  | 
||
93  | 
(* lookup consts *)  | 
|
94  | 
||
| 
30284
 
907da436f8a9
eliminated Consts.eq_consts tuning -- this is built into tables and name spaces already;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
95  | 
fun the_const (Consts {decls = (_, tab), ...}) c =
 | 
| 19364 | 96  | 
(case Symtab.lookup tab c of  | 
| 
33092
 
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
97  | 
SOME decl => decl  | 
| 25041 | 98  | 
  | NONE => raise TYPE ("Unknown constant: " ^ quote c, [], []));
 | 
| 18935 | 99  | 
|
| 25041 | 100  | 
fun the_type consts c =  | 
| 24772 | 101  | 
(case the_const consts c of  | 
| 24909 | 102  | 
    ({T, ...}, NONE) => T
 | 
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
103  | 
  | _ => raise TYPE ("Not a logical constant: " ^ quote c, [], []));
 | 
| 18060 | 104  | 
|
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
105  | 
fun the_abbreviation consts c =  | 
| 24772 | 106  | 
(case the_const consts c of  | 
| 25048 | 107  | 
    ({T, ...}, SOME {rhs, ...}) => (T, rhs)
 | 
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
108  | 
  | _ => raise TYPE ("Not an abbreviated constant: " ^ quote c, [], []));
 | 
| 
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
109  | 
|
| 25041 | 110  | 
val the_decl = #1 oo the_const;  | 
111  | 
val type_scheme = #T oo the_decl;  | 
|
112  | 
val type_arguments = #typargs oo the_decl;  | 
|
113  | 
||
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
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: 
30280 
diff
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: 
30280 
diff
changeset
 | 
124  | 
fun space_of (Consts {decls = (space, _), ...}) = space;
 | 
| 19364 | 125  | 
|
| 35680 | 126  | 
fun alias naming binding name = map_consts (fn ((space, decls), constraints, rev_abbrevs) =>  | 
127  | 
((Name_Space.alias naming binding name space, decls), constraints, rev_abbrevs));  | 
|
128  | 
||
| 33158 | 129  | 
val is_concealed = Name_Space.is_concealed o space_of;  | 
130  | 
||
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
131  | 
val intern = Name_Space.intern o space_of;  | 
| 
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
132  | 
val extern = Name_Space.extern o space_of;  | 
| 19364 | 133  | 
|
| 35554 | 134  | 
fun intern_syntax consts s =  | 
135  | 
(case try Syntax.unmark_const s of  | 
|
| 35255 | 136  | 
SOME c => c  | 
| 35554 | 137  | 
| NONE => intern consts s);  | 
138  | 
||
139  | 
fun extern_syntax consts s =  | 
|
140  | 
(case try Syntax.unmark_const s of  | 
|
141  | 
SOME c => extern consts c  | 
|
142  | 
| NONE => s);  | 
|
| 21181 | 143  | 
|
| 18060 | 144  | 
|
| 19364 | 145  | 
(* read_const *)  | 
146  | 
||
147  | 
fun read_const consts raw_c =  | 
|
148  | 
let  | 
|
149  | 
val c = intern consts raw_c;  | 
|
| 25041 | 150  | 
val T = type_scheme consts c handle TYPE (msg, _, _) => error msg;  | 
| 21205 | 151  | 
in Const (c, T) end;  | 
| 19364 | 152  | 
|
153  | 
||
154  | 
(* certify *)  | 
|
155  | 
||
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
156  | 
fun certify pp tsig do_expand consts =  | 
| 18965 | 157  | 
let  | 
158  | 
fun err msg (c, T) =  | 
|
159  | 
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
 | 
160  | 
val certT = Type.cert_typ tsig;  | 
| 18965 | 161  | 
fun cert tm =  | 
162  | 
let  | 
|
163  | 
val (head, args) = Term.strip_comb tm;  | 
|
| 21694 | 164  | 
val args' = map cert args;  | 
165  | 
fun comb head' = Term.list_comb (head', args');  | 
|
| 18965 | 166  | 
in  | 
167  | 
(case head of  | 
|
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
168  | 
Abs (x, T, t) => comb (Abs (x, certT T, cert t))  | 
| 19364 | 169  | 
| Const (c, T) =>  | 
170  | 
let  | 
|
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
171  | 
val T' = certT T;  | 
| 24909 | 172  | 
              val ({T = U, ...}, abbr) = the_const consts c;
 | 
| 25048 | 173  | 
fun expand u =  | 
174  | 
Term.betapplys (Envir.expand_atom T' (U, u) handle TYPE _ =>  | 
|
175  | 
err "Illegal type for abbreviation" (c, T), args');  | 
|
| 19364 | 176  | 
in  | 
| 
19433
 
c7a2b7a8c4cb
certify: ignore sort constraints of declarations (MAJOR CHANGE);
 
wenzelm 
parents: 
19364 
diff
changeset
 | 
177  | 
if not (Type.raw_instance (T', U)) then  | 
| 19364 | 178  | 
err "Illegal type for constant" (c, T)  | 
179  | 
else  | 
|
| 24909 | 180  | 
(case abbr of  | 
| 25048 | 181  | 
                  SOME {rhs, normal_rhs, force_expand} =>
 | 
182  | 
if do_expand then expand normal_rhs  | 
|
183  | 
else if force_expand then expand rhs  | 
|
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
184  | 
else comb head  | 
| 19364 | 185  | 
| _ => comb head)  | 
186  | 
end  | 
|
187  | 
| _ => comb head)  | 
|
| 18965 | 188  | 
end;  | 
189  | 
in cert end;  | 
|
190  | 
||
191  | 
||
| 18060 | 192  | 
(* typargs -- view actual const type as instance of declaration *)  | 
193  | 
||
| 24909 | 194  | 
local  | 
195  | 
||
196  | 
fun args_of (Type (_, Ts)) pos = args_of_list Ts 0 pos  | 
|
197  | 
| args_of (TVar v) pos = insert (eq_fst op =) (v, rev pos)  | 
|
198  | 
| args_of (TFree _) _ = I  | 
|
199  | 
and args_of_list (T :: Ts) i is = args_of T (i :: is) #> args_of_list Ts (i + 1) is  | 
|
200  | 
| args_of_list [] _ _ = I;  | 
|
201  | 
||
| 19364 | 202  | 
fun subscript (Type (_, Ts)) (i :: is) = subscript (nth Ts i) is  | 
203  | 
| subscript T [] = T  | 
|
| 32784 | 204  | 
| subscript _ _ = raise Subscript;  | 
| 18060 | 205  | 
|
| 24909 | 206  | 
in  | 
207  | 
||
208  | 
fun typargs_of T = map #2 (rev (args_of T [] []));  | 
|
209  | 
||
| 19364 | 210  | 
fun typargs consts (c, T) = map (subscript T) (type_arguments consts c);  | 
| 18060 | 211  | 
|
| 24909 | 212  | 
end;  | 
213  | 
||
| 18163 | 214  | 
fun instance consts (c, Ts) =  | 
215  | 
let  | 
|
| 25041 | 216  | 
val declT = type_scheme consts c;  | 
| 18163 | 217  | 
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
 | 
218  | 
val inst = vars ~~ Ts handle UnequalLengths =>  | 
| 
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
219  | 
      raise TYPE ("Consts.instance", Ts, [Const (c, dummyT)]);
 | 
| 31977 | 220  | 
in declT |> Term_Subst.instantiateT inst end;  | 
| 18163 | 221  | 
|
| 18060 | 222  | 
|
223  | 
||
| 19364 | 224  | 
(** build consts **)  | 
| 18060 | 225  | 
|
| 18935 | 226  | 
(* name space *)  | 
| 18060 | 227  | 
|
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
228  | 
fun hide fully c = map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
229  | 
(apfst (Name_Space.hide fully c) decls, constraints, rev_abbrevs));  | 
| 18935 | 230  | 
|
231  | 
||
232  | 
(* declarations *)  | 
|
233  | 
||
| 35255 | 234  | 
fun declare naming (b, declT) =  | 
| 
33092
 
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
235  | 
map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
| 
 
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
236  | 
let  | 
| 35255 | 237  | 
      val decl = {T = declT, typargs = typargs_of declT};
 | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
238  | 
val (_, decls') = decls |> Name_Space.define true naming (b, (decl, NONE));  | 
| 
33092
 
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
239  | 
in (decls', constraints, rev_abbrevs) end);  | 
| 19364 | 240  | 
|
241  | 
||
242  | 
(* constraints *)  | 
|
243  | 
||
244  | 
fun constrain (c, C) consts =  | 
|
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
245  | 
consts |> map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
| 19364 | 246  | 
(the_const consts c handle TYPE (msg, _, _) => error msg;  | 
247  | 
(decls,  | 
|
248  | 
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
 | 
249  | 
rev_abbrevs)));  | 
| 18060 | 250  | 
|
| 18935 | 251  | 
|
252  | 
(* abbreviations *)  | 
|
253  | 
||
| 19027 | 254  | 
local  | 
255  | 
||
| 
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: 
30566 
diff
changeset
 | 
256  | 
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: 
30566 
diff
changeset
 | 
257  | 
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: 
30566 
diff
changeset
 | 
258  | 
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: 
30566 
diff
changeset
 | 
259  | 
| strip_abss t = ([], t);  | 
| 19027 | 260  | 
|
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
261  | 
fun rev_abbrev lhs rhs =  | 
| 18060 | 262  | 
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: 
30566 
diff
changeset
 | 
263  | 
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: 
30566 
diff
changeset
 | 
264  | 
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: 
30566 
diff
changeset
 | 
265  | 
in (Term.subst_bounds (rev vars, body), Term.list_comb (lhs, vars)) end;  | 
| 19027 | 266  | 
|
267  | 
in  | 
|
| 18965 | 268  | 
|
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33158 
diff
changeset
 | 
269  | 
fun abbreviate pp tsig naming mode (b, raw_rhs) consts =  | 
| 18965 | 270  | 
let  | 
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
271  | 
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
 | 
272  | 
val expand_term = certify pp tsig true consts;  | 
| 
37146
 
f652333bbf8e
renamed structure PrintMode to Print_Mode, keeping the old name as legacy alias for some time;
 
wenzelm 
parents: 
35680 
diff
changeset
 | 
273  | 
val force_expand = mode = Print_Mode.internal;  | 
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
274  | 
|
| 
30286
 
cf89a03ee308
Consts.abbreviate: reject schematic term variables, prevent schematic type variables (hidden polymorphism) via Term.close_schematic_term -- see also 8f84a608883d;
 
wenzelm 
parents: 
30284 
diff
changeset
 | 
275  | 
val _ = Term.exists_subterm Term.is_Var raw_rhs andalso  | 
| 
33092
 
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
276  | 
      error ("Illegal schematic variables on rhs of abbreviation " ^ quote (Binding.str_of b));
 | 
| 
30286
 
cf89a03ee308
Consts.abbreviate: reject schematic term variables, prevent schematic type variables (hidden polymorphism) via Term.close_schematic_term -- see also 8f84a608883d;
 
wenzelm 
parents: 
30284 
diff
changeset
 | 
277  | 
|
| 19364 | 278  | 
val rhs = raw_rhs  | 
| 
20548
 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 
wenzelm 
parents: 
20509 
diff
changeset
 | 
279  | 
|> 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: 
30284 
diff
changeset
 | 
280  | 
|> 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: 
30284 
diff
changeset
 | 
281  | 
|> Term.close_schematic_term;  | 
| 25404 | 282  | 
val normal_rhs = expand_term rhs;  | 
| 21794 | 283  | 
val T = Term.fastype_of rhs;  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
284  | 
val lhs = Const (Name_Space.full_name naming b, T);  | 
| 19364 | 285  | 
in  | 
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
286  | 
consts |> map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
| 19364 | 287  | 
let  | 
| 35255 | 288  | 
        val decl = {T = T, typargs = typargs_of T};
 | 
| 25404 | 289  | 
        val abbr = {rhs = rhs, normal_rhs = normal_rhs, force_expand = force_expand};
 | 
| 28861 | 290  | 
val (_, decls') = decls  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
291  | 
|> Name_Space.define true naming (b, (decl, SOME abbr));  | 
| 19364 | 292  | 
val rev_abbrevs' = rev_abbrevs  | 
| 33373 | 293  | 
|> update_abbrevs mode (rev_abbrev lhs rhs);  | 
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
294  | 
in (decls', constraints, rev_abbrevs') end)  | 
| 21794 | 295  | 
|> pair (lhs, rhs)  | 
| 19364 | 296  | 
end;  | 
| 18935 | 297  | 
|
| 25048 | 298  | 
fun revert_abbrev mode c consts = consts |> map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
299  | 
let  | 
|
300  | 
val (T, rhs) = the_abbreviation consts c;  | 
|
301  | 
val rev_abbrevs' = rev_abbrevs  | 
|
| 33373 | 302  | 
|> update_abbrevs mode (rev_abbrev (Const (c, T)) rhs);  | 
| 25048 | 303  | 
in (decls, constraints, rev_abbrevs') end);  | 
304  | 
||
| 19027 | 305  | 
end;  | 
306  | 
||
| 18060 | 307  | 
|
308  | 
(* empty and merge *)  | 
|
309  | 
||
| 33096 | 310  | 
val empty = make_consts (Name_Space.empty_table "constant", Symtab.empty, Symtab.empty);  | 
| 18060 | 311  | 
|
312  | 
fun merge  | 
|
| 
30284
 
907da436f8a9
eliminated Consts.eq_consts tuning -- this is built into tables and name spaces already;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
313  | 
   (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: 
30280 
diff
changeset
 | 
314  | 
    Consts {decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}) =
 | 
| 18060 | 315  | 
let  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
316  | 
val decls' = Name_Space.merge_tables (decls1, decls2);  | 
| 25037 | 317  | 
val constraints' = Symtab.join (K fst) (constraints1, constraints2);  | 
| 
30566
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
318  | 
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: 
23655 
diff
changeset
 | 
319  | 
in make_consts (decls', constraints', rev_abbrevs') end;  | 
| 18060 | 320  | 
|
321  | 
end;  |