| author | wenzelm | 
| Fri, 01 Apr 2016 17:49:03 +0200 | |
| changeset 62799 | 46e6f91c4da1 | 
| parent 61262 | 7bd1eb4b056e | 
| child 63573 | 8976c5bc9e97 | 
| 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  | 
| 
56056
 
4d46d53566e6
more efficient local theory operations, by imposing a linear change discipline on the main types/consts tables, in order to speed-up Proof_Context.transfer_syntax required for Local_Theory.raw_theory_result;
 
wenzelm 
parents: 
56025 
diff
changeset
 | 
12  | 
val change_base: bool -> T -> T  | 
| 
56139
 
b7add947a6ef
more frugal recording of changes: join merely requires information from one side;
 
wenzelm 
parents: 
56062 
diff
changeset
 | 
13  | 
val change_ignore: T -> T  | 
| 
30566
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
14  | 
val retrieve_abbrevs: T -> string list -> term -> (term * term) list  | 
| 18935 | 15  | 
val dest: T ->  | 
| 56025 | 16  | 
   {const_space: Name_Space.T,
 | 
| 56062 | 17  | 
constants: (string * (typ * term option)) list,  | 
18  | 
constraints: (string * typ) list}  | 
|
| 
43794
 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
19  | 
val the_const: T -> string -> string * typ (*exception TYPE*)  | 
| 25048 | 20  | 
val the_abbreviation: T -> string -> typ * term (*exception TYPE*)  | 
21  | 
val type_scheme: T -> string -> typ (*exception TYPE*)  | 
|
22  | 
val is_monomorphic: T -> string -> bool (*exception TYPE*)  | 
|
23  | 
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
 | 
24  | 
val space_of: T -> Name_Space.T  | 
| 35680 | 25  | 
val alias: Name_Space.naming -> binding -> string -> T -> T  | 
| 33158 | 26  | 
val is_concealed: T -> string -> bool  | 
| 19364 | 27  | 
val intern: T -> xstring -> string  | 
| 35255 | 28  | 
val intern_syntax: T -> xstring -> string  | 
| 
55956
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
29  | 
val check_const: Context.generic -> T -> xstring * Position.T list -> term * Position.report list  | 
| 
61262
 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 
wenzelm 
parents: 
56139 
diff
changeset
 | 
30  | 
val certify: Context.generic -> Type.tsig -> bool -> T -> term -> term (*exception TYPE*)  | 
| 18146 | 31  | 
val typargs: T -> string * typ -> typ list  | 
| 18163 | 32  | 
val instance: T -> string * typ list -> typ  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
33  | 
val declare: Context.generic -> binding * typ -> T -> T  | 
| 
19098
 
fc736dbbe333
constrain: assert const declaration, optional type (i.e. may delete constraints);
 
wenzelm 
parents: 
19027 
diff
changeset
 | 
34  | 
val constrain: string * typ option -> T -> T  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
35  | 
val abbreviate: Context.generic -> Type.tsig -> string -> 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  | 
|
| 
56056
 
4d46d53566e6
more efficient local theory operations, by imposing a linear change discipline on the main types/consts tables, in order to speed-up Proof_Context.transfer_syntax required for Local_Theory.raw_theory_result;
 
wenzelm 
parents: 
56025 
diff
changeset
 | 
70  | 
fun change_base begin = map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
| 
 
4d46d53566e6
more efficient local theory operations, by imposing a linear change discipline on the main types/consts tables, in order to speed-up Proof_Context.transfer_syntax required for Local_Theory.raw_theory_result;
 
wenzelm 
parents: 
56025 
diff
changeset
 | 
71  | 
(Name_Space.change_base begin decls, constraints, rev_abbrevs));  | 
| 
 
4d46d53566e6
more efficient local theory operations, by imposing a linear change discipline on the main types/consts tables, in order to speed-up Proof_Context.transfer_syntax required for Local_Theory.raw_theory_result;
 
wenzelm 
parents: 
56025 
diff
changeset
 | 
72  | 
|
| 
56139
 
b7add947a6ef
more frugal recording of changes: join merely requires information from one side;
 
wenzelm 
parents: 
56062 
diff
changeset
 | 
73  | 
val change_ignore = map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
| 
 
b7add947a6ef
more frugal recording of changes: join merely requires information from one side;
 
wenzelm 
parents: 
56062 
diff
changeset
 | 
74  | 
(Name_Space.change_ignore decls, constraints, rev_abbrevs));  | 
| 
 
