author | wenzelm |
Sun, 07 Jan 2018 16:55:45 +0100 | |
changeset 67362 | 221612c942de |
parent 64556 | 851ae0e7b09c |
child 68276 | cbee43ff4ceb |
permissions | -rw-r--r-- |
31325 | 1 |
(* Title: Pure/ML/ml_env.ML |
2 |
Author: Makarius |
|
3 |
||
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
4 |
Toplevel environment for Standard ML and Isabelle/ML within the |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
5 |
implicit context. |
31325 | 6 |
*) |
7 |
||
8 |
signature ML_ENV = |
|
9 |
sig |
|
10 |
val inherit: Context.generic -> Context.generic -> Context.generic |
|
59127
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents:
59058
diff
changeset
|
11 |
val forget_structure: string -> Context.generic -> Context.generic |
60746 | 12 |
val get_breakpoint: Context.generic -> serial -> (bool Unsynchronized.ref * Position.T) option |
62941 | 13 |
val add_breakpoints: (serial * (bool Unsynchronized.ref * Thread_Position.T)) list -> unit |
62873
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
14 |
val init_bootstrap: Context.generic -> Context.generic |
62902
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
15 |
val SML_environment: bool Config.T |
60858 | 16 |
val add_name_space: {SML: bool} -> ML_Name_Space.T -> Context.generic -> Context.generic |
62495 | 17 |
val make_name_space: {SML: bool, exchange: bool} -> ML_Name_Space.T |
18 |
val context: ML_Compiler0.context |
|
19 |
val name_space: ML_Name_Space.T |
|
36163
823c9400eb62
proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents:
33519
diff
changeset
|
20 |
val check_functor: string -> unit |
31325 | 21 |
end |
22 |
||
23 |
structure ML_Env: ML_ENV = |
|
24 |
struct |
|
25 |
||
31328 | 26 |
(* context data *) |
27 |
||
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
28 |
type tables = |
62934 | 29 |
PolyML.NameSpace.Values.value Symtab.table * |
30 |
PolyML.NameSpace.TypeConstrs.typeConstr Symtab.table * |
|
31 |
PolyML.NameSpace.Infixes.fixity Symtab.table * |
|
32 |
PolyML.NameSpace.Structures.structureVal Symtab.table * |
|
33 |
PolyML.NameSpace.Signatures.signatureVal Symtab.table * |
|
34 |
PolyML.NameSpace.Functors.functorVal Symtab.table; |
|
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
35 |
|
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
36 |
fun merge_tables |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
37 |
((val1, type1, fixity1, structure1, signature1, functor1), |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
38 |
(val2, type2, fixity2, structure2, signature2, functor2)) : tables = |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
39 |
(Symtab.merge (K true) (val1, val2), |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
40 |
Symtab.merge (K true) (type1, type2), |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
41 |
Symtab.merge (K true) (fixity1, fixity2), |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
42 |
Symtab.merge (K true) (structure1, structure2), |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
43 |
Symtab.merge (K true) (signature1, signature2), |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
44 |
Symtab.merge (K true) (functor1, functor2)); |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
45 |
|
60746 | 46 |
type data = |
47 |
{bootstrap: bool, |
|
48 |
tables: tables, |
|
49 |
sml_tables: tables, |
|
50 |
breakpoints: (bool Unsynchronized.ref * Position.T) Inttab.table}; |
|
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
51 |
|
60746 | 52 |
fun make_data (bootstrap, tables, sml_tables, breakpoints) : data = |
53 |
{bootstrap = bootstrap, tables = tables, sml_tables = sml_tables, breakpoints = breakpoints}; |
|
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
54 |
|
33519 | 55 |
structure Env = Generic_Data |
31325 | 56 |
( |
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
57 |
type T = data |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
58 |
val empty = |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
59 |
make_data (true, |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
60 |
(Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty), |
62839
ea9f12e422c7
clarified SML name space: no access to structure PolyML;
wenzelm
parents:
62716
diff
changeset
|
61 |
(Symtab.make ML_Name_Space.sml_val, |
ea9f12e422c7
clarified SML name space: no access to structure PolyML;
wenzelm
parents:
62716
diff
changeset
|
62 |
Symtab.make ML_Name_Space.sml_type, |
ea9f12e422c7
clarified SML name space: no access to structure PolyML;
wenzelm
parents:
62716
diff
changeset
|
63 |
Symtab.make ML_Name_Space.sml_fixity, |
ea9f12e422c7
clarified SML name space: no access to structure PolyML;
wenzelm
parents:
62716
diff
changeset
|
64 |
Symtab.make ML_Name_Space.sml_structure, |
ea9f12e422c7
clarified SML name space: no access to structure PolyML;
wenzelm
parents:
62716
diff
changeset
|
65 |
Symtab.make ML_Name_Space.sml_signature, |
ea9f12e422c7
clarified SML name space: no access to structure PolyML;
wenzelm
parents:
62716
diff
changeset
|
66 |
Symtab.make ML_Name_Space.sml_functor), |
60746 | 67 |
Inttab.empty); |
68 |
fun extend (data : T) = make_data (false, #tables data, #sml_tables data, #breakpoints data); |
|
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
69 |
fun merge (data : T * T) = |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
70 |
make_data (false, |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
56618
diff
changeset
|
71 |
merge_tables (apply2 #tables data), |
60746 | 72 |
merge_tables (apply2 #sml_tables data), |
73 |
Inttab.merge (K true) (apply2 #breakpoints data)); |
|
31325 | 74 |
); |
75 |
||
76 |
val inherit = Env.put o Env.get; |
|
77 |
||
59127
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents:
59058
diff
changeset
|
78 |
fun forget_structure name = |
60746 | 79 |
Env.map (fn {bootstrap, tables, sml_tables, breakpoints} => |
59127
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents:
59058
diff
changeset
|
80 |
let |
62657 | 81 |
val _ = if bootstrap then ML_Name_Space.forget_structure name else (); |
59127
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents:
59058
diff
changeset
|
82 |
val tables' = |
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents:
59058
diff
changeset
|
83 |
(#1 tables, #2 tables, #3 tables, Symtab.delete_safe name (#4 tables), #5 tables, #6 tables); |
60746 | 84 |
in make_data (bootstrap, tables', sml_tables, breakpoints) end); |
85 |
||
62941 | 86 |
val get_breakpoint = Inttab.lookup o #breakpoints o Env.get; |
60746 | 87 |
|
62941 | 88 |
fun add_breakpoints more_breakpoints = |
89 |
if is_some (Context.get_generic_context ()) then |
|
90 |
Context.>> |
|
91 |
(Env.map (fn {bootstrap, tables, sml_tables, breakpoints} => |
|
92 |
let val breakpoints' = |
|
93 |
fold (Inttab.update_new o (apsnd o apsnd) Position.make) more_breakpoints breakpoints; |
|
94 |
in make_data (bootstrap, tables, sml_tables, breakpoints') end)) |
|
95 |
else (); |
|
59127
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents:
59058
diff
changeset
|
96 |
|
31328 | 97 |
|
62902
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
98 |
(* SML environment for Isabelle/ML *) |
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
99 |
|
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
100 |
val SML_environment = |
64556 | 101 |
Config.bool (Config.declare ("SML_environment", \<^here>) (fn _ => Config.Bool false)); |
62902
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
102 |
|
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
103 |
fun sml_env SML = |
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
104 |
SML orelse |
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
105 |
(case Context.get_generic_context () of |
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
106 |
NONE => false |
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
107 |
| SOME context => Config.get_generic context SML_environment); |
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
108 |
|
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
109 |
|
56618
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56275
diff
changeset
|
110 |
(* name space *) |
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
111 |
|
62873
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
112 |
val init_bootstrap = |
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
113 |
Env.map (fn {bootstrap, tables, sml_tables, breakpoints} => |
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
114 |
let |
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
115 |
val sml_tables' = |
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
116 |
sml_tables |> (fn (val1, type1, fixity1, structure1, signature1, functor1) => |
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
117 |
let |
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
118 |
val val2 = |
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
119 |
fold (fn (x, y) => |
62875 | 120 |
member (op =) ML_Name_Space.bootstrap_values x ? Symtab.update (x, y)) |
62873
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
121 |
(#allVal ML_Name_Space.global ()) val1; |
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
122 |
val structure2 = |
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
123 |
fold (fn (x, y) => |
62875 | 124 |
member (op =) ML_Name_Space.bootstrap_structures x ? Symtab.update (x, y)) |
62873
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
125 |
(#allStruct ML_Name_Space.global ()) structure1; |
62910
f37878ebba65
explicit handling of recursive ML name space, e.g. relevant for ML_Bootstrap;
wenzelm
parents:
62902
diff
changeset
|
126 |
val signature2 = |
f37878ebba65
explicit handling of recursive ML name space, e.g. relevant for ML_Bootstrap;
wenzelm
parents:
62902
diff
changeset
|
127 |
fold (fn (x, y) => |
f37878ebba65
explicit handling of recursive ML name space, e.g. relevant for ML_Bootstrap;
wenzelm
parents:
62902
diff
changeset
|
128 |
member (op =) ML_Name_Space.bootstrap_signatures x ? Symtab.update (x, y)) |
f37878ebba65
explicit handling of recursive ML name space, e.g. relevant for ML_Bootstrap;
wenzelm
parents:
62902
diff
changeset
|
129 |
(#allSig ML_Name_Space.global ()) signature1; |
f37878ebba65
explicit handling of recursive ML name space, e.g. relevant for ML_Bootstrap;
wenzelm
parents:
62902
diff
changeset
|
130 |
in (val2, type1, fixity1, structure2, signature2, functor1) end); |
62873
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
131 |
in make_data (bootstrap, tables, sml_tables', breakpoints) end); |
2f9c8a18f832
support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents:
62839
diff
changeset
|
132 |
|
60858 | 133 |
fun add_name_space {SML} (space: ML_Name_Space.T) = |
134 |
Env.map (fn {bootstrap, tables, sml_tables, breakpoints} => |
|
135 |
let |
|
136 |
val (tables', sml_tables') = |
|
62902
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
137 |
(tables, sml_tables) |> (if sml_env SML then apsnd else apfst) |
60858 | 138 |
(fn (val1, type1, fixity1, structure1, signature1, functor1) => |
139 |
let |
|
140 |
val val2 = fold Symtab.update (#allVal space ()) val1; |
|
141 |
val type2 = fold Symtab.update (#allType space ()) type1; |
|
142 |
val fixity2 = fold Symtab.update (#allFix space ()) fixity1; |
|
143 |
val structure2 = fold Symtab.update (#allStruct space ()) structure1; |
|
144 |
val signature2 = fold Symtab.update (#allSig space ()) signature1; |
|
145 |
val functor2 = fold Symtab.update (#allFunct space ()) functor1; |
|
146 |
in (val2, type2, fixity2, structure2, signature2, functor2) end); |
|
147 |
in make_data (bootstrap, tables', sml_tables', breakpoints) end); |
|
148 |
||
62495 | 149 |
fun make_name_space {SML, exchange} : ML_Name_Space.T = |
31325 | 150 |
let |
151 |
fun lookup sel1 sel2 name = |
|
62902
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
152 |
if sml_env SML then |
62876 | 153 |
Context.the_generic_context () |
56618
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56275
diff
changeset
|
154 |
|> (fn context => Symtab.lookup (sel1 (#sml_tables (Env.get context))) name) |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56275
diff
changeset
|
155 |
else |
62889 | 156 |
Context.get_generic_context () |
56618
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56275
diff
changeset
|
157 |
|> (fn NONE => NONE | SOME context => Symtab.lookup (sel1 (#tables (Env.get context))) name) |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56275
diff
changeset
|
158 |
|> (fn NONE => sel2 ML_Name_Space.global name | some => some); |
31325 | 159 |
|
160 |
fun all sel1 sel2 () = |
|
62902
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
161 |
(if sml_env SML then |
62876 | 162 |
Context.the_generic_context () |
56618
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56275
diff
changeset
|
163 |
|> (fn context => Symtab.dest (sel1 (#sml_tables (Env.get context)))) |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56275
diff
changeset
|
164 |
else |
62889 | 165 |
Context.get_generic_context () |
56618
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56275
diff
changeset
|
166 |
|> (fn NONE => [] | SOME context => Symtab.dest (sel1 (#tables (Env.get context)))) |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56275
diff
changeset
|
167 |
|> append (sel2 ML_Name_Space.global ())) |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
56618
diff
changeset
|
168 |
|> sort_distinct (string_ord o apply2 #1); |
31325 | 169 |
|
170 |
fun enter ap1 sel2 entry = |
|
62902
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
171 |
if sml_env SML <> exchange then |
60746 | 172 |
Context.>> (Env.map (fn {bootstrap, tables, sml_tables, breakpoints} => |
56618
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56275
diff
changeset
|
173 |
let val sml_tables' = ap1 (Symtab.update entry) sml_tables |
60746 | 174 |
in make_data (bootstrap, tables, sml_tables', breakpoints) end)) |
62889 | 175 |
else if is_some (Context.get_generic_context ()) then |
60746 | 176 |
Context.>> (Env.map (fn {bootstrap, tables, sml_tables, breakpoints} => |
56203
76c72f4d0667
clarified bootstrap process: switch to ML with context and antiquotations earlier;
wenzelm
parents:
48992
diff
changeset
|
177 |
let |
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
178 |
val _ = if bootstrap then sel2 ML_Name_Space.global entry else (); |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56203
diff
changeset
|
179 |
val tables' = ap1 (Symtab.update entry) tables; |
60746 | 180 |
in make_data (bootstrap, tables', sml_tables, breakpoints) end)) |
31325 | 181 |
else sel2 ML_Name_Space.global entry; |
182 |
in |
|
183 |
{lookupVal = lookup #1 #lookupVal, |
|
184 |
lookupType = lookup #2 #lookupType, |
|
185 |
lookupFix = lookup #3 #lookupFix, |
|
186 |
lookupStruct = lookup #4 #lookupStruct, |
|
187 |
lookupSig = lookup #5 #lookupSig, |
|
188 |
lookupFunct = lookup #6 #lookupFunct, |
|
189 |
enterVal = enter (fn h => fn (a, b, c, d, e, f) => (h a, b, c, d, e, f)) #enterVal, |
|
190 |
enterType = enter (fn h => fn (a, b, c, d, e, f) => (a, h b, c, d, e, f)) #enterType, |
|
191 |
enterFix = enter (fn h => fn (a, b, c, d, e, f) => (a, b, h c, d, e, f)) #enterFix, |
|
192 |
enterStruct = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, h d, e, f)) #enterStruct, |
|
193 |
enterSig = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, h e, f)) #enterSig, |
|
194 |
enterFunct = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, e, h f)) #enterFunct, |
|
195 |
allVal = all #1 #allVal, |
|
196 |
allType = all #2 #allType, |
|
197 |
allFix = all #3 #allFix, |
|
198 |
allStruct = all #4 #allStruct, |
|
199 |
allSig = all #5 #allSig, |
|
200 |
allFunct = all #6 #allFunct} |
|
201 |
end; |
|
202 |
||
62495 | 203 |
val context: ML_Compiler0.context = |
204 |
{name_space = make_name_space {SML = false, exchange = false}, |
|
62716
d80b9f4990e4
explicit print_depth for the sake of Spec_Check.determine_type;
wenzelm
parents:
62657
diff
changeset
|
205 |
print_depth = NONE, |
62493 | 206 |
here = Position.here oo Position.line_file, |
31325 | 207 |
print = writeln, |
208 |
error = error}; |
|
209 |
||
62495 | 210 |
val name_space = #name_space context; |
56618
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56275
diff
changeset
|
211 |
|
62495 | 212 |
val is_functor = is_some o #lookupFunct name_space; |
36163
823c9400eb62
proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents:
33519
diff
changeset
|
213 |
|
823c9400eb62
proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents:
33519
diff
changeset
|
214 |
fun check_functor name = |
36165 | 215 |
if not (is_functor "Table") (*mask dummy version of name_space*) orelse is_functor name then () |
36163
823c9400eb62
proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents:
33519
diff
changeset
|
216 |
else error ("Unknown ML functor: " ^ quote name); |
823c9400eb62
proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents:
33519
diff
changeset
|
217 |
|
31325 | 218 |
end; |