| author | wenzelm | 
| Thu, 13 Aug 2015 11:05:19 +0200 | |
| changeset 60924 | 610794dff23c | 
| parent 60693 | 044f8bb3dd30 | 
| child 61579 | 634cd44bb1d3 | 
| permissions | -rw-r--r-- | 
| 5829 | 1  | 
(* Title: Pure/Isar/outer_syntax.ML  | 
2  | 
Author: Markus Wenzel, TU Muenchen  | 
|
3  | 
||
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
4  | 
Isabelle/Isar outer syntax.  | 
| 5829 | 5  | 
*)  | 
6  | 
||
7  | 
signature OUTER_SYNTAX =  | 
|
8  | 
sig  | 
|
| 58930 | 9  | 
val help: theory -> string list -> unit  | 
10  | 
val print_commands: theory -> unit  | 
|
| 59935 | 11  | 
type command_keyword = string * Position.T  | 
12  | 
val command: command_keyword -> string ->  | 
|
| 29311 | 13  | 
(Toplevel.transition -> Toplevel.transition) parser -> unit  | 
| 59935 | 14  | 
val local_theory': command_keyword -> string ->  | 
| 29380 | 15  | 
(bool -> local_theory -> local_theory) parser -> unit  | 
| 59935 | 16  | 
val local_theory: command_keyword -> string ->  | 
| 29311 | 17  | 
(local_theory -> local_theory) parser -> unit  | 
| 59935 | 18  | 
val local_theory_to_proof': command_keyword -> string ->  | 
| 29311 | 19  | 
(bool -> local_theory -> Proof.state) parser -> unit  | 
| 59935 | 20  | 
val local_theory_to_proof: command_keyword -> string ->  | 
| 29311 | 21  | 
(local_theory -> Proof.state) parser -> unit  | 
| 60095 | 22  | 
val bootstrap: bool Config.T  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
23  | 
val parse: theory -> Position.T -> string -> Toplevel.transition list  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
24  | 
val parse_tokens: theory -> Token.T list -> Toplevel.transition list  | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
25  | 
val parse_spans: Token.T list -> Command_Span.span list  | 
| 52510 | 26  | 
val side_comments: Token.T list -> Token.T list  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
27  | 
val command_reports: theory -> Token.T -> Position.report_text list  | 
| 5829 | 28  | 
end;  | 
29  | 
||
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
30  | 
structure Outer_Syntax: OUTER_SYNTAX =  | 
| 5829 | 31  | 
struct  | 
32  | 
||
33  | 
(** outer syntax **)  | 
|
34  | 
||
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
35  | 
(* errors *)  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
36  | 
|
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
37  | 
fun err_command msg name ps =  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
38  | 
error (msg ^ quote (Markup.markup Markup.keyword1 name) ^ Position.here_list ps);  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
39  | 
|
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
40  | 
fun err_dup_command name ps =  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
41  | 
err_command "Duplicate outer syntax command " name ps;  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
42  | 
|
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
43  | 
|
| 29311 | 44  | 
(* command parsers *)  | 
| 5829 | 45  | 
|
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
46  | 
datatype command_parser =  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
47  | 
Parser of (Toplevel.transition -> Toplevel.transition) parser |  | 
| 60691 | 48  | 
Restricted_Parser of  | 
| 
59939
 
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
 
wenzelm 
parents: 
59935 
diff
changeset
 | 
49  | 
(bool * Position.T) option -> (Toplevel.transition -> Toplevel.transition) parser;  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
50  | 
|
| 29311 | 51  | 
datatype command = Command of  | 
| 24868 | 52  | 
 {comment: string,
 | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
53  | 
command_parser: command_parser,  | 
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
54  | 
pos: Position.T,  | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
55  | 
id: serial};  | 
| 5829 | 56  | 
|
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
57  | 
fun eq_command (Command {id = id1, ...}, Command {id = id2, ...}) = id1 = id2;
 | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
58  | 
|
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
59  | 
fun new_command comment command_parser pos =  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
60  | 
  Command {comment = comment, command_parser = command_parser, pos = pos, id = serial ()};
 | 
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
61  | 
|
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
62  | 
fun command_pos (Command {pos, ...}) = pos;
 | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
63  | 
|
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
64  | 
fun command_markup def (name, Command {pos, id, ...}) =
 | 
| 
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
65  | 
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
 | 
66  | 
(Markup.entity Markup.commandN name);  | 
| 5829 | 67  | 
|
| 50213 | 68  | 
fun pretty_command (cmd as (name, Command {comment, ...})) =
 | 
