| author | wenzelm | 
| Fri, 05 Apr 2013 20:54:55 +0200 | |
| changeset 51627 | 589daaf48dba | 
| parent 51271 | e8d2ecf6aaac | 
| child 52509 | 2193d2c7f586 | 
| permissions | -rw-r--r-- | 
| 5829 | 1  | 
(* Title: Pure/Isar/outer_syntax.ML  | 
2  | 
Author: Markus Wenzel, TU Muenchen  | 
|
3  | 
||
| 
27353
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
26990 
diff
changeset
 | 
4  | 
The global Isabelle/Isar outer syntax.  | 
| 
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
26990 
diff
changeset
 | 
5  | 
|
| 
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
26990 
diff
changeset
 | 
6  | 
Note: the syntax for files is statically determined at the very  | 
| 
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
26990 
diff
changeset
 | 
7  | 
beginning; for interactive processing it may change dynamically.  | 
| 5829 | 8  | 
*)  | 
9  | 
||
10  | 
signature OUTER_SYNTAX =  | 
|
11  | 
sig  | 
|
| 43711 | 12  | 
type outer_syntax  | 
| 
43712
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43711 
diff
changeset
 | 
13  | 
val is_markup: outer_syntax -> Thy_Output.markup -> string -> bool  | 
| 43711 | 14  | 
val get_syntax: unit -> (Scan.lexicon * Scan.lexicon) * outer_syntax  | 
| 
46970
 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
15  | 
val check_syntax: unit -> unit  | 
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
16  | 
type command_spec = (string * Keyword.T) * Position.T  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
17  | 
val command: command_spec -> string ->  | 
| 29311 | 18  | 
(Toplevel.transition -> Toplevel.transition) parser -> unit  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
19  | 
val markup_command: Thy_Output.markup -> command_spec -> string ->  | 
| 29311 | 20  | 
(Toplevel.transition -> Toplevel.transition) parser -> unit  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
21  | 
val improper_command: command_spec -> string ->  | 
| 43711 | 22  | 
(Toplevel.transition -> Toplevel.transition) parser -> unit  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
23  | 
val local_theory': command_spec -> string ->  | 
| 29380 | 24  | 
(bool -> local_theory -> local_theory) parser -> unit  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
25  | 
val local_theory: command_spec -> string ->  | 
| 29311 | 26  | 
(local_theory -> local_theory) parser -> unit  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
27  | 
val local_theory_to_proof': command_spec -> string ->  | 
| 29311 | 28  | 
(bool -> local_theory -> Proof.state) parser -> unit  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
29  | 
val local_theory_to_proof: command_spec -> string ->  | 
| 29311 | 30  | 
(local_theory -> Proof.state) parser -> unit  | 
| 50213 | 31  | 
val help_outer_syntax: string list -> unit  | 
| 5883 | 32  | 
val print_outer_syntax: unit -> unit  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
33  | 
val scan: Position.T -> string -> Token.T list  | 
| 25580 | 34  | 
val parse: Position.T -> string -> Toplevel.transition list  | 
| 26600 | 35  | 
type isar  | 
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
36  | 
val isar: TextIO.instream -> bool -> isar  | 
| 
48918
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
37  | 
val span_cmts: Token.T list -> Token.T list  | 
| 
51268
 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 
wenzelm 
parents: 
51225 
diff
changeset
 | 
38  | 
val read_span: outer_syntax -> Token.T list -> Toplevel.transition  | 
| 5829 | 39  | 
end;  | 
40  | 
||
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
41  | 
structure Outer_Syntax: OUTER_SYNTAX =  | 
| 5829 | 42  | 
struct  | 
43  | 
||
44  | 
(** outer syntax **)  | 
|
45  | 
||
| 29311 | 46  | 
(* command parsers *)  | 
| 5829 | 47  | 
|
| 29311 | 48  | 
datatype command = Command of  | 
| 24868 | 49  | 
 {comment: string,
 | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
36959 
diff
changeset
 | 
50  | 
markup: Thy_Output.markup option,  | 
| 24868 | 51  | 
int_only: bool,  | 
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
52  | 
parse: (Toplevel.transition -> Toplevel.transition) parser,  | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
53  | 
pos: Position.T,  | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
54  | 
id: serial};  | 
| 5829 | 55  | 
|
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
56  | 
fun new_command comment markup int_only parse pos =  | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
57  | 
  Command {comment = comment, markup = markup, int_only = int_only, parse = parse,
 | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
58  | 
pos = pos, id = serial ()};  | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
59  | 
|
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
60  | 
fun command_markup def (name, Command {pos, id, ...}) =
 | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
