src/Pure/Isar/outer_syntax.ML
author wenzelm
Fri, 08 Jan 2021 14:40:04 +0100
changeset 73098 8a20737e4ebf
parent 72434 cc27cf7e51c6
child 73106 3df45de0c079
permissions -rw-r--r--
support more command positions, analogous to Command.core_range in Isabelle/Scala;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/outer_syntax.ML
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
     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
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
     5
*)
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
     6
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
     7
signature OUTER_SYNTAX =
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
     8
sig
58930
13d3eb07a17a tuned signature;
wenzelm
parents: 58928
diff changeset
     9
  val help: theory -> string list -> unit
13d3eb07a17a tuned signature;
wenzelm
parents: 58928
diff changeset
    10
  val print_commands: theory -> unit
59935
343905de27b1 clarified command keyword markup;
wenzelm
parents: 59932
diff changeset
    11
  type command_keyword = string * Position.T
343905de27b1 clarified command keyword markup;
wenzelm
parents: 59932
diff changeset
    12
  val command: command_keyword -> string ->
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    13
    (Toplevel.transition -> Toplevel.transition) parser -> unit
63273
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63022
diff changeset
    14
  val maybe_begin_local_theory: command_keyword -> string ->
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63022
diff changeset
    15
    (local_theory -> local_theory) parser -> (theory -> local_theory) parser -> unit
59935
343905de27b1 clarified command keyword markup;
wenzelm
parents: 59932
diff changeset
    16
  val local_theory': command_keyword -> string ->
29380
a9ee3475abf4 added local_theory';
wenzelm
parents: 29315
diff changeset
    17
    (bool -> local_theory -> local_theory) parser -> unit
59935
343905de27b1 clarified command keyword markup;
wenzelm
parents: 59932
diff changeset
    18
  val local_theory: command_keyword -> string ->
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    19
    (local_theory -> local_theory) parser -> unit
59935
343905de27b1 clarified command keyword markup;
wenzelm
parents: 59932
diff changeset
    20
  val local_theory_to_proof': command_keyword -> string ->
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    21
    (bool -> local_theory -> Proof.state) parser -> unit
59935
343905de27b1 clarified command keyword markup;
wenzelm
parents: 59932
diff changeset
    22
  val local_theory_to_proof: command_keyword -> string ->
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    23
    (local_theory -> Proof.state) parser -> unit
60095
35f626b11422 more explicit bootstrap_thy;
wenzelm
parents: 60073
diff changeset
    24
  val bootstrap: bool Config.T
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
68184
6c693b2700b3 support for dynamic document output while editing;
wenzelm
parents: 67499
diff changeset
    26
  val make_span: Token.T list -> Command_Span.span
69891
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
    27
  val parse_span: theory -> (unit -> theory) -> Token.T list -> Toplevel.transition
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
    28
  val parse_text: theory -> (unit -> theory) -> Position.T -> string -> Toplevel.transition list
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    29
  val command_reports: theory -> Token.T -> Position.report_text list
63274
wenzelm
parents: 63273
diff changeset
    30
  val check_command: Proof.context -> command_keyword -> string
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    31
end;
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    32
36953
2af1ad9aa1a3 renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents: 36950
diff changeset
    33
structure Outer_Syntax: OUTER_SYNTAX =
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    34
struct
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    35
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    36
(** outer syntax **)
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    37
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    38
(* errors *)
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_command msg name ps =
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    41
  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
    42
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    43
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
    44
  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
    45
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    46
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    47
(* command parsers *)
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    48
59924
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59923
diff changeset
    49
datatype command_parser =
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59923
diff changeset
    50
  Parser of (Toplevel.transition -> Toplevel.transition) parser |
60691
0568c7a2b5db tuned according to a81dc82ecba3;
wenzelm
parents: 60095
diff changeset
    51
  Restricted_Parser of