69  | 
Pretty.block  | 
|
70  | 
(Pretty.marks_str  | 
|
| 
50450
 
358b6020f8b6
generalized notion of active area, where sendback is just one application;
 
wenzelm 
parents: 
50215 
diff
changeset
 | 
71  | 
      ([Active.make_markup Markup.sendbackN {implicit = true, properties = [Markup.padding_line]},
 | 
| 50215 | 72  | 
command_markup false cmd], name) :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment);  | 
| 50213 | 73  | 
|
| 5829 | 74  | 
|
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
75  | 
(* theory data *)  | 
| 43711 | 76  | 
|
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
77  | 
structure Data = Theory_Data  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
78  | 
(  | 
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
79  | 
type T = command Symtab.table;  | 
| 
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
80  | 
val empty = Symtab.empty;  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
81  | 
val extend = I;  | 
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
82  | 
fun merge data : T =  | 
| 
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
83  | 
data |> Symtab.join (fn name => fn (cmd1, cmd2) =>  | 
| 
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
84  | 
if eq_command (cmd1, cmd2) then raise Symtab.SAME  | 
| 
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
85  | 
else err_dup_command name [command_pos cmd1, command_pos cmd2]);  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
86  | 
);  | 
| 43711 | 87  | 
|
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
88  | 
val get_commands = Data.get;  | 
| 
60924
 
610794dff23c
tuned signature, in accordance to sortBy in Scala;
 
wenzelm 
parents: 
60693 
diff
changeset
 | 