61  | 
Markup.properties (Position.entity_properties_of def id pos)  | 
| 
50201
 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 
wenzelm 
parents: 
49564 
diff
changeset
 | 
62  | 
(Markup.entity Markup.commandN name);  | 
| 5829 | 63  | 
|
| 50213 | 64  | 
fun pretty_command (cmd as (name, Command {comment, ...})) =
 | 
65  | 
Pretty.block  | 
|
66  | 
(Pretty.marks_str  | 
|
| 
50450
 
358b6020f8b6
generalized notion of active area, where sendback is just one application;
 
wenzelm 
parents: 
50215 
diff
changeset
 | 
67  | 
      ([Active.make_markup Markup.sendbackN {implicit = true, properties = [Markup.padding_line]},
 | 
| 50215 | 68  | 
command_markup false cmd], name) :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment);  | 
| 50213 | 69  | 
|
| 5829 | 70  | 
|
71  | 
(* parse command *)  | 
|
72  | 
||
| 6860 | 73  | 
local  | 
| 6199 | 74  | 
|
| 
14925
 
0f86a8a694f8
added read (provides transition names and sources);
 
wenzelm 
parents: 
14687 
diff
changeset
 | 
75  | 
fun terminate false = Scan.succeed ()  | 
| 44357 | 76  | 
| terminate true =  | 
77  | 
Parse.group (fn () => "end of input")  | 
|
78  | 
(Scan.option Parse.sync -- Parse.semicolon >> K ());  | 
|
| 
14925
 
0f86a8a694f8
added read (provides transition names and sources);
 
wenzelm 
parents: 
14687 
diff
changeset
 | 
79  | 
|
| 26620 | 80  | 
fun body cmd (name, _) =  | 
| 7026 | 81  | 
(case cmd name of  | 
| 29311 | 82  | 
    SOME (Command {int_only, parse, ...}) =>
 | 
| 36950 | 83  | 
Parse.!!! (Scan.prompt (name ^ "# ") (Parse.tags |-- parse >> pair int_only))  | 
| 
48191
 
c1def7433a72
internalize error into command transaction -- relevant for commands that are declared via 'keywords', but not defined yet;
 
wenzelm 
parents: 
47416 
diff
changeset
 | 
84  | 
| NONE =>  | 
| 
 
c1def7433a72
internalize error into command transaction -- relevant for commands that are declared via 'keywords', but not defined yet;
 
wenzelm 
parents: 
47416 
diff
changeset
 | 
85  | 
Scan.succeed (false, Toplevel.imperative (fn () =>  | 
| 
 
c1def7433a72
internalize error into command transaction -- relevant for commands that are declared via 'keywords', but not defined yet;
 
wenzelm 
parents: 
47416 
diff
changeset
 | 
86  | 
        error ("Bad parser for outer syntax command " ^ quote name))));
 | 
| 6860 | 87  | 
|
88  | 
in  | 
|
| 5829 | 89  | 
|
| 26620 | 90  | 
fun parse_command do_terminate cmd =  | 
| 36950 | 91  | 
Parse.semicolon >> K NONE ||  | 
92  | 
Parse.sync >> K NONE ||  | 
|
93  | 
(Parse.position Parse.command :-- body cmd) --| terminate do_terminate  | 
|
| 6860 | 94  | 
>> (fn ((name, pos), (int_only, f)) =>  | 
| 15531 | 95  | 
SOME (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>  | 
| 6860 | 96  | 
Toplevel.interactive int_only |> f));  | 
| 5829 | 97  | 
|
| 6199 | 98  | 
end;  | 
99  | 
||
| 5829 | 100  | 
|
| 43711 | 101  | 
(* type outer_syntax *)  | 
102  | 
||
103  | 
datatype outer_syntax = Outer_Syntax of  | 
|
104  | 
 {commands: command Symtab.table,
 | 
|
105  | 
markups: (string * Thy_Output.markup) list};  | 
|
106  | 
||
107  | 
fun make_outer_syntax commands markups =  | 
|
108  | 
  Outer_Syntax {commands = commands, markups = markups};
 | 
