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
|
5829
|
12 |
val help: unit -> unit
|
5923
|
13 |
val load: string -> unit
|
5829
|
14 |
end;
|
|
15 |
|
|
16 |
signature OUTER_SYNTAX =
|
|
17 |
sig
|
|
18 |
include BASIC_OUTER_SYNTAX
|
|
19 |
type token
|
|
20 |
type parser
|
|
21 |
val parser: bool -> string -> string ->
|
|
22 |
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
|
5883
|
23 |
val print_outer_syntax: unit -> unit
|
5952
|
24 |
val commands: unit -> string list
|
5829
|
25 |
val add_keywords: string list -> unit
|
|
26 |
val add_parsers: parser list -> unit
|
|
27 |
val get_header: Position.T -> (string, 'a) Source.source -> (bstring * bstring list) option
|
|
28 |
val read_text: Position.T -> (string, 'a) Source.source -> Toplevel.transition list
|
|
29 |
val read_file: string -> Toplevel.transition list
|
|
30 |
val isar: Toplevel.isar
|
|
31 |
end;
|
|
32 |
|
|
33 |
structure OuterSyntax: OUTER_SYNTAX =
|
|
34 |
struct
|
|
35 |
|
|
36 |
open OuterParse;
|
|
37 |
|
|
38 |
|
|
39 |
(** outer syntax **)
|
|
40 |
|
|
41 |
(* parsers *)
|
|
42 |
|
|
43 |
type token = OuterLex.token;
|
|
44 |
type parser_fn = token list -> (Toplevel.transition -> Toplevel.transition) * token list;
|
|
45 |
|
|
46 |
datatype parser =
|
|
47 |
Parser of string * string * bool * parser_fn;
|
|
48 |
|
|
49 |
fun parser int_only name comment parse = Parser (name, comment, int_only, parse);
|
|
50 |
|
|
51 |
|
|
52 |
(* parse command *)
|
|
53 |
|
|
54 |
fun command_name cmd =
|
|
55 |
group "command"
|
|
56 |
(position (Scan.one (OuterLex.keyword_pred (is_some o cmd)) >> OuterLex.val_of));
|
|
57 |
|
|
58 |
fun command_body cmd (name, _) =
|
|
59 |
let val (int_only, parse) = the (cmd name)
|
|
60 |
in !!! (Scan.prompt (name ^ "# ") (parse >> pair int_only)) end;
|
|
61 |
|
|
62 |
fun command cmd =
|
|
63 |
$$$ ";" >> K None ||
|
|
64 |
command_name cmd :-- command_body cmd >> (fn ((name, pos), (int_only, f)) =>
|
|
65 |
Some (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>
|
|
66 |
Toplevel.interactive int_only |> f));
|
|
67 |
|
|
68 |
|
|
69 |
|
|
70 |
(** global syntax state **)
|
|
71 |
|
|
72 |
val global_lexicon = ref Scan.empty_lexicon;
|
|
73 |
val global_parsers = ref (Symtab.empty: (string * (bool * parser_fn)) Symtab.table);
|
|
74 |
|
5952
|
75 |
fun commands () = Symtab.keys (! global_parsers);
|
|
76 |
|
5829
|
77 |
|
|
78 |
(* print syntax *)
|
|
79 |
|
|
80 |
fun print_outer_syntax () =
|
|
81 |
let
|
|
82 |
val keywords = map implode (Scan.dest_lexicon (! global_lexicon));
|
|
83 |
fun pretty_cmd (name, (comment, _)) =
|
|
84 |
Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
|
|
85 |
val (int_cmds, cmds) = partition (#1 o #2 o #2) (Symtab.dest (! global_parsers));
|
|
86 |
in
|
|
87 |
Pretty.writeln (Pretty.strs ("keywords:" :: map quote keywords));
|
6095
|
88 |
Pretty.writeln (Pretty.big_list "proper commands:" (map pretty_cmd cmds));
|
|
89 |
Pretty.writeln (Pretty.big_list "improper commands (interactive-only):"
|
|
90 |
(map pretty_cmd int_cmds))
|
5829
|
91 |
end;
|
|
92 |
|
|
93 |
|
|
94 |
(* augment syntax *)
|
|
95 |
|
|
96 |
fun add_keywords keywords =
|
|
97 |
global_lexicon := Scan.extend_lexicon (! global_lexicon) (map Symbol.explode keywords);
|
|
98 |
|
|
99 |
fun add_parser (tab, Parser (name, comment, int_only, parse)) =
|
|
100 |
(if is_none (Symtab.lookup (tab, name)) then ()
|
|
101 |
else warning ("Redefined outer syntax command " ^ quote name);
|
|
102 |
Symtab.update ((name, (comment, (int_only, parse))), tab));
|
|
103 |
|
|
104 |
fun add_parsers parsers =
|
|
105 |
(global_parsers := foldl add_parser (! global_parsers, parsers);
|
|
106 |
add_keywords (map (fn Parser (name, _, _, _) => name) parsers));
|
|
107 |
|
|
108 |
|
|
109 |
(* get current lexer / parser *)
|
|
110 |
|
|
111 |
(*Note: the syntax for files is statically determined at the very
|
|
112 |
beginning; for interactive processing it may change dynamically.*)
|
|
113 |
|
|
114 |
fun get_lexicon () = ! global_lexicon;
|
|
115 |
fun get_parser () = apsome snd o curry Symtab.lookup (! global_parsers);
|
|
116 |
|
|
117 |
|
|
118 |
|
|
119 |
(** read theory **)
|
|
120 |
|
|
121 |
(* source *)
|
|
122 |
|
|
123 |
fun no_command cmd =
|
|
124 |
Scan.one ((not o OuterLex.keyword_pred ((is_some o cmd) orf equal ";")) andf OuterLex.not_eof);
|
|
125 |
|
|
126 |
fun recover cmd =
|
|
127 |
Scan.prompt "recover# " (Scan.one OuterLex.not_eof -- Scan.repeat (no_command cmd));
|
|
128 |
|
|
129 |
fun source do_recover cmd src =
|
|
130 |
src
|
|
131 |
|> Source.source OuterLex.stopper (Scan.bulk (fn xs => !!! (command (cmd ())) xs))
|
|
132 |
(if do_recover then Some (fn xs => recover (cmd ()) xs) else None)
|
|
133 |
|> Source.mapfilter I;
|
|
134 |
|
|
135 |
|
|
136 |
(* detect header *)
|
|
137 |
|
|
138 |
val head_lexicon =
|
|
139 |
Scan.make_lexicon (map Symbol.explode ["+", ":", "=", "theory"]);
|
|
140 |
|
|
141 |
val head_parser =
|
|
142 |
$$$ "theory" |-- !!! (name -- ($$$ "=" |-- enum "+" name --| (Scan.ahead eof || $$$ ":")));
|
|
143 |
|
|
144 |
fun get_header pos src =
|
|
145 |
src
|
|
146 |
|> Symbol.source false
|
|
147 |
|> OuterLex.source false (K head_lexicon) pos
|
|
148 |
|> Source.source OuterLex.stopper (Scan.single (Scan.option head_parser)) None
|
|
149 |
|> (fst o the o Source.get_single);
|
|
150 |
|
|
151 |
|
|
152 |
(* read text (including header) *)
|
|
153 |
|
|
154 |
fun read_text pos src =
|
|
155 |
src
|
|
156 |
|> Symbol.source false
|
|
157 |
|> OuterLex.source false (K (get_lexicon ())) pos
|
|
158 |
|> source false (K (get_parser ()))
|
|
159 |
|> Source.exhaust;
|
|
160 |
|
|
161 |
fun read_file name = read_text (Position.line_name 1 (quote name)) (Source.of_file name);
|
|
162 |
|
|
163 |
|
|
164 |
(* interactive source of state transformers *)
|
|
165 |
|
|
166 |
val isar =
|
|
167 |
Source.tty
|
|
168 |
|> Symbol.source true
|
|
169 |
|> OuterLex.source true get_lexicon (Position.line_name 1 "stdin")
|
|
170 |
|> source true get_parser;
|
|
171 |
|
|
172 |
|
|
173 |
|
|
174 |
(** the read-eval-print loop **)
|
|
175 |
|
5923
|
176 |
(* main loop *)
|
|
177 |
|
5883
|
178 |
fun loop () = (Context.reset_context (); Toplevel.loop isar);
|
5829
|
179 |
|
|
180 |
fun main () =
|
|
181 |
(Toplevel.set_state Toplevel.toplevel;
|
|
182 |
ml_prompts "ML> " "ML# ";
|
5992
|
183 |
writeln (Context.welcome ());
|
5883
|
184 |
loop ());
|
5829
|
185 |
|
|
186 |
|
5923
|
187 |
(* load *)
|
|
188 |
|
|
189 |
fun load name = Toplevel.excursion (read_file (name ^ ".thy"))
|
|
190 |
handle exn => error (Toplevel.exn_message exn);
|
|
191 |
|
|
192 |
|
5829
|
193 |
(* help *)
|
|
194 |
|
|
195 |
fun help () =
|
|
196 |
writeln ("This is Isabelle's underlying ML system (" ^ ml_system ^ ");\n\
|
5883
|
197 |
\invoke 'loop();' to enter the Isar loop.");
|
5829
|
198 |
|
|
199 |
|
|
200 |
end;
|
|
201 |
|
|
202 |
structure BasicOuterSyntax: BASIC_OUTER_SYNTAX = OuterSyntax;
|
|
203 |
open BasicOuterSyntax;
|