89  | 
val dest_commands = get_commands #> Symtab.dest #> sort_by #1;  | 
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
90  | 
val lookup_commands = Symtab.lookup o get_commands;  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
91  | 
|
| 58930 | 92  | 
fun help thy pats =  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
93  | 
dest_commands thy  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
94  | 
|> filter (fn (name, _) => forall (fn pat => match_string pat name) pats)  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
95  | 
|> map pretty_command  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
96  | 
|> Pretty.writeln_chunks;  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
97  | 
|
| 58930 | 98  | 
fun print_commands thy =  | 
| 43711 | 99  | 
let  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
100  | 
val keywords = Thy_Header.get_keywords thy;  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
101  | 
val minor = Scan.dest_lexicon (Keyword.minor_keywords keywords);  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
102  | 
val commands = dest_commands thy;  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
103  | 
in  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
104  | 
    [Pretty.strs ("keywords:" :: map quote minor),
 | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
105  | 
Pretty.big_list "commands:" (map pretty_command commands)]  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
106  | 
|> Pretty.writeln_chunks  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
107  | 
end;  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
108  | 
|
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
109  | 
|
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
110  | 
(* maintain commands *)  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
111  | 
|
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
112  | 
fun add_command name cmd thy =  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
113  | 
let  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
114  | 
val _ =  | 
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
115  | 
Keyword.is_command (Thy_Header.get_keywords thy) name orelse  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
116  | 
err_command "Undeclared outer syntax command " name [command_pos cmd];  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
117  | 
val _ =  | 
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
118  | 
(case lookup_commands thy name of  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
119  | 
NONE => ()  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
120  | 
| SOME cmd' => err_dup_command name [command_pos cmd, command_pos cmd']);  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
121  | 
val _ =  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
122  | 
Context_Position.report_generic (ML_Context.the_generic_context ())  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
123  | 
(command_pos cmd) (command_markup true (name, cmd));  | 
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
124  | 
in Data.map (Symtab.update (name, cmd)) thy end;  | 
| 43711 | 125  | 
|
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
126  | 
val _ = Theory.setup (Theory.at_end (fn thy =>  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
127  | 
let  | 
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
128  | 
val command_keywords =  | 
| 
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
129  | 
Scan.dest_lexicon (Keyword.major_keywords (Thy_Header.get_keywords thy));  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
130  | 
val _ =  | 
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58934 
diff
changeset
 | 
131  | 
(case subtract (op =) (map #1 (dest_commands thy)) command_keywords of  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
132  | 
[] => ()  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
133  | 
      | missing => error ("Missing outer syntax command(s) " ^ commas_quote missing))
 | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
134  | 
in NONE end));  | 
| 43711 | 135  | 
|
| 5829 | 136  | 
|
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
137  | 
(* implicit theory setup *)  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
138  | 
|
| 59935 | 139  | 
type command_keyword = string * Position.T;  | 
| 5952 | 140  | 
|
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
141  | 
fun raw_command (name, pos) comment command_parser =  | 
| 59932 | 142  | 
let val setup = add_command name (new_command comment command_parser pos)  | 
143  | 
in Context.>> (Context.mapping setup (Local_Theory.background_theory setup)) end;  | 
|
| 26990 | 144  | 
|
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
145  | 
fun command (name, pos) comment parse =  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
146  | 
raw_command (name, pos) comment (Parser parse);  | 
| 
59923
 
b21c82422d65
support private scope for individual local theory commands;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
147  | 
|
| 59935 | 148  | 
fun local_theory_command trans command_keyword comment parse =  | 
149  | 
raw_command command_keyword comment  | 
|
| 60691 | 150  | 
(Restricted_Parser (fn restricted =>  | 
151  | 
Parse.opt_target -- parse >> (fn (target, f) => trans restricted target f)));  | 
|
| 26990 | 152  | 
|
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56334 
diff
changeset
 | 
153  | 
val local_theory' = local_theory_command Toplevel.local_theory';  | 
| 
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56334 
diff
changeset
 | 
154  | 
val local_theory = local_theory_command Toplevel.local_theory;  | 
| 
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56334 
diff
changeset
 | 
155  | 
val local_theory_to_proof' = local_theory_command Toplevel.local_theory_to_proof';  | 
| 
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56334 
diff
changeset
 | 
156  | 
val local_theory_to_proof = local_theory_command Toplevel.local_theory_to_proof;  | 
| 26990 | 157  | 
|
158  | 
||
| 5829 | 159  | 
|
| 9132 | 160  | 
(** toplevel parsing **)  | 
| 5829 | 161  | 
|
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
162  | 
(* parse commands *)  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
163  | 
|
| 60095 | 164  | 
val bootstrap =  | 
165  | 
  Config.bool (Config.declare ("Outer_Syntax.bootstrap", @{here}) (K (Config.Bool true)));
 | 
|
166  | 
||
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
167  | 
local  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
168  | 
|
| 
59939
 
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
 
wenzelm 
parents: 
59935 
diff
changeset
 | 
169  | 
val before_command =  | 
| 
59990
 
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
 
wenzelm 
parents: 
59939 
diff
changeset
 | 
170  | 
Scan.option (Parse.position (Parse.private >> K true || Parse.qualified >> K false));  | 
| 
59939
 
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
 
wenzelm 
parents: 
59935 
diff
changeset
 | 
171  | 
|
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
172  | 
fun parse_command thy =  | 
| 
59939
 
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
 
wenzelm 
parents: 
59935 
diff
changeset
 | 
173  | 
Scan.ahead (before_command |-- Parse.position Parse.command_) :|-- (fn (name, pos) =>  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
174  | 
let  | 
| 
60693
 
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
 
wenzelm 
parents: 
60691 
diff
changeset
 | 
175  | 
val keywords = Thy_Header.get_keywords thy;  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
176  | 
val command_tags = Parse.command_ -- Parse.tags;  | 
| 
60693
 
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
 
wenzelm 
parents: 
60691 
diff
changeset
 | 
177  | 
val tr0 =  | 
| 
 
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
 
wenzelm 
parents: 
60691 
diff
changeset
 | 
178  | 
Toplevel.empty  | 
| 
 
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
 
wenzelm 
parents: 
60691 
diff
changeset
 | 
179  | 
|> Toplevel.name name  | 
| 
 
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
 
wenzelm 
parents: 
60691 
diff
changeset
 | 
180  | 
|> Toplevel.position pos  | 
| 
 
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
 
wenzelm 
parents: 
60691 
diff
changeset
 | 
181  | 
|> Keyword.is_proof_open keywords name ? Toplevel.skip_proof_open  | 
| 
 
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
 
wenzelm 
parents: 
60691 
diff
changeset
 | 
182  | 
|> Keyword.is_proof_close keywords name ? Toplevel.skip_proof_close;  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
183  | 
in  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
184  | 
(case lookup_commands thy name of  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
185  | 
        SOME (Command {command_parser = Parser parse, ...}) =>
 | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
186  | 
Parse.!!! (command_tags |-- parse) >> (fn f => f tr0)  | 
| 60691 | 187  | 
      | SOME (Command {command_parser = Restricted_Parser parse, ...}) =>
 | 
188  | 
before_command :|-- (fn restricted =>  | 
|
189  | 
Parse.!!! (command_tags |-- parse restricted)) >> (fn f => f tr0)  | 
|
| 60095 | 190  | 
| NONE =>  | 
191  | 
Scan.fail_with (fn _ => fn _ =>  | 
|
192  | 
let  | 
|
193  | 
val msg =  | 
|
194  | 
if Config.get_global thy bootstrap  | 
|
195  | 
then "missing theory context for command "  | 
|
196  | 
else "undefined command ";  | 
|
197  | 
in msg ^ quote (Markup.markup Markup.keyword1 name) end))  | 
|
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
198  | 
end);  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
199  | 
|
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
200  | 
val parse_cmt = Parse.$$$ "--" -- Parse.!!! Parse.document_source;  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
201  | 
|
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
202  | 
fun commands_source thy =  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
203  | 
Token.source_proper #>  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
204  | 
Source.source Token.stopper (Scan.bulk (parse_cmt >> K NONE || Parse.not_eof >> SOME)) #>  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
205  | 
Source.map_filter I #>  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
206  | 
Source.source Token.stopper (Scan.bulk (fn xs => Parse.!!! (parse_command thy) xs));  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
207  | 
|
| 
60073
 
76a8400a58d9
more robust error handling of commands that are declared but not yet defined;
 
wenzelm 
parents: 
59990 
diff
changeset
 | 
208  | 
in  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
209  | 
|
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
210  | 
fun parse thy pos str =  | 
| 
57918
 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 
wenzelm 
parents: 
57905 
diff
changeset
 | 
211  | 
Source.of_string str  | 
| 
 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 
wenzelm 
parents: 
57905 
diff
changeset
 | 
212  | 
|> Symbol.source  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
213  | 
|> Token.source (Thy_Header.get_keywords thy) pos  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
214  | 
|> commands_source thy  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
215  | 
|> Source.exhaust;  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
216  | 
|
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
217  | 
fun parse_tokens thy toks =  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
218  | 
Source.of_list toks  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
219  | 
|> commands_source thy  | 
| 
57918
 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 
wenzelm 
parents: 
57905 
diff
changeset
 | 
220  | 
|> Source.exhaust;  | 
| 
 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 
wenzelm 
parents: 
57905 
diff
changeset
 | 
221  | 
|
| 
60073
 
76a8400a58d9
more robust error handling of commands that are declared but not yet defined;
 
wenzelm 
parents: 
59990 
diff
changeset
 | 
222  | 
end;  | 
| 
 
76a8400a58d9
more robust error handling of commands that are declared but not yet defined;
 
wenzelm 
parents: 
59990 
diff
changeset
 | 
223  | 
|
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
224  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
225  | 
(* parse spans *)  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
226  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
227  | 
local  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
228  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
229  | 
fun ship span =  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
230  | 
let  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
231  | 
val kind =  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
232  | 
if forall Token.is_improper span then Command_Span.Ignored_Span  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
233  | 
else if exists Token.is_error span then Command_Span.Malformed_Span  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
234  | 
else  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
235  | 
(case find_first Token.is_command span of  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
236  | 
NONE => Command_Span.Malformed_Span  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
237  | 
| SOME cmd => Command_Span.Command_Span (Token.content_of cmd, Token.pos_of cmd));  | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
238  | 
in cons (Command_Span.Span (kind, span)) end;  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
239  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
240  | 
fun flush (result, content, improper) =  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
241  | 
result  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
242  | 
|> not (null content) ? ship (rev content)  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
243  | 
|> not (null improper) ? ship (rev improper);  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
244  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
245  | 
fun parse tok (result, content, improper) =  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
246  | 
if Token.is_improper tok then (result, content, tok :: improper)  | 
| 
59939
 
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
 
wenzelm 
parents: 
59935 
diff
changeset
 | 
247  | 
else if Token.is_command_modifier tok orelse  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
248  | 
Token.is_command tok andalso  | 
| 
59939
 
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
 
wenzelm 
parents: 
59935 
diff
changeset
 | 
249  | 
(not (exists Token.is_command_modifier content) orelse exists Token.is_command content)  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59923 
diff
changeset
 | 
250  | 
then (flush (result, content, improper), [tok], [])  | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
251  | 
else (result, tok :: (improper @ content), []);  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
252  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
253  | 
in  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
254  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
255  | 
fun parse_spans toks =  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
256  | 
fold parse toks ([], [], []) |> flush |> rev;  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
257  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
258  | 
end;  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57623 
diff
changeset
 | 
259  | 
|
| 14091 | 260  | 
|
| 52510 | 261  | 
(* side-comments *)  | 
| 
48749
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48647 
diff
changeset
 | 
262  | 
|
| 52510 | 263  | 
fun cmts (t1 :: t2 :: toks) =  | 
264  | 
if Token.keyword_with (fn s => s = "--") t1 then t2 :: cmts toks  | 
|
265  | 
else cmts (t2 :: toks)  | 
|
266  | 
| cmts _ = [];  | 
|
267  | 
||
268  | 
val side_comments = filter Token.is_proper #> cmts;  | 
|
269  | 
||
270  | 
||
271  | 
(* read commands *)  | 
|
| 
48647
 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 
wenzelm 
parents: 
48646 
diff
changeset
 | 
272  | 
|
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
273  | 
fun command_reports thy tok =  | 
| 52510 | 274  | 
if Token.is_command tok then  | 
275  | 
let val name = Token.content_of tok in  | 
|
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58918 
diff
changeset
 | 
276  | 
(case lookup_commands thy name of  | 
| 52510 | 277  | 
NONE => []  | 
| 55708 | 278  | 
| SOME cmd => [((Token.pos_of tok, command_markup false (name, cmd)), "")])  | 
| 52510 | 279  | 
end  | 
280  | 
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
 | 
281  | 
|
| 5829 | 282  | 
end;  |