|
109  | 
||
110  | 
val empty_outer_syntax = make_outer_syntax Symtab.empty [];  | 
|
111  | 
||
112  | 
||
113  | 
fun map_commands f (Outer_Syntax {commands, ...}) =
 | 
|
114  | 
let  | 
|
115  | 
val commands' = f commands;  | 
|
116  | 
val markups' =  | 
|
117  | 
      Symtab.fold (fn (name, Command {markup = SOME m, ...}) => cons (name, m) | _ => I)
 | 
|
118  | 
commands' [];  | 
|
119  | 
in make_outer_syntax commands' markups' end;  | 
|
120  | 
||
121  | 
fun dest_commands (Outer_Syntax {commands, ...}) =
 | 
|
| 50213 | 122  | 
commands |> Symtab.dest |> sort_wrt #1;  | 
| 43711 | 123  | 
|
124  | 
fun lookup_commands (Outer_Syntax {commands, ...}) = Symtab.lookup commands;
 | 
|
125  | 
||
126  | 
fun is_markup (Outer_Syntax {markups, ...}) kind name =
 | 
|
127  | 
AList.lookup (op =) markups name = SOME kind;  | 
|
128  | 
||
129  | 
||
| 5829 | 130  | 
|
| 9132 | 131  | 
(** global outer syntax **)  | 
| 5829 | 132  | 
|
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
133  | 
type command_spec = (string * Keyword.T) * Position.T;  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
134  | 
|
| 7026 | 135  | 
local  | 
136  | 
||
| 43711 | 137  | 
(*synchronized wrt. Keywords*)  | 
138  | 
val global_outer_syntax = Unsynchronized.ref empty_outer_syntax;  | 
|
| 5952 | 139  | 
|
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
140  | 
fun add_command (name, kind) cmd = CRITICAL (fn () =>  | 
| 
46950
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46876 
diff
changeset
 | 
141  | 
let  | 
| 
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46876 
diff
changeset
 | 
142  | 
val thy = ML_Context.the_global_context ();  | 
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
143  | 
    val Command {pos, ...} = cmd;
 | 
| 
46950
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46876 
diff
changeset
 | 
144  | 
val _ =  | 
| 
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46876 
diff
changeset
 | 
145  | 
(case try (Thy_Header.the_keyword thy) name of  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
146  | 
SOME spec =>  | 
| 
48864
 
3ee314ae1e0a
added keyword kind "thy_load" (with optional list of file extensions);
 
wenzelm 
parents: 
48771 
diff
changeset
 | 
147  | 
if Option.map #1 spec = SOME (Keyword.kind_files_of kind) then ()  | 
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
148  | 
          else error ("Inconsistent outer syntax keyword declaration " ^
 | 
| 48992 | 149  | 
quote name ^ Position.here pos)  | 
| 
46950
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46876 
diff
changeset
 | 
150  | 
| NONE =>  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
151  | 
if Context.theory_name thy = Context.PureN  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
152  | 
then Keyword.define (name, SOME kind)  | 
| 48992 | 153  | 
          else error ("Undeclared outer syntax command " ^ quote name ^ Position.here pos));
 | 
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
154  | 
val _ = Position.report pos (command_markup true (name, cmd));  | 
| 
46950
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46876 
diff
changeset
 | 
155  | 
in  | 
| 
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46876 
diff
changeset
 | 
156  | 
Unsynchronized.change global_outer_syntax (map_commands (fn commands =>  | 
| 
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46876 
diff
changeset
 | 
157  | 
(if not (Symtab.defined commands name) then ()  | 
| 
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46876 
diff
changeset
 | 
158  | 
      else warning ("Redefining outer syntax command " ^ quote name);
 | 
| 
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46876 
diff
changeset
 | 
159  | 
Symtab.update (name, cmd) commands)))  | 
| 
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46876 
diff
changeset
 | 
160  | 
end);  | 
| 6722 | 161  | 
|
| 7026 | 162  | 
in  | 
163  | 
||
| 43711 | 164  | 
fun get_syntax () = CRITICAL (fn () => (Keyword.get_lexicons (), ! global_outer_syntax));  | 
| 7789 | 165  | 
|
| 
46970
 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
166  | 
fun check_syntax () =  | 
| 
 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
167  | 
let  | 
| 
 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
168  | 
val ((_, major), syntax) = CRITICAL (fn () => (Keyword.dest (), ! global_outer_syntax));  | 
| 
 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