b7add947a6ef
more frugal recording of changes: join merely requires information from one side;
 
wenzelm 
parents: 
56062 
diff
changeset
 | 
75  | 
|
| 
30566
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
76  | 
|
| 
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
77  | 
(* reverted abbrevs *)  | 
| 
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
78  | 
|
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33158 
diff
changeset
 | 
79  | 
val empty_abbrevs =  | 
| 33373 | 80  | 
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
 | 
81  | 
|
| 33373 | 82  | 
fun update_abbrevs mode abbrs =  | 
83  | 
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
 | 
84  | 
|
| 
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
85  | 
fun retrieve_abbrevs (Consts {rev_abbrevs, ...}) modes =
 | 
| 
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
86  | 
let val nets = map_filter (Symtab.lookup rev_abbrevs) modes  | 
| 
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
87  | 
in fn t => maps (fn net => Item_Net.retrieve net t) nets end;  | 
| 19364 | 88  | 
|
| 18965 | 89  | 
|
| 19364 | 90  | 
(* dest consts *)  | 
91  | 
||
| 56025 | 92  | 
fun dest (Consts {decls, constraints, ...}) =
 | 
93  | 
 {const_space = Name_Space.space_of_table decls,
 | 
|
94  | 
constants =  | 
|
95  | 
    Name_Space.fold_table (fn (c, ({T, ...}, abbr)) =>
 | 
|
| 56062 | 96  | 
cons (c, (T, Option.map #rhs abbr))) decls [],  | 
97  | 
constraints = Symtab.dest constraints};  | 
|
| 18060 | 98  | 
|
99  | 
||
100  | 
(* lookup consts *)  | 
|
101  | 
||
| 56025 | 102  | 
fun the_entry (Consts {decls, ...}) c =
 | 
103  | 
(case Name_Space.lookup_key decls c of  | 
|
| 
43794
 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
104  | 
SOME entry => entry  | 
| 25041 | 105  | 
  | NONE => raise TYPE ("Unknown constant: " ^ quote c, [], []));
 | 
| 18935 | 106  | 
|
| 
43794
 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
107  | 
fun the_const consts c =  | 
| 
 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
108  | 
(case the_entry consts c of  | 
| 
 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
109  | 
    (c', ({T, ...}, NONE)) => (c', T)
 | 
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
110  | 
  | _ => raise TYPE ("Not a logical constant: " ^ quote c, [], []));
 | 
| 18060 | 111  | 
|
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
112  | 
fun the_abbreviation consts c =  | 
| 
43794
 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
113  | 
(case the_entry consts c of  | 
| 
 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
114  | 
    (_, ({T, ...}, SOME {rhs, ...})) => (T, rhs)
 | 
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
115  | 
  | _ => raise TYPE ("Not an abbreviated constant: " ^ quote c, [], []));
 | 
| 
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
116  | 
|
| 
43794
 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
117  | 
fun the_decl consts = #1 o #2 o the_entry consts;  | 
| 25041 | 118  | 
val type_scheme = #T oo the_decl;  | 
119  | 
val type_arguments = #typargs oo the_decl;  | 
|
120  | 
||
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
121  | 
val is_monomorphic = null oo type_arguments;  | 
| 18935 | 122  | 
|
| 
30284
 
907da436f8a9
eliminated Consts.eq_consts tuning -- this is built into tables and name spaces already;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
123  | 
fun the_constraint (consts as Consts {constraints, ...}) c =
 | 
| 18060 | 124  | 
(case Symtab.lookup constraints c of  | 
125  | 
SOME T => T  | 
|
| 25041 | 126  | 
| NONE => type_scheme consts c);  | 
| 18935 | 127  | 
|
128  | 
||
| 19657 | 129  | 
(* name space and syntax *)  | 
| 19364 | 130  | 
|
| 56025 | 131  | 
fun space_of (Consts {decls, ...}) = Name_Space.space_of_table decls;
 | 
| 19364 | 132  | 
|
| 56025 | 133  | 
fun alias naming binding name = map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
134  | 
((Name_Space.alias_table naming binding name decls), constraints, rev_abbrevs));  | 
|
| 35680 | 135  | 
|
| 33158 | 136  | 
val is_concealed = Name_Space.is_concealed o space_of;  | 
137  | 
||
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
138  | 
val intern = Name_Space.intern o space_of;  | 
| 19364 | 139  | 
|
| 35554 | 140  | 
fun intern_syntax consts s =  | 
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42083 
diff
changeset
 | 
141  | 
(case try Lexicon.unmark_const s of  | 
| 35255 | 142  | 
SOME c => c  | 
| 35554 | 143  | 
| NONE => intern consts s);  | 
144  | 
||
| 18060 | 145  | 
|
| 
55843
 
3fa61f39d1f2
prefer Name_Space.check with its builtin reports (including completion);
 
wenzelm 
parents: 
50768 
diff
changeset
 | 
146  | 
(* check_const *)  | 
| 19364 | 147  | 
|
| 
55956
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
148  | 
fun check_const context consts (xname, ps) =  | 
| 19364 | 149  | 
let  | 
| 
55843
 
3fa61f39d1f2
prefer Name_Space.check with its builtin reports (including completion);
 
wenzelm 
parents: 
50768 
diff
changeset
 | 
150  | 
    val Consts {decls, ...} = consts;
 | 
| 
55956
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
151  | 
val ((c, reports), _) = Name_Space.check_reports context decls (xname, ps);  | 
| 
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
152  | 
val T = type_scheme consts c handle TYPE (msg, _, _) => error (msg ^ Position.here_list ps);  | 
| 
55950
 
3bb5c7179234
clarified treatment of consts -- prefer value-oriented reports;
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
153  | 
in (Const (c, T), reports) end;  | 
| 19364 | 154  | 
|
155  | 
||
156  | 
(* certify *)  | 
|
157  | 
||
| 
61262
 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 
wenzelm 
parents: 
56139 
diff
changeset
 | 
158  | 
fun certify context tsig do_expand consts =  | 
| 18965 | 159  | 
let  | 
160  | 
fun err msg (c, T) =  | 
|
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
161  | 
raise TYPE (msg ^ " " ^ quote c ^ " :: " ^  | 
| 
61262
 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 
wenzelm 
parents: 
56139 
diff
changeset
 | 
162  | 
Syntax.string_of_typ (Syntax.init_pretty context) T, [], []);  | 
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
163  | 
val certT = Type.cert_typ tsig;  | 
| 18965 | 164  | 
fun cert tm =  | 
165  | 
let  | 
|
166  | 
val (head, args) = Term.strip_comb tm;  | 
|
| 21694 | 167  | 
val args' = map cert args;  | 
168  | 
fun comb head' = Term.list_comb (head', args');  | 
|
| 18965 | 169  | 
in  | 
170  | 
(case head of  | 
|
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
171  | 
Abs (x, T, t) => comb (Abs (x, certT T, cert t))  | 
| 19364 | 172  | 
| Const (c, T) =>  | 
173  | 
let  | 
|
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
174  | 
val T' = certT T;  | 
| 
43794
 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
175  | 
              val (_, ({T = U, ...}, abbr)) = the_entry consts c;
 | 
| 25048 | 176  | 
fun expand u =  | 
177  | 
Term.betapplys (Envir.expand_atom T' (U, u) handle TYPE _ =>  | 
|
178  | 
err "Illegal type for abbreviation" (c, T), args');  | 
|
| 19364 | 179  | 
in  | 
| 
19433
 
c7a2b7a8c4cb
certify: ignore sort constraints of declarations (MAJOR CHANGE);
 
wenzelm 
parents: 
19364 
diff
changeset
 | 
180  | 
if not (Type.raw_instance (T', U)) then  | 
| 19364 | 181  | 
err "Illegal type for constant" (c, T)  | 
182  | 
else  | 
|
| 24909 | 183  | 
(case abbr of  | 
| 25048 | 184  | 
                  SOME {rhs, normal_rhs, force_expand} =>
 | 
185  | 
if do_expand then expand normal_rhs  | 
|
186  | 
else if force_expand then expand rhs  | 
|
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
187  | 
else comb head  | 
| 19364 | 188  | 
| _ => comb head)  | 
189  | 
end  | 
|
190  | 
| _ => comb head)  | 
|
| 18965 | 191  | 
end;  | 
192  | 
in cert end;  | 
|
193  | 
||
194  | 
||
| 18060 | 195  | 
(* typargs -- view actual const type as instance of declaration *)  | 
196  | 
||
| 24909 | 197  | 
local  | 
198  | 
||
199  | 
fun args_of (Type (_, Ts)) pos = args_of_list Ts 0 pos  | 
|
200  | 
| args_of (TVar v) pos = insert (eq_fst op =) (v, rev pos)  | 
|
201  | 
| args_of (TFree _) _ = I  | 
|
202  | 
and args_of_list (T :: Ts) i is = args_of T (i :: is) #> args_of_list Ts (i + 1) is  | 
|
203  | 
| args_of_list [] _ _ = I;  | 
|
204  | 
||
| 19364 | 205  | 
fun subscript (Type (_, Ts)) (i :: is) = subscript (nth Ts i) is  | 
206  | 
| subscript T [] = T  | 
|
| 32784 | 207  | 
| subscript _ _ = raise Subscript;  | 
| 18060 | 208  | 
|
| 24909 | 209  | 
in  | 
210  | 
||
211  | 
fun typargs_of T = map #2 (rev (args_of T [] []));  | 
|
212  | 
||
| 19364 | 213  | 
fun typargs consts (c, T) = map (subscript T) (type_arguments consts c);  | 
| 18060 | 214  | 
|
| 24909 | 215  | 
end;  | 
216  | 
||
| 18163 | 217  | 
fun instance consts (c, Ts) =  | 
218  | 
let  | 
|
| 25041 | 219  | 
val declT = type_scheme consts c;  | 
| 18163 | 220  | 
val vars = map Term.dest_TVar (typargs consts (c, declT));  | 
| 
40722
 
441260986b63
make two copies (!) of Library.UnequalLengths coincide with ListPair.UnequalLengths;
 
wenzelm 
parents: 
37146 
diff
changeset
 | 
221  | 
val inst = vars ~~ Ts handle ListPair.UnequalLengths =>  | 
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
222  | 
      raise TYPE ("Consts.instance", Ts, [Const (c, dummyT)]);
 | 
| 31977 | 223  | 
in declT |> Term_Subst.instantiateT inst end;  | 
| 18163 | 224  | 
|
| 18060 | 225  | 
|
226  | 
||
| 19364 | 227  | 
(** build consts **)  | 
| 18060 | 228  | 
|
| 18935 | 229  | 
(* name space *)  | 
| 18060 | 230  | 
|
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
231  | 
fun hide fully c = map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
| 56025 | 232  | 
(Name_Space.hide_table fully c decls, constraints, rev_abbrevs));  | 
| 18935 | 233  | 
|
234  | 
||
235  | 
(* declarations *)  | 
|
236  | 
||
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
237  | 
fun declare context (b, declT) =  | 
| 
33092
 
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
238  | 
map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
| 
 
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
239  | 
let  | 
| 35255 | 240  | 
      val decl = {T = declT, typargs = typargs_of declT};
 | 
| 
41254
 
78c3e472bb35
extra checking of name bindings for classes, types, consts;
 
wenzelm 
parents: 
40722 
diff
changeset
 | 
241  | 
val _ = Binding.check b;  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
242  | 
val (_, decls') = decls |> Name_Space.define context true (b, (decl, NONE));  | 
| 
33092
 
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
243  | 
in (decls', constraints, rev_abbrevs) end);  | 
| 19364 | 244  | 
|
245  | 
||
246  | 
(* constraints *)  | 
|
247  | 
||
248  | 
fun constrain (c, C) consts =  | 
|
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
249  | 
consts |> map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
| 
43794
 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
250  | 
(#2 (the_entry consts c) handle TYPE (msg, _, _) => error msg;  | 
| 19364 | 251  | 
(decls,  | 
252  | 
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
 | 
253  | 
rev_abbrevs)));  | 
| 18060 | 254  | 
|
| 18935 | 255  | 
|
256  | 
(* abbreviations *)  | 
|
257  | 
||
| 19027 | 258  | 
local  | 
259  | 
||
| 
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
 | 
260  | 
fun strip_abss (t as Abs (x, T, b)) =  | 
| 
42083
 
e1209fc7ecdc
added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
 
wenzelm 
parents: 
41254 
diff
changeset
 | 
261  | 
if Term.is_dependent b then strip_abss b |>> cons (x, T) (* FIXME decr!? *)  | 
| 
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
 | 
262  | 
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
 | 
263  | 
| strip_abss t = ([], t);  | 
| 19027 | 264  | 
|
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
265  | 
fun rev_abbrev lhs rhs =  | 
| 18060 | 266  | 
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
 | 
267  | 
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
 | 
268  | 
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
 | 
269  | 
in (Term.subst_bounds (rev vars, body), Term.list_comb (lhs, vars)) end;  | 
| 19027 | 270  | 
|
271  | 
in  | 
|
| 18965 | 272  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
273  | 
fun abbreviate context tsig mode (b, raw_rhs) consts =  | 
| 18965 | 274  | 
let  | 
| 
61262
 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 
wenzelm 
parents: 
56139 
diff
changeset
 | 
275  | 
val cert_term = certify context tsig false consts;  | 
| 
 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 
wenzelm 
parents: 
56139 
diff
changeset
 | 
276  | 
val expand_term = certify context 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
 | 
277  | 
val force_expand = mode = Print_Mode.internal;  | 
| 
21720
 
059e6b8cee8e
abbreviate: always authentic, force expansion of internal abbreviations;
 
wenzelm 
parents: 
21694 
diff
changeset
 | 
278  | 
|
| 
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
 | 
279  | 
val _ = Term.exists_subterm Term.is_Var raw_rhs andalso  | 
| 
42381
 
309ec68442c6
added Binding.print convenience, which includes quote already;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
280  | 
      error ("Illegal schematic variables on rhs of abbreviation " ^ Binding.print 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
 | 
281  | 
|
| 19364 | 282  | 
val rhs = raw_rhs  | 
| 
20548
 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 
wenzelm 
parents: 
20509 
diff
changeset
 | 
283  | 
|> 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
 | 
284  | 
|> 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
 | 
285  | 
|> Term.close_schematic_term;  | 
| 25404 | 286  | 
val normal_rhs = expand_term rhs;  | 
| 21794 | 287  | 
val T = Term.fastype_of rhs;  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
288  | 
val lhs = Const (Name_Space.full_name (Name_Space.naming_of context) b, T);  | 
| 19364 | 289  | 
in  | 
| 
24673
 
62b75508eb66
certify: do_expand as explicit argument, actually certify type of abstractions;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
290  | 
consts |> map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
| 19364 | 291  | 
let  | 
| 35255 | 292  | 
        val decl = {T = T, typargs = typargs_of T};
 | 
| 25404 | 293  | 
        val abbr = {rhs = rhs, normal_rhs = normal_rhs, force_expand = force_expand};
 | 
| 
41254
 
78c3e472bb35
extra checking of name bindings for classes, types, consts;
 
wenzelm 
parents: 
40722 
diff
changeset
 | 
294  | 
val _ = Binding.check b;  | 
| 28861 | 295  | 
val (_, decls') = decls  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
296  | 
|> Name_Space.define context true (b, (decl, SOME abbr));  | 
| 19364 | 297  | 
val rev_abbrevs' = rev_abbrevs  | 
| 33373 | 298  | 
|> 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
 | 
299  | 
in (decls', constraints, rev_abbrevs') end)  | 
| 21794 | 300  | 
|> pair (lhs, rhs)  | 
| 19364 | 301  | 
end;  | 
| 18935 | 302  | 
|
| 25048 | 303  | 
fun revert_abbrev mode c consts = consts |> map_consts (fn (decls, constraints, rev_abbrevs) =>  | 
304  | 
let  | 
|
305  | 
val (T, rhs) = the_abbreviation consts c;  | 
|
306  | 
val rev_abbrevs' = rev_abbrevs  | 
|
| 33373 | 307  | 
|> update_abbrevs mode (rev_abbrev (Const (c, T)) rhs);  | 
| 25048 | 308  | 
in (decls, constraints, rev_abbrevs') end);  | 
309  | 
||
| 19027 | 310  | 
end;  | 
311  | 
||
| 18060 | 312  | 
|
313  | 
(* empty and merge *)  | 
|
314  | 
||
| 45666 | 315  | 
val empty =  | 
| 
50201
 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
316  | 
make_consts (Name_Space.empty_table Markup.constantN, 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: 
30280 
diff
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: 
30280 
diff
changeset
 | 
320  | 
    Consts {decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}) =
 | 
| 18060 | 321  | 
let  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
322  | 
val decls' = Name_Space.merge_tables (decls1, decls2);  | 
| 
50768
 
2172f82de515
tuned -- prefer high-level Table.merge with its slightly more conservative update;
 
wenzelm 
parents: 
50201 
diff
changeset
 | 
323  | 
val constraints' = Symtab.merge (K true) (constraints1, constraints2);  | 
| 
30566
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
324  | 
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
 | 
325  | 
in make_consts (decls', constraints', rev_abbrevs') end;  | 
| 18060 | 326  | 
|
327  | 
end;  |