59939
7d46aa03696e support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents: 59935
diff changeset
    52
    (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
    53
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    54
datatype command = Command of
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    55
 {comment: string,
59924
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59923
diff changeset
    56
  command_parser: command_parser,
48647
a5144c4c26a2 report commands as formal entities, with def/ref positions;
wenzelm
parents: 48646
diff changeset
    57
  pos: Position.T,
a5144c4c26a2 report commands as formal entities, with def/ref positions;
wenzelm
parents: 48646
diff changeset
    58
  id: serial};
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    59
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    60
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
    61
59924
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59923
diff changeset
    62
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
    63
  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
    64
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    65
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
    66
48647
a5144c4c26a2 report commands as formal entities, with def/ref positions;
wenzelm
parents: 48646
diff changeset
    67
fun command_markup def (name, Command {pos, id, ...}) =
a5144c4c26a2 report commands as formal entities, with def/ref positions;
wenzelm
parents: 48646
diff changeset
    68
  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
    69
    (Markup.entity Markup.commandN name);
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    70
50213
7b73c0509835 refined outer syntax 'help' command;
wenzelm
parents: 50201
diff changeset
    71
fun pretty_command (cmd as (name, Command {comment, ...})) =
7b73c0509835 refined outer syntax 'help' command;
wenzelm
parents: 50201
diff changeset
    72
  Pretty.block
7b73c0509835 refined outer syntax 'help' command;
wenzelm
parents: 50201
diff changeset
    73
    (Pretty.marks_str
50450
358b6020f8b6 generalized notion of active area, where sendback is just one application;
wenzelm
parents: 50215
diff changeset
    74
      ([Active.make_markup Markup.sendbackN {implicit = true, properties = [Markup.padding_line]},
50215
97959912840a more general sendback properties;
wenzelm
parents: 50213
diff changeset
    75
        command_markup false cmd], name) :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment);
50213
7b73c0509835 refined outer syntax 'help' command;
wenzelm
parents: 50201
diff changeset
    76
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    77
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58934
diff changeset
    78
(* theory data *)
43711
608d1b451f67 less stateful outer_syntax;
wenzelm
parents: 43621
diff changeset
    79
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    80
structure Data = Theory_Data
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    81
(
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58934
diff changeset
    82
  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
    83
  val empty = Symtab.empty;
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    84
  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
    85
  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
    86
    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
    87
      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
    88
      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
    89
);
43711
608d1b451f67 less stateful outer_syntax;
wenzelm
parents: 43621
diff changeset
    90
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58934
diff changeset
    91
val get_commands = Data.get;
60924
610794dff23c tuned signature, in accordance to sortBy in Scala;
wenzelm
parents: 60693
diff changeset
    92
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
    93
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
    94
58930
13d3eb07a17a tuned signature;
wenzelm
parents: 58928
diff changeset
    95
fun help thy pats =
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    96
  dest_commands thy
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    97
  |> 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
    98
  |> map pretty_command
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    99
  |> Pretty.writeln_chunks;
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
   100
58930
13d3eb07a17a tuned signature;
wenzelm
parents: 58928
diff changeset
   101
fun print_commands thy =
43711
608d1b451f67 less stateful outer_syntax;
wenzelm
parents: 43621
diff changeset
   102
  let
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
   103
    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
   104
    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
   105
    val commands = dest_commands thy;
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
   106
  in
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
   107
    [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
   108
      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
   109
    |> Pretty.writeln_chunks
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
   110
  end;
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
   111
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
   112
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58934
diff changeset
   113
(* maintain commands *)
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
   114
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58934
diff changeset
   115
fun add_command name cmd thy =
63022
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   116
  if member (op =) Thy_Header.bootstrap_thys (Context.theory_name thy) then thy
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   117
  else
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   118
    let
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   119
      val _ =
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   120
        Keyword.is_command (Thy_Header.get_keywords thy) name orelse
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   121
          err_command "Undeclared outer syntax command " name [command_pos cmd];
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   122
      val _ =
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   123
        (case lookup_commands thy name of
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   124
          NONE => ()
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   125
        | SOME cmd' => err_dup_command name [command_pos cmd, command_pos cmd']);
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   126
      val _ =
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   127
        Context_Position.report_generic (Context.the_generic_context ())
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   128
          (command_pos cmd) (command_markup true (name, cmd));
785a59235a15 more IDE support for Isabelle/Pure bootstrap;
wenzelm
parents: 62876
diff changeset
   129
    in Data.map (Symtab.update (name, cmd)) thy end;
43711
608d1b451f67 less stateful outer_syntax;
wenzelm
parents: 43621
diff changeset
   130
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
   131
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
   132
  let
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58934
diff changeset
   133
    val command_keywords =
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58934
diff changeset
   134
      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
   135
    val _ =
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58934
diff changeset
   136
      (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
   137
        [] => ()
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
   138
      | 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
   139
  in NONE end));
43711
608d1b451f67 less stateful outer_syntax;
wenzelm
parents: 43621
diff changeset
   140
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   141
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
   142
(* implicit theory setup *)
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46958
diff changeset
   143
59935
343905de27b1 clarified command keyword markup;
wenzelm
parents: 59932
diff changeset
   144