169  | 
in  | 
| 
 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
170  | 
(case subtract (op =) (map #1 (dest_commands syntax)) major of  | 
| 
 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
171  | 
[] => ()  | 
| 
 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
172  | 
    | missing => error ("Missing outer syntax command(s) " ^ commas_quote missing))
 | 
| 
 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
173  | 
end;  | 
| 
 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
174  | 
|
| 43711 | 175  | 
fun lookup_commands_dynamic () = lookup_commands (! global_outer_syntax);  | 
| 5829 | 176  | 
|
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
177  | 
fun command (spec, pos) comment parse =  | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
178  | 
add_command spec (new_command comment NONE false parse pos);  | 
| 24868 | 179  | 
|
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
180  | 
fun markup_command markup (spec, pos) comment parse =  | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
181  | 
add_command spec (new_command comment (SOME markup) false parse pos);  | 
| 7026 | 182  | 
|
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
183  | 
fun improper_command (spec, pos) comment parse =  | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
184  | 
add_command spec (new_command comment NONE true parse pos);  | 
| 29311 | 185  | 
|
| 43711 | 186  | 
end;  | 
187  | 
||
| 5829 | 188  | 
|
| 26990 | 189  | 
(* local_theory commands *)  | 
190  | 
||
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
191  | 
fun local_theory_command do_print trans command_spec comment parse =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
192  | 
command command_spec comment (Parse.opt_target -- parse  | 
| 26990 | 193  | 
>> (fn (loc, f) => (if do_print then Toplevel.print else I) o trans loc f));  | 
194  | 
||
| 29380 | 195  | 
val local_theory' = local_theory_command false Toplevel.local_theory';  | 
| 29311 | 196  | 
val local_theory = local_theory_command false Toplevel.local_theory;  | 
| 26990 | 197  | 
val local_theory_to_proof' = local_theory_command true Toplevel.local_theory_to_proof';  | 
| 29311 | 198  | 
val local_theory_to_proof = local_theory_command true Toplevel.local_theory_to_proof;  | 
| 26990 | 199  | 
|
200  | 
||
| 24872 | 201  | 
(* inspect syntax *)  | 
| 7026 | 202  | 
|
| 50213 | 203  | 
fun help_outer_syntax pats =  | 
204  | 
dest_commands (#2 (get_syntax ()))  | 
|
205  | 
|> filter (fn (name, _) => forall (fn pat => match_string pat name) pats)  | 
|
206  | 
|> map pretty_command  | 
|
207  | 
|> Pretty.chunks |> Pretty.writeln;  | 
|
208  | 
||
| 9223 | 209  | 
fun print_outer_syntax () =  | 
| 7026 | 210  | 
let  | 
| 
46957
 
0c15caf47040
clarified Keyword.is_keyword: union of minor and major;
 
wenzelm 
parents: 
46950 
diff
changeset
 | 
211  | 
val ((keywords, _), outer_syntax) =  | 
| 
 
0c15caf47040
clarified Keyword.is_keyword: union of minor and major;
 
wenzelm 
parents: 
46950 
diff
changeset
 | 
212  | 
CRITICAL (fn () => (Keyword.dest (), #2 (get_syntax ())));  | 
| 50213 | 213  | 
val (int_cmds, cmds) =  | 
214  | 
      List.partition (fn (_, Command {int_only, ...}) => int_only) (dest_commands outer_syntax);
 | 
|
| 7026 | 215  | 
in  | 
| 43711 | 216  | 
    [Pretty.strs ("syntax keywords:" :: map quote keywords),
 | 
| 50213 | 217  | 
Pretty.big_list "commands:" (map pretty_command cmds),  | 
218  | 
Pretty.big_list "interactive-only commands:" (map pretty_command int_cmds)]  | 
|
| 9223 | 219  | 
|> Pretty.chunks |> Pretty.writeln  | 
| 7026 | 220  | 
end;  | 
| 5829 | 221  | 
|
222  | 
||
223  | 
||
| 9132 | 224  | 
(** toplevel parsing **)  | 
| 5829 | 225  | 
|
| 9132 | 226  | 
(* basic sources *)  | 
| 6860 | 227  | 
|
| 26620 | 228  | 
fun toplevel_source term do_recover cmd src =  | 
| 9132 | 229  | 
let  | 
230  | 
val no_terminator =  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
231  | 
Scan.unless Parse.semicolon (Scan.one (Token.not_sync andf Token.not_eof));  | 
| 
23682
 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 
wenzelm 
parents: 
23679 
diff
changeset
 | 
232  | 
fun recover int =  | 
| 
 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 
wenzelm 
parents: 
23679 
diff
changeset
 | 
233  | 
(int, fn _ => Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [NONE]);  | 
| 9132 | 234  | 
in  | 
235  | 
src  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
236  | 
|> Token.source_proper  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
237  | 
|> Source.source Token.stopper  | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51271 
diff
changeset
 | 
238  | 
(Scan.bulk (Parse.$$$ "--" -- Parse.!!! Parse.document_source >> K NONE || Parse.not_eof >> SOME))  | 
| 
23682
 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 
wenzelm 
parents: 
23679 
diff
changeset
 | 
239  | 
(Option.map recover do_recover)  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19060 
diff
changeset
 | 
240  | 
|> Source.map_filter I  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
241  | 
|> Source.source Token.stopper  | 
| 36950 | 242  | 
(Scan.bulk (fn xs => Parse.!!! (parse_command term (cmd ())) xs))  | 
| 
23682
 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 
wenzelm 
parents: 
23679 
diff
changeset
 | 
243  | 
(Option.map recover do_recover)  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19060 
diff
changeset
 | 
244  | 
|> Source.map_filter I  | 
| 9132 | 245  | 
end;  | 
| 5829 | 246  | 
|
| 7746 | 247  | 
|
| 25580 | 248  | 
(* off-line scanning/parsing *)  | 
| 
14925
 
0f86a8a694f8
added read (provides transition names and sources);
 
wenzelm 
parents: 
14687 
diff
changeset
 | 
249  | 
|
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
250  | 
fun scan pos str =  | 
| 16195 | 251  | 
Source.of_string str  | 
| 
40523
 
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
 
wenzelm 
parents: 
38422 
diff
changeset
 | 
252  | 
|> Symbol.source  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
253  | 
  |> Token.source {do_recover = SOME false} Keyword.get_lexicons pos
 | 
| 16195 | 254  | 
|> Source.exhaust;  | 
255  | 
||
| 25580 | 256  | 
fun parse pos str =  | 
257  | 
Source.of_string str  | 
|
| 
40523
 
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
 
wenzelm 
parents: 
38422 
diff
changeset
 | 
258  | 
|> Symbol.source  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
259  | 
  |> Token.source {do_recover = SOME false} Keyword.get_lexicons pos
 | 
| 43711 | 260  | 
|> toplevel_source false NONE lookup_commands_dynamic  | 
| 25580 | 261  | 
|> Source.exhaust;  | 
262  | 
||
| 14091 | 263  | 
|
| 24868 | 264  | 
(* interactive source of toplevel transformers *)  | 
265  | 
||
| 26600 | 266  | 
type isar =  | 
267  | 
(Toplevel.transition, (Toplevel.transition option,  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
268  | 
(Token.T, (Token.T option, (Token.T, (Token.T,  | 
| 30573 | 269  | 
(Symbol_Pos.T, Position.T * (Symbol.symbol, (string, unit) Source.source)  | 
| 27770 | 270  | 
Source.source) Source.source) Source.source) Source.source)  | 
271  | 
Source.source) Source.source) Source.source) Source.source;  | 
|
| 26600 | 272  | 
|
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
273  | 
fun isar in_stream term : isar =  | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
274  | 
Source.tty in_stream  | 
| 
40523
 
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
 
wenzelm 
parents: 
38422 
diff
changeset
 | 
275  | 
|> Symbol.source  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
276  | 
  |> Token.source {do_recover = SOME true} Keyword.get_lexicons Position.none
 | 
| 43711 | 277  | 
|> toplevel_source term (SOME true) lookup_commands_dynamic;  | 
| 24868 | 278  | 
|
279  | 
||
| 
48918
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
280  | 
(* side-comments *)  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
281  | 
|
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
282  | 
local  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
283  | 
|
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
284  | 
fun cmts (t1 :: t2 :: toks) =  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
285  | 
if Token.keyword_with (fn s => s = "--") t1 then t2 :: cmts toks  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
286  | 
else cmts (t2 :: toks)  | 
| 49564 | 287  | 
| cmts _ = [];  | 
| 
48918
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
288  | 
|
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
289  | 
in  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
290  | 
|
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
291  | 
val span_cmts = filter Token.is_proper #> cmts;  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
292  | 
|
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
293  | 
end;  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
294  | 
|
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
295  | 
|
| 51271 | 296  | 
(* read command span -- fail-safe *)  | 
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
297  | 
|
| 
44658
 
5bec9c15ef29
more direct Token.range_pos and Outer_Syntax.read_command, bypassing Thy_Syntax.span;
 
wenzelm 
parents: 
44478 
diff
changeset
 | 
298  | 
fun read_span outer_syntax toks =  | 
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
299  | 
let  | 
| 43711 | 300  | 
val commands = lookup_commands outer_syntax;  | 
| 
48749
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48647 
diff
changeset
 | 
301  | 
|
| 
48771
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
48768 
diff
changeset
 | 
302  | 
val proper_range = Position.set_range (Command.proper_range toks);  | 
| 
48749
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48647 
diff
changeset
 | 
303  | 
val pos =  | 
| 
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48647 
diff
changeset
 | 
304  | 
(case find_first Token.is_command toks of  | 
| 
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48647 
diff
changeset
 | 
305  | 
SOME tok => Token.position_of tok  | 
| 
48771
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
48768 
diff
changeset
 | 
306  | 
| NONE => proper_range);  | 
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
307  | 
|
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
308  | 
fun command_reports tok =  | 
| 
48749
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48647 
diff
changeset
 | 
309  | 
if Token.is_command tok then  | 
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
310  | 
let val name = Token.content_of tok in  | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
311  | 
(case commands name of  | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
312  | 
NONE => []  | 
| 
48768
 
abc45de5bb22
clarified "bad" markup: proper body text, invent missing serial on Scala side (counting backwards);
 
wenzelm 
parents: 
48749 
diff
changeset
 | 
313  | 
| SOME cmd => [((Token.position_of tok, command_markup false (name, cmd)), "")])  | 
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
314  | 
end  | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
315  | 
else [];  | 
| 
48749
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48647 
diff
changeset
 | 
316  | 
|
| 
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48647 
diff
changeset
 | 
317  | 
val (is_malformed, token_reports) = Thy_Syntax.reports_of_tokens toks;  | 
| 
48768
 
abc45de5bb22
clarified "bad" markup: proper body text, invent missing serial on Scala side (counting backwards);
 
wenzelm 
parents: 
48749 
diff
changeset
 | 
318  | 
val _ = Position.reports_text (token_reports @ maps command_reports toks);  | 
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
319  | 
in  | 
| 
51268
 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 
wenzelm 
parents: 
51225 
diff
changeset
 | 
320  | 
if is_malformed then Toplevel.malformed pos "Malformed command syntax"  | 
| 
48749
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48647 
diff
changeset
 | 
321  | 
else  | 
| 
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48647 
diff
changeset
 | 
322  | 
(case Source.exhaust (toplevel_source false NONE (K commands) (Source.of_list toks)) of  | 
| 
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48647 
diff
changeset
 | 
323  | 
[tr] =>  | 
| 
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48647 
diff
changeset
 | 
324  | 
if Keyword.is_control (Toplevel.name_of tr) then  | 
| 
51268
 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 
wenzelm 
parents: 
51225 
diff
changeset
 | 
325  | 
Toplevel.malformed pos "Illegal control command"  | 
| 
 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 
wenzelm 
parents: 
51225 
diff
changeset
 | 
326  | 
else tr  | 
| 
 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 
wenzelm 
parents: 
51225 
diff
changeset
 | 
327  | 
| [] => Toplevel.ignored (Position.set_range (Command.range toks))  | 
| 
 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 
wenzelm 
parents: 
51225 
diff
changeset
 | 
328  | 
| _ => Toplevel.malformed proper_range "Exactly one command expected")  | 
| 
 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 
wenzelm 
parents: 
51225 
diff
changeset
 | 
329  | 
handle ERROR msg => Toplevel.malformed proper_range msg  | 
| 
28436
 
4faf705a177d
load_thy: more precise treatment of improper cmd or proof (notably 'oops');
 
wenzelm 
parents: 
28432 
diff
changeset
 | 
330  | 
end;  | 
| 
28432
 
944cb67f809a
load_thy: Toplevel.excursion based on units of command/proof pairs;
 
wenzelm 
parents: 
28424 
diff
changeset
 | 
331  | 
|
| 5829 | 332  | 
end;  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
333  |