author | skalberg |
Thu, 03 Mar 2005 12:43:01 +0100 | |
changeset 15570 | 8d8c70b41bab |
parent 15531 | 08c8dad8e399 |
child 15830 | 74d8412b1a27 |
permissions | -rw-r--r-- |
5829 | 1 |
(* Title: Pure/Isar/outer_syntax.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
5 |
The global Isabelle/Isar outer syntax. |
|
6 |
*) |
|
7 |
||
8 |
signature BASIC_OUTER_SYNTAX = |
|
9 |
sig |
|
10 |
val main: unit -> unit |
|
5883 | 11 |
val loop: unit -> unit |
6860 | 12 |
val sync_main: unit -> unit |
13 |
val sync_loop: unit -> unit |
|
5829 | 14 |
end; |
15 |
||
16 |
signature OUTER_SYNTAX = |
|
17 |
sig |
|
18 |
include BASIC_OUTER_SYNTAX |
|
6722 | 19 |
structure Keyword: |
20 |
sig |
|
21 |
val control: string |
|
22 |
val diag: string |
|
23 |
val thy_begin: string |
|
7104 | 24 |
val thy_switch: string |
6722 | 25 |
val thy_end: string |
26 |
val thy_heading: string |
|
27 |
val thy_decl: string |
|
9588 | 28 |
val thy_script: string |
6722 | 29 |
val thy_goal: string |
30 |
val qed: string |
|
6733 | 31 |
val qed_block: string |
8209 | 32 |
val qed_global: string |
9552 | 33 |
val prf_heading: string |
6722 | 34 |
val prf_goal: string |
35 |
val prf_block: string |
|
9056 | 36 |
val prf_open: string |
37 |
val prf_close: string |
|
6722 | 38 |
val prf_chain: string |
39 |
val prf_decl: string |
|
6868 | 40 |
val prf_asm: string |
7676 | 41 |
val prf_asm_goal: string |
6722 | 42 |
val prf_script: string |
43 |
val kinds: string list |
|
44 |
end |
|
5829 | 45 |
type token |
46 |
type parser |
|
6722 | 47 |
val command: string -> string -> string -> |
6373 | 48 |
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser |
9132 | 49 |
val markup_command: IsarOutput.markup -> string -> string -> string -> |
7789 | 50 |
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser |
6722 | 51 |
val improper_command: string -> string -> string -> |
6373 | 52 |
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser |
14687 | 53 |
val is_keyword: string -> bool |
7026 | 54 |
val dest_keywords: unit -> string list |
55 |
val dest_parsers: unit -> (string * string * string * bool) list |
|
5883 | 56 |
val print_outer_syntax: unit -> unit |
9223 | 57 |
val print_commands: Toplevel.transition -> Toplevel.transition |
5829 | 58 |
val add_keywords: string list -> unit |
59 |
val add_parsers: parser list -> unit |
|
12943 | 60 |
val check_text: string * Position.T -> bool -> Toplevel.state -> unit |
7940 | 61 |
val deps_thy: string -> bool -> Path.T -> string list * Path.T list |
6199 | 62 |
val load_thy: string -> bool -> bool -> Path.T -> unit |
14091 | 63 |
val isar: bool -> bool -> unit Toplevel.isar |
15224
1bd35fd87963
Allow scanning to recover and reconstruct bad input
aspinall
parents:
15156
diff
changeset
|
64 |
val scan: string -> OuterLex.token list |
15144
85929e1b307d
Remove isar_readstring. Split read into scanner and parser.
aspinall
parents:
14981
diff
changeset
|
65 |
val read: OuterLex.token list -> |
85929e1b307d
Remove isar_readstring. Split read into scanner and parser.
aspinall
parents:
14981
diff
changeset
|
66 |
(string * OuterLex.token list * Toplevel.transition) list |
5829 | 67 |
end; |
68 |
||
15224
1bd35fd87963
Allow scanning to recover and reconstruct bad input
aspinall
parents:
15156
diff
changeset
|
69 |
structure OuterSyntax : OUTER_SYNTAX = |
5829 | 70 |
struct |
71 |
||
7750 | 72 |
structure T = OuterLex; |
6860 | 73 |
structure P = OuterParse; |
74 |
||
5829 | 75 |
|
76 |
(** outer syntax **) |
|
77 |
||
6722 | 78 |
(* command keyword classification *) |
79 |
||
80 |
structure Keyword = |
|
81 |
struct |
|
82 |
val control = "control"; |
|
83 |
val diag = "diag"; |
|
84 |
val thy_begin = "theory-begin"; |
|
7104 | 85 |
val thy_switch = "theory-switch"; |
6722 | 86 |
val thy_end = "theory-end"; |
87 |
val thy_heading = "theory-heading"; |
|
88 |
val thy_decl = "theory-decl"; |
|
9588 | 89 |
val thy_script = "theory-script"; |
6722 | 90 |
val thy_goal = "theory-goal"; |
91 |
val qed = "qed"; |
|
6733 | 92 |
val qed_block = "qed-block"; |
8209 | 93 |
val qed_global = "qed-global"; |
9552 | 94 |
val prf_heading = "proof-heading"; |
6722 | 95 |
val prf_goal = "proof-goal"; |
96 |
val prf_block = "proof-block"; |
|
9056 | 97 |
val prf_open = "proof-open"; |
98 |
val prf_close = "proof-close"; |
|
6722 | 99 |
val prf_chain = "proof-chain"; |
100 |
val prf_decl = "proof-decl"; |
|
6868 | 101 |
val prf_asm = "proof-asm"; |
7676 | 102 |
val prf_asm_goal = "proof-asm-goal"; |
6722 | 103 |
val prf_script = "proof-script"; |
104 |
||
9588 | 105 |
val kinds = [control, diag, thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script, |
106 |
thy_goal, qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open, prf_close, |
|
107 |
prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script]; |
|
6722 | 108 |
end; |
109 |
||
110 |
||
5829 | 111 |
(* parsers *) |
112 |
||
7750 | 113 |
type token = T.token; |
5829 | 114 |
type parser_fn = token list -> (Toplevel.transition -> Toplevel.transition) * token list; |
115 |
||
116 |
datatype parser = |
|
9132 | 117 |
Parser of string * (string * string * IsarOutput.markup option) * bool * parser_fn; |
5829 | 118 |
|
7750 | 119 |
fun parser int_only markup name comment kind parse = |
120 |
Parser (name, (comment, kind, markup), int_only, parse); |
|
5829 | 121 |
|
122 |
||
123 |
(* parse command *) |
|
124 |
||
6860 | 125 |
local |
6199 | 126 |
|
14925
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
127 |
fun terminate false = Scan.succeed () |
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
128 |
| terminate true = P.group "end of input" (Scan.option P.sync -- P.semicolon >> K ()); |
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
129 |
|
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
130 |
fun trace false parse = parse |
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
131 |
| trace true parse = Scan.trace parse >> (fn (f, toks) => f o Toplevel.source toks); |
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
132 |
|
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
133 |
fun body cmd trc (name, _) = |
7026 | 134 |
(case cmd name of |
15531 | 135 |
SOME (int_only, parse) => |
14925
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
136 |
P.!!! (Scan.prompt (name ^ "# ") (trace trc parse >> pair int_only)) |
15531 | 137 |
| NONE => sys_error ("no parser for outer syntax command " ^ quote name)); |
6860 | 138 |
|
139 |
in |
|
5829 | 140 |
|
14925
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
141 |
fun command do_terminate do_trace cmd = |
15531 | 142 |
P.semicolon >> K NONE || |
143 |
P.sync >> K NONE || |
|
14925
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
144 |
(P.position P.command :-- body cmd do_trace) --| terminate do_terminate |
6860 | 145 |
>> (fn ((name, pos), (int_only, f)) => |
15531 | 146 |
SOME (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |> |
6860 | 147 |
Toplevel.interactive int_only |> f)); |
5829 | 148 |
|
6199 | 149 |
end; |
150 |
||
5829 | 151 |
|
152 |
||
9132 | 153 |
(** global outer syntax **) |
5829 | 154 |
|
7026 | 155 |
local |
156 |
||
157 |
val global_lexicons = ref (Scan.empty_lexicon, Scan.empty_lexicon); |
|
7750 | 158 |
val global_parsers = |
9132 | 159 |
ref (Symtab.empty: (((string * string) * (bool * parser_fn)) * IsarOutput.markup option) |
160 |
Symtab.table); |
|
161 |
val global_markups = ref ([]: (string * IsarOutput.markup) list); |
|
5952 | 162 |
|
7026 | 163 |
fun change_lexicons f = |
164 |
let val lexs = f (! global_lexicons) in |
|
165 |
(case (op inter_string) (pairself Scan.dest_lexicon lexs) of |
|
166 |
[] => global_lexicons := lexs |
|
167 |
| bads => error ("Clash of outer syntax commands and keywords: " ^ commas_quote bads)) |
|
168 |
end; |
|
5829 | 169 |
|
15531 | 170 |
fun get_markup (ms, (name, (_, SOME m))) = (name, m) :: ms |
7789 | 171 |
| get_markup (ms, _) = ms; |
7750 | 172 |
|
173 |
fun make_markups () = global_markups := Symtab.foldl get_markup ([], ! global_parsers); |
|
9132 | 174 |
fun change_parsers f = (Library.change global_parsers f; make_markups ()); |
6722 | 175 |
|
7026 | 176 |
in |
177 |
||
7750 | 178 |
|
9132 | 179 |
(* access current syntax *) |
7026 | 180 |
|
181 |
(*Note: the syntax for files is statically determined at the very |
|
182 |
beginning; for interactive processing it may change dynamically.*) |
|
183 |
||
184 |
fun get_lexicons () = ! global_lexicons; |
|
185 |
fun get_parsers () = ! global_parsers; |
|
15570 | 186 |
fun get_parser () = Option.map (#2 o #1) o curry Symtab.lookup (! global_parsers); |
7789 | 187 |
|
9132 | 188 |
fun is_markup kind name = |
15531 | 189 |
(case assoc (! global_markups, name) of SOME k => k = kind | NONE => false); |
9132 | 190 |
fun markup kind = Scan.one (T.is_kind T.Command andf is_markup kind o T.val_of); |
5829 | 191 |
|
192 |
||
193 |
(* augment syntax *) |
|
194 |
||
7026 | 195 |
fun add_keywords keywords = change_lexicons (apfst (fn lex => |
196 |
(Scan.extend_lexicon lex (map Symbol.explode keywords)))); |
|
5829 | 197 |
|
7750 | 198 |
fun add_parser (tab, Parser (name, (comment, kind, markup), int_only, parse)) = |
5829 | 199 |
(if is_none (Symtab.lookup (tab, name)) then () |
200 |
else warning ("Redefined outer syntax command " ^ quote name); |
|
7750 | 201 |
Symtab.update ((name, (((comment, kind), (int_only, parse)), markup)), tab)); |
5829 | 202 |
|
203 |
fun add_parsers parsers = |
|
15570 | 204 |
(change_parsers (fn tab => Library.foldl add_parser (tab, parsers)); |
7026 | 205 |
change_lexicons (apsnd (fn lex => Scan.extend_lexicon lex |
206 |
(map (fn Parser (name, _, _, _) => Symbol.explode name) parsers)))); |
|
207 |
||
208 |
end; |
|
5829 | 209 |
|
210 |
||
7026 | 211 |
(* print syntax *) |
212 |
||
14687 | 213 |
fun is_keyword s = Scan.is_literal (#1 (get_lexicons ())) (Symbol.explode s); |
7026 | 214 |
fun dest_keywords () = Scan.dest_lexicon (#1 (get_lexicons ())); |
215 |
||
216 |
fun dest_parsers () = |
|
7750 | 217 |
map (fn (name, (((cmt, kind), (int_only, _)), _)) => (name, cmt, kind, int_only)) |
7026 | 218 |
(Symtab.dest (get_parsers ())); |
5829 | 219 |
|
9223 | 220 |
fun print_outer_syntax () = |
7026 | 221 |
let |
222 |
fun pretty_cmd (name, comment, _, _) = |
|
223 |
Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; |
|
15570 | 224 |
val (int_cmds, cmds) = List.partition #4 (dest_parsers ()); |
7026 | 225 |
in |
8720 | 226 |
[Pretty.strs ("syntax keywords:" :: map quote (dest_keywords ())), |
227 |
Pretty.big_list "proper commands:" (map pretty_cmd cmds), |
|
228 |
Pretty.big_list "improper commands (interactive-only):" (map pretty_cmd int_cmds)] |
|
9223 | 229 |
|> Pretty.chunks |> Pretty.writeln |
7026 | 230 |
end; |
5829 | 231 |
|
9223 | 232 |
val print_commands = Toplevel.imperative print_outer_syntax; |
7367 | 233 |
|
5829 | 234 |
|
235 |
||
9132 | 236 |
(** toplevel parsing **) |
5829 | 237 |
|
9132 | 238 |
(* basic sources *) |
6860 | 239 |
|
14925
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
240 |
fun toplevel_source term trc do_recover cmd src = |
9132 | 241 |
let |
242 |
val no_terminator = |
|
243 |
Scan.unless P.semicolon (Scan.one (T.not_sync andf T.not_eof)); |
|
15531 | 244 |
fun recover x = (Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [NONE]) x; |
9132 | 245 |
in |
246 |
src |
|
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
247 |
|> T.source_proper |
9132 | 248 |
|> Source.source T.stopper |
15531 | 249 |
(Scan.bulk (P.$$$ "--" -- P.!!! P.text >> K NONE || P.not_eof >> SOME)) |
250 |
(if do_recover then SOME recover else NONE) |
|
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
251 |
|> Source.mapfilter I |
14925
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
252 |
|> Source.source T.stopper (Scan.bulk (fn xs => P.!!! (command term trc (cmd ())) xs)) |
15531 | 253 |
(if do_recover then SOME recover else NONE) |
9132 | 254 |
|> Source.mapfilter I |
255 |
end; |
|
5829 | 256 |
|
7746 | 257 |
|
9132 | 258 |
(* interactive source of toplevel transformers *) |
5829 | 259 |
|
9132 | 260 |
fun isar term no_pos = |
261 |
Source.tty |
|
262 |
|> Symbol.source true |
|
263 |
|> T.source true get_lexicons |
|
264 |
(if no_pos then Position.none else Position.line_name 1 "stdin") |
|
14925
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
265 |
|> toplevel_source term false true get_parser; |
6199 | 266 |
|
267 |
||
15144
85929e1b307d
Remove isar_readstring. Split read into scanner and parser.
aspinall
parents:
14981
diff
changeset
|
268 |
(* scan text, read tokens with trace (for Proof General) *) |
14925
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
269 |
|
15144
85929e1b307d
Remove isar_readstring. Split read into scanner and parser.
aspinall
parents:
14981
diff
changeset
|
270 |
fun scan str = |
85929e1b307d
Remove isar_readstring. Split read into scanner and parser.
aspinall
parents:
14981
diff
changeset
|
271 |
Source.of_string str |
85929e1b307d
Remove isar_readstring. Split read into scanner and parser.
aspinall
parents:
14981
diff
changeset
|
272 |
|> Symbol.source false |
15224
1bd35fd87963
Allow scanning to recover and reconstruct bad input
aspinall
parents:
15156
diff
changeset
|
273 |
|> T.source true get_lexicons Position.none |
15144
85929e1b307d
Remove isar_readstring. Split read into scanner and parser.
aspinall
parents:
14981
diff
changeset
|
274 |
|> Source.exhaust |
14925
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
275 |
|
15144
85929e1b307d
Remove isar_readstring. Split read into scanner and parser.
aspinall
parents:
14981
diff
changeset
|
276 |
fun read toks = |
85929e1b307d
Remove isar_readstring. Split read into scanner and parser.
aspinall
parents:
14981
diff
changeset
|
277 |
Source.of_list toks |
14925
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
278 |
|> toplevel_source false true true get_parser |
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
279 |
|> Source.exhaust |
15570 | 280 |
|> map (fn tr => (Toplevel.name_of tr, valOf (Toplevel.source_of tr), tr)); |
14091 | 281 |
|
282 |
||
9132 | 283 |
(** read theory **) |
6247 | 284 |
|
12943 | 285 |
(* check_text *) |
286 |
||
287 |
fun check_text s true state = (IsarOutput.eval_antiquote (#1 (get_lexicons ())) state s; ()) |
|
288 |
| check_text _ false _ = (); |
|
289 |
||
290 |
||
9132 | 291 |
(* deps_thy *) |
6199 | 292 |
|
7940 | 293 |
fun deps_thy name ml path = |
6199 | 294 |
let |
7735 | 295 |
val src = Source.of_string (File.read path); |
296 |
val pos = Path.position path; |
|
9132 | 297 |
val (name', parents, files) = ThyHeader.scan (src, pos); |
6199 | 298 |
val ml_path = ThyLoad.ml_path name; |
15570 | 299 |
val ml_file = if ml andalso isSome (ThyLoad.check_file NONE ml_path) then [ml_path] else []; |
6199 | 300 |
in |
301 |
if name <> name' then |
|
7940 | 302 |
error ("Filename " ^ quote (Path.pack path) ^ |
8078 | 303 |
" does not agree with theory name " ^ quote name') |
6247 | 304 |
else (parents, map (Path.unpack o #1) files @ ml_file) |
6199 | 305 |
end; |
306 |
||
7746 | 307 |
|
9132 | 308 |
(* load_thy *) |
6199 | 309 |
|
7746 | 310 |
local |
311 |
||
7940 | 312 |
fun try_ml_file name time = |
6199 | 313 |
let |
314 |
val path = ThyLoad.ml_path name; |
|
7940 | 315 |
val tr = Toplevel.imperative (fn () => ThyInfo.load_file time path); |
6247 | 316 |
val tr_name = if time then "time_use" else "use"; |
6199 | 317 |
in |
15531 | 318 |
if is_none (ThyLoad.check_file NONE path) then () |
9132 | 319 |
else Toplevel.excursion [Toplevel.empty |> Toplevel.name tr_name |> tr] |
6199 | 320 |
end; |
321 |
||
9132 | 322 |
fun parse_thy src = |
323 |
src |
|
14925
0f86a8a694f8
added read (provides transition names and sources);
wenzelm
parents:
14687
diff
changeset
|
324 |
|> toplevel_source false false false (K (get_parser ())) |
7746 | 325 |
|> Source.exhaust; |
5829 | 326 |
|
6247 | 327 |
fun run_thy name path = |
7683 | 328 |
let |
7735 | 329 |
val pos = Path.position path; |
9132 | 330 |
val text = Library.untabify (explode (File.read path)); |
331 |
val text_src = Source.of_list text; |
|
332 |
fun present_text () = Source.exhaust (Symbol.source false text_src); |
|
7683 | 333 |
in |
7735 | 334 |
Present.init_theory name; |
9132 | 335 |
Present.verbatim_source name present_text; |
336 |
if ThyHeader.is_old (text_src, pos) then (ThySyn.load_thy name text; |
|
337 |
Present.old_symbol_source name present_text) (*note: text presented twice*) |
|
338 |
else |
|
339 |
let |
|
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
340 |
val tok_src = text_src |
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
341 |
|> Symbol.source false |
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
342 |
|> T.source false (K (get_lexicons ())) pos |
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
343 |
|> Source.exhausted; |
9132 | 344 |
val out = Toplevel.excursion_result |
345 |
(IsarOutput.parse_thy markup (#1 (get_lexicons ())) |
|
346 |
(parse_thy tok_src) tok_src); |
|
347 |
in Present.theory_output name (Buffer.content out) end |
|
6247 | 348 |
end; |
6199 | 349 |
|
7746 | 350 |
in |
351 |
||
6199 | 352 |
fun load_thy name ml time path = |
6247 | 353 |
(if time then |
354 |
timeit (fn () => |
|
355 |
(writeln ("\n**** Starting theory " ^ quote name ^ " ****"); |
|
9036 | 356 |
run_thy name path; |
6247 | 357 |
writeln ("**** Finished theory " ^ quote name ^ " ****\n"))) |
358 |
else run_thy name path; |
|
359 |
Context.context (ThyInfo.get_theory name); |
|
7940 | 360 |
if ml then try_ml_file name time else ()); |
5829 | 361 |
|
7746 | 362 |
end; |
363 |
||
5829 | 364 |
|
365 |
||
366 |
(** the read-eval-print loop **) |
|
367 |
||
5923 | 368 |
(* main loop *) |
369 |
||
7333 | 370 |
fun gen_loop term no_pos = |
371 |
(Context.reset_context (); |
|
372 |
Toplevel.loop (isar term no_pos)); |
|
5829 | 373 |
|
7333 | 374 |
fun gen_main term no_pos = |
5829 | 375 |
(Toplevel.set_state Toplevel.toplevel; |
6199 | 376 |
writeln (Session.welcome ()); |
7333 | 377 |
gen_loop term no_pos); |
6860 | 378 |
|
7333 | 379 |
fun main () = gen_main false false; |
380 |
fun loop () = gen_loop false false; |
|
381 |
fun sync_main () = gen_main true true; |
|
382 |
fun sync_loop () = gen_loop true true; |
|
5829 | 383 |
|
384 |
||
6373 | 385 |
(*final declarations of this structure!*) |
15531 | 386 |
val command = parser false NONE; |
387 |
val markup_command = parser false o SOME; |
|
388 |
val improper_command = parser true NONE; |
|
6685 | 389 |
|
6373 | 390 |
|
5829 | 391 |
end; |
392 |
||
6199 | 393 |
(*setup theory syntax dependent operations*) |
394 |
ThyLoad.deps_thy_fn := OuterSyntax.deps_thy; |
|
395 |
ThyLoad.load_thy_fn := OuterSyntax.load_thy; |
|
396 |
structure ThyLoad: THY_LOAD = ThyLoad; |
|
397 |
||
5829 | 398 |
structure BasicOuterSyntax: BASIC_OUTER_SYNTAX = OuterSyntax; |
399 |
open BasicOuterSyntax; |