type command_keyword = string * Position.T;
5952
7d4ec8992b23 added commands;
wenzelm
parents: 5923
diff changeset
   145
59924
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59923
diff changeset
   146
fun raw_command (name, pos) comment command_parser =
59932
74872a13f628 support local command setup;
wenzelm
parents: 59924
diff changeset
   147
  let val setup = add_command name (new_command comment command_parser pos)
74872a13f628 support local command setup;
wenzelm
parents: 59924
diff changeset
   148
  in Context.>> (Context.mapping setup (Local_Theory.background_theory setup)) end;
26990
a91f7741967a added local_theory command wrappers;
wenzelm
parents: 26881
diff changeset
   149
59924
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59923
diff changeset
   150
fun command (name, pos) comment parse =
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59923
diff changeset
   151
  raw_command (name, pos) comment (Parser parse);
59923
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59083
diff changeset
   152
63273
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63022
diff changeset
   153
fun maybe_begin_local_theory command_keyword comment parse_local parse_global =
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63022
diff changeset
   154
  raw_command command_keyword comment
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63022
diff changeset
   155
    (Restricted_Parser (fn restricted =>
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63022
diff changeset
   156
      Parse.opt_target -- parse_local
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63022
diff changeset
   157
        >> (fn (target, f) => Toplevel.local_theory restricted target f) ||
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63022
diff changeset
   158
      (if is_some restricted then Scan.fail
72434
cc27cf7e51c6 consolidated terminology
haftmann
parents: 70134
diff changeset
   159
       else parse_global >> Toplevel.begin_main_target true)));
63273
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63022
diff changeset
   160
59935
343905de27b1 clarified command keyword markup;
wenzelm
parents: 59932
diff changeset
   161
fun local_theory_command trans command_keyword comment parse =
343905de27b1 clarified command keyword markup;
wenzelm
parents: 59932
diff changeset
   162
  raw_command command_keyword comment
60691
0568c7a2b5db tuned according to a81dc82ecba3;
wenzelm
parents: 60095
diff changeset
   163
    (Restricted_Parser (fn restricted =>
0568c7a2b5db tuned according to a81dc82ecba3;
wenzelm
parents: 60095
diff changeset
   164
      Parse.opt_target -- parse >> (fn (target, f) => trans restricted target f)));
26990
a91f7741967a added local_theory command wrappers;
wenzelm
parents: 26881
diff changeset
   165
56895
f058120aaad4 discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
wenzelm
parents: 56334
diff changeset
   166
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
   167
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
   168
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
   169
val local_theory_to_proof = local_theory_command Toplevel.local_theory_to_proof;
26990
a91f7741967a added local_theory command wrappers;
wenzelm
parents: 26881
diff changeset
   170
a91f7741967a added local_theory command wrappers;
wenzelm
parents: 26881
diff changeset
   171
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   172
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   173
(** toplevel parsing **)
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   174
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   175
(* parse spans *)
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   176
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   177
local
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   178
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   179
fun ship span =
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   180
  let
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   181
    val kind =
68729
3a02b424d5fb clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
wenzelm
parents: 68184
diff changeset
   182
      if forall Token.is_ignored span then Command_Span.Ignored_Span
59924
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59923
diff changeset
   183
      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
   184
      else
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59923
diff changeset
   185
        (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
   186
          NONE => Command_Span.Malformed_Span
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59923
diff changeset
   187
        | 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
   188
  in cons (Command_Span.Span (kind, span)) end;
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   189
68729
3a02b424d5fb clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
wenzelm
parents: 68184
diff changeset
   190
fun flush (result, content, ignored) =
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   191
  result
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   192
  |> not (null content) ? ship (rev content)
68729
3a02b424d5fb clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
wenzelm
parents: 68184
diff changeset
   193
  |> not (null ignored) ? ship (rev ignored);
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   194
68729
3a02b424d5fb clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
wenzelm
parents: 68184
diff changeset
   195
fun parse tok (result, content, ignored) =
3a02b424d5fb clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
wenzelm
parents: 68184
diff changeset
   196
  if Token.is_ignored tok then (result, content, tok :: ignored)
59939
7d46aa03696e support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents: 59935
diff changeset
   197
  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
   198
    Token.is_command tok andalso
59939
7d46aa03696e support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents: 59935
diff changeset
   199
      (not (exists Token.is_command_modifier content) orelse exists Token.is_command content)
68729
3a02b424d5fb clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
wenzelm
parents: 68184
diff changeset
   200
  then (flush (result, content, ignored), [tok], [])
3a02b424d5fb clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
wenzelm
parents: 68184
diff changeset
   201
  else (result, tok :: (ignored @ content), []);
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   202
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   203
in
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   204
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   205
fun parse_spans toks =
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   206
  fold parse toks ([], [], []) |> flush |> rev;
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   207
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   208
end;
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57623
diff changeset
   209
68184
6c693b2700b3 support for dynamic document output while editing;
wenzelm
parents: 67499
diff changeset
   210
fun make_span toks =
6c693b2700b3 support for dynamic document output while editing;
wenzelm
parents: 67499
diff changeset
   211
  (case parse_spans toks of
6c693b2700b3 support for dynamic document output while editing;
wenzelm
parents: 67499
diff changeset
   212
    [span] => span
6c693b2700b3 support for dynamic document output while editing;
wenzelm
parents: 67499
diff changeset
   213
  | _ => Command_Span.Span (Command_Span.Malformed_Span, toks));
6c693b2700b3 support for dynamic document output while editing;
wenzelm
parents: 67499
diff changeset
   214
14091
ad6ba9c55190 A patch by david aspinall
nipkow
parents: 12943
diff changeset
   215
69891
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   216
(* parse commands *)
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   217
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   218
val bootstrap = Config.declare_bool ("Outer_Syntax.bootstrap", \<^here>) (K true);
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   219
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   220
local
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   221
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   222
val before_command =
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   223
  Scan.option (Parse.position (Parse.private >> K true || Parse.qualified >> K false));
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   224
73098
8a20737e4ebf support more command positions, analogous to Command.core_range in Isabelle/Scala;
wenzelm
parents: 72434
diff changeset
   225
fun parse_command thy core_range markers =
69891
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   226
  Scan.ahead (before_command |-- Parse.position Parse.command) :|-- (fn (name, pos) =>
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   227
    let
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   228
      val keywords = Thy_Header.get_keywords thy;
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   229
      val tr0 =
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   230
        Toplevel.empty
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   231
        |> Toplevel.name name
73098
8a20737e4ebf support more command positions, analogous to Command.core_range in Isabelle/Scala;
wenzelm
parents: 72434
diff changeset
   232
        |> Toplevel.positions pos core_range
69891
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   233
        |> Keyword.is_proof_open keywords name ? Toplevel.skip_proof_open
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   234
        |> Keyword.is_proof_close keywords name ? Toplevel.skip_proof_close;
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   235
      val command_markers =
70134
e69f00f4b897 support "tag" marker with scope;
wenzelm
parents: 69891
diff changeset
   236
        Parse.command |-- Document_Source.old_tags >>
69891
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   237
          (fn tags => Toplevel.markers (map Document_Marker.legacy_tag tags @ markers) tr0);
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   238
    in
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   239
      (case lookup_commands thy name of
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   240
        SOME (Command {command_parser = Parser parse, ...}) =>
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   241
          Parse.!!! (command_markers -- parse) >> (op |>)
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   242
      | SOME (Command {command_parser = Restricted_Parser parse, ...}) =>
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   243
          before_command :|-- (fn restricted =>
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   244
            Parse.!!! (command_markers -- parse restricted)) >> (op |>)
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   245
      | NONE =>
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   246
          Scan.fail_with (fn _ => fn _ =>
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   247
            let
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   248
              val msg =
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   249
                if Config.get_global thy bootstrap
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   250
                then "missing theory context for command "
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   251
                else "undefined command ";
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   252
            in msg ^ quote (Markup.markup Markup.keyword1 name) end))
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   253
    end);
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   254
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   255
in
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   256
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   257
fun parse_span thy init span =
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   258
  let
73098
8a20737e4ebf support more command positions, analogous to Command.core_range in Isabelle/Scala;
wenzelm
parents: 72434
diff changeset
   259
    val full_range = Token.range_of span;
69891
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   260
    val core_range = Token.core_range_of span;
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   261
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   262
    val markers = map Token.input_of (filter Token.is_document_marker span);
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   263
    fun parse () =
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   264
      filter Token.is_proper span
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   265
      |> Source.of_list
73098
8a20737e4ebf support more command positions, analogous to Command.core_range in Isabelle/Scala;
wenzelm
parents: 72434
diff changeset
   266
      |> Source.source Token.stopper
8a20737e4ebf support more command positions, analogous to Command.core_range in Isabelle/Scala;
wenzelm
parents: 72434
diff changeset
   267
          (Scan.bulk (fn xs => Parse.!!! (parse_command thy core_range markers) xs))
69891
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   268
      |> Source.exhaust;
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   269
  in
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   270
    (case parse () of
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   271
      [tr] => Toplevel.modify_init init tr
73098
8a20737e4ebf support more command positions, analogous to Command.core_range in Isabelle/Scala;
wenzelm
parents: 72434
diff changeset
   272
    | [] => Toplevel.ignored full_range
8a20737e4ebf support more command positions, analogous to Command.core_range in Isabelle/Scala;
wenzelm
parents: 72434
diff changeset
   273
    | _ => Toplevel.malformed core_range "Exactly one command expected")
8a20737e4ebf support more command positions, analogous to Command.core_range in Isabelle/Scala;
wenzelm
parents: 72434
diff changeset
   274
    handle ERROR msg => Toplevel.malformed core_range msg
69891
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   275
  end;
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   276
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   277
fun parse_text thy init pos text =
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   278
  Symbol_Pos.explode (text, pos)
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   279
  |> Token.tokenize (Thy_Header.get_keywords thy) {strict = false}
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   280
  |> parse_spans
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   281
  |> map (Command_Span.content #> parse_span thy init);
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   282
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   283
end;
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   284
def3ec9cdb7e document markers are formal comments, and may thus occur anywhere in the command-span;
wenzelm
parents: 69887
diff changeset
   285
61618
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   286
(* check commands *)
48647
a5144c4c26a2 report commands as formal entities, with def/ref positions;
wenzelm
parents: 48646
diff changeset
   287
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
   288
fun command_reports thy tok =
52510
a4a102237ded tuned signature;
wenzelm
parents: 52509
diff changeset
   289
  if Token.is_command tok then
a4a102237ded tuned signature;
wenzelm
parents: 52509
diff changeset
   290
    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
   291
      (case lookup_commands thy name of
52510
a4a102237ded tuned signature;
wenzelm
parents: 52509
diff changeset
   292
        NONE => []
55708
f4b114070675 tuned signature;
wenzelm
parents: 55489
diff changeset
   293
      | SOME cmd => [((Token.pos_of tok, command_markup false (name, cmd)), "")])
52510
a4a102237ded tuned signature;
wenzelm
parents: 52509
diff changeset
   294
    end
a4a102237ded tuned signature;
wenzelm
parents: 52509
diff changeset
   295
  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
   296
61618
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   297
fun check_command ctxt (name, pos) =
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   298
  let
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   299
    val thy = Proof_Context.theory_of ctxt;
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   300
    val keywords = Thy_Header.get_keywords thy;
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   301
  in
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   302
    if Keyword.is_command keywords name then
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   303
      let
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   304
        val markup =
67495
90d760fa8f34 clarified operations;
wenzelm
parents: 67446
diff changeset
   305
          Token.explode0 keywords name
61618
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   306
          |> maps (command_reports thy)
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   307
          |> map (#2 o #1);
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   308
        val _ = Context_Position.reports ctxt (map (pair pos) markup);
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   309
      in name end
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   310
    else
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   311
      let
69289
bf6937af7fe8 clarified signature;
wenzelm
parents: 68729
diff changeset
   312
        val completion_report =
bf6937af7fe8 clarified signature;
wenzelm
parents: 68729
diff changeset
   313
          Completion.make_report (name, pos)
61618
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   314
            (fn completed =>
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   315
              Keyword.dest_commands keywords
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   316
              |> filter completed
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   317
              |> sort_strings
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   318
              |> map (fn a => (a, (Markup.commandN, a))));
69289
bf6937af7fe8 clarified signature;
wenzelm
parents: 68729
diff changeset
   319
      in error ("Bad command " ^ quote name ^ Position.here pos ^ completion_report) end
61618
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   320
  end;
27af754f50ca more thorough check_command, including completion;
wenzelm
parents: 61579
diff changeset
   321
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 61618
diff changeset
   322
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 61618
diff changeset
   323
(* 'ML' command -- required for bootstrapping Isar *)
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 61618
diff changeset
   324
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 61618
diff changeset
   325
val _ =
64556
851ae0e7b09c more symbols;
wenzelm
parents: 63936
diff changeset
   326
  command ("ML", \<^here>) "ML text within theory or local theory"
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 61618
diff changeset
   327
    (Parse.ML_source >> (fn source =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 61618
diff changeset
   328
      Toplevel.generic_theory
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 61618
diff changeset
   329
        (ML_Context.exec (fn () =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 61618
diff changeset
   330
            ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags) source) #>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 61618
diff changeset
   331
          Local_Theory.propagate_ml_env)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 61618
diff changeset
   332
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   333
end;