src/Pure/Isar/outer_syntax.ML
author wenzelm
Mon, 09 Aug 2010 18:18:32 +0200
changeset 38253 3d4e521014f7
parent 38236 d8c7be27e01d
child 38422 f96394dba335
permissions -rw-r--r--
Isabelle_Process: separate input fifo for commands (still using the old tty protocol); some partial workarounds for Cygwin;
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
27353
71c4dd53d4cb moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
wenzelm
parents: 26990
diff changeset
     4
The global Isabelle/Isar outer syntax.
71c4dd53d4cb moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
wenzelm
parents: 26990
diff changeset
     5
71c4dd53d4cb moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
wenzelm
parents: 26990
diff changeset
     6
Note: the syntax for files is statically determined at the very
71c4dd53d4cb moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
wenzelm
parents: 26990
diff changeset
     7
beginning; for interactive processing it may change dynamically.
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
     8
*)
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
     9
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    10
signature OUTER_SYNTAX =
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    11
sig
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
    12
  val command: string -> string -> Keyword.T ->
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    13
    (Toplevel.transition -> Toplevel.transition) parser -> unit
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36959
diff changeset
    14
  val markup_command: Thy_Output.markup -> string -> string -> Keyword.T ->
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    15
    (Toplevel.transition -> Toplevel.transition) parser -> unit
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
    16
  val improper_command: string -> string -> Keyword.T ->
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    17
    (Toplevel.transition -> Toplevel.transition) parser -> unit
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    18
  val internal_command: string -> (Toplevel.transition -> Toplevel.transition) parser -> unit
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
    19
  val local_theory': string -> string -> Keyword.T ->
29380
a9ee3475abf4 added local_theory';
wenzelm
parents: 29315
diff changeset
    20
    (bool -> local_theory -> local_theory) parser -> unit
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
    21
  val local_theory: string -> string -> Keyword.T ->
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    22
    (local_theory -> local_theory) parser -> unit
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
    23
  val local_theory_to_proof': string -> string -> Keyword.T ->
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    24
    (bool -> local_theory -> Proof.state) parser -> unit
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
    25
  val local_theory_to_proof: string -> string -> Keyword.T ->
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    26
    (local_theory -> Proof.state) parser -> unit
5883
975c984526b0 tuned names;
wenzelm
parents: 5829
diff changeset
    27
  val print_outer_syntax: unit -> unit
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36955
diff changeset
    28
  val scan: Position.T -> string -> Token.T list
25580
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
    29
  val parse: Position.T -> string -> Toplevel.transition list
26431
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
    30
  val process_file: Path.T -> theory -> theory
26600
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
    31
  type isar
38253
3d4e521014f7 Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
wenzelm
parents: 38236
diff changeset
    32
  val isar: TextIO.instream -> bool -> isar
27839
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
    33
  val prepare_command: Position.T -> string -> Toplevel.transition
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37953
diff changeset
    34
  val load_thy: string -> (unit -> theory) -> Position.T -> string -> (unit -> unit) * theory
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    35
end;
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    36
36953
2af1ad9aa1a3 renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents: 36950
diff changeset
    37
structure Outer_Syntax: OUTER_SYNTAX =
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    38
struct
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    39
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    40
(** outer syntax **)
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    41
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    42
(* command parsers *)
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    43
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    44
datatype command = Command of
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    45
 {comment: string,
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36959
diff changeset
    46
  markup: Thy_Output.markup option,
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    47
  int_only: bool,
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    48
  parse: (Toplevel.transition -> Toplevel.transition) parser};
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    49
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    50
fun make_command comment markup int_only parse =
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    51
  Command {comment = comment, markup = markup, int_only = int_only, parse = parse};
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    52
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    53
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    54
(* parse command *)
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    55
6860
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
    56
local
6199
9b1be867e21a removed load;
wenzelm
parents: 6107
diff changeset
    57
14925
0f86a8a694f8 added read (provides transition names and sources);
wenzelm
parents: 14687
diff changeset
    58
fun terminate false = Scan.succeed ()
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
    59
  | terminate true = Parse.group "end of input" (Scan.option Parse.sync -- Parse.semicolon >> K ());
14925
0f86a8a694f8 added read (provides transition names and sources);
wenzelm
parents: 14687
diff changeset
    60
26620
722cf4fdd4dd eliminated unused trace, read;
wenzelm
parents: 26611
diff changeset
    61
fun body cmd (name, _) =
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
    62
  (case cmd name of
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    63
    SOME (Command {int_only, parse, ...}) =>
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
    64
      Parse.!!! (Scan.prompt (name ^ "# ") (Parse.tags |-- parse >> pair int_only))
37852
a902f158b4fc eliminated old-style sys_error/SYS_ERROR in favour of exception Fail -- after careful checking that there is no overlap with existing handling of that;
wenzelm
parents: 37713
diff changeset
    65
  | NONE => raise Fail ("No parser for outer syntax command " ^ quote name));
6860
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
    66
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
    67
in
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    68
26620
722cf4fdd4dd eliminated unused trace, read;
wenzelm
parents: 26611
diff changeset
    69
fun parse_command do_terminate cmd =
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
    70
  Parse.semicolon >> K NONE ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
    71
  Parse.sync >> K NONE ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
    72
  (Parse.position Parse.command :-- body cmd) --| terminate do_terminate
6860
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
    73
    >> (fn ((name, pos), (int_only, f)) =>
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15224
diff changeset
    74
      SOME (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>
6860
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
    75
        Toplevel.interactive int_only |> f));
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    76
6199
9b1be867e21a removed load;
wenzelm
parents: 6107
diff changeset
    77
end;
9b1be867e21a removed load;
wenzelm
parents: 6107
diff changeset
    78
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    79
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    80
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
    81
(** global outer syntax **)
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    82
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
    83
local
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
    84
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 30573
diff changeset
    85
val global_commands = Unsynchronized.ref (Symtab.empty: command Symtab.table);
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36959
diff changeset
    86
val global_markups = Unsynchronized.ref ([]: (string * Thy_Output.markup) list);
5952
7d4ec8992b23 added commands;
wenzelm
parents: 5923
diff changeset
    87
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    88
fun change_commands f = CRITICAL (fn () =>
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 30573
diff changeset
    89
 (Unsynchronized.change global_commands f;
23939
e543359fe8b6 marked some CRITICAL sections;
wenzelm
parents: 23884
diff changeset
    90
  global_markups :=
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    91
    Symtab.fold (fn (name, Command {markup = SOME m, ...}) => cons (name, m) | _ => I)
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
    92
      (! global_commands) []));
6722
5e82c7196789 added keyword classification;
wenzelm
parents: 6685
diff changeset
    93
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
    94
in
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
    95
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
    96
(* access current syntax *)
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
    97
33223
d27956b4d3b4 non-critical atomic accesses;
wenzelm
parents: 32738
diff changeset
    98
fun get_commands () = ! global_commands;
d27956b4d3b4 non-critical atomic accesses;
wenzelm
parents: 32738
diff changeset
    99
fun get_markups () = ! global_markups;
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   100
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   101
fun get_command () = Symtab.lookup (get_commands ());
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
   102
fun get_syntax () = CRITICAL (fn () => (Keyword.get_lexicons (), get_command ()));
7789
57d20133224e verbatim markup tokens;
wenzelm
parents: 7774
diff changeset
   103
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   104
fun is_markup kind name = AList.lookup (op =) (get_markups ()) name = SOME kind;
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   105
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   106
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   107
(* augment syntax *)
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   108
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   109
fun add_command name kind cmd = CRITICAL (fn () =>
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
   110
 (Keyword.command name kind;
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   111
  if not (Symtab.defined (get_commands ()) name) then ()
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   112
  else warning ("Redefining outer syntax command " ^ quote name);
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   113
  change_commands (Symtab.update (name, cmd))));
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   114
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   115
fun command name comment kind parse =
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   116
  add_command name kind (make_command comment NONE false parse);
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   117
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   118
fun markup_command markup name comment kind parse =
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   119
  add_command name kind (make_command comment (SOME markup) false parse);
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   120
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   121
fun improper_command name comment kind parse =
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   122
  add_command name kind (make_command comment NONE true parse);
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   123
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   124
end;
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   125
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   126
fun internal_command name parse =
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
   127
  command name "(internal)" Keyword.control (parse >> (fn tr => Toplevel.no_timing o tr));
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   128
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   129
26990
a91f7741967a added local_theory command wrappers;
wenzelm
parents: 26881
diff changeset
   130
(* local_theory commands *)
a91f7741967a added local_theory command wrappers;
wenzelm
parents: 26881
diff changeset
   131
a91f7741967a added local_theory command wrappers;
wenzelm
parents: 26881
diff changeset
   132
fun local_theory_command do_print trans name comment kind parse =
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
   133
  command name comment kind (Parse.opt_target -- parse
26990
a91f7741967a added local_theory command wrappers;
wenzelm
parents: 26881
diff changeset
   134
    >> (fn (loc, f) => (if do_print then Toplevel.print else I) o trans loc f));
a91f7741967a added local_theory command wrappers;
wenzelm
parents: 26881
diff changeset
   135
29380
a9ee3475abf4 added local_theory';
wenzelm
parents: 29315
diff changeset
   136
val local_theory' = local_theory_command false Toplevel.local_theory';
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   137
val local_theory = local_theory_command false Toplevel.local_theory;
26990
a91f7741967a added local_theory command wrappers;
wenzelm
parents: 26881
diff changeset
   138
val local_theory_to_proof' = local_theory_command true Toplevel.local_theory_to_proof';
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   139
val local_theory_to_proof = local_theory_command true Toplevel.local_theory_to_proof;
26990
a91f7741967a added local_theory command wrappers;
wenzelm
parents: 26881
diff changeset
   140
a91f7741967a added local_theory command wrappers;
wenzelm
parents: 26881
diff changeset
   141
24872
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
   142
(* inspect syntax *)
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   143
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   144
fun dest_commands () =
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   145
  get_commands () |> Symtab.dest |> sort_wrt #1
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   146
  |> map (fn (name, Command {comment, int_only, ...}) => (name, comment, int_only));
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   147
9223
eb752c2fac22 removed help;
wenzelm
parents: 9213
diff changeset
   148
fun print_outer_syntax () =
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   149
  let
27353
71c4dd53d4cb moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
wenzelm
parents: 26990
diff changeset
   150
    fun pretty_cmd (name, comment, _) =
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   151
      Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   152
    val (int_cmds, cmds) = List.partition #3 (dest_commands ());
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   153
  in
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
   154
    [Pretty.strs ("syntax keywords:" :: map quote (Keyword.dest_keywords ())),
18326
2f57579e618f tuned msg;
wenzelm
parents: 18064
diff changeset
   155
      Pretty.big_list "commands:" (map pretty_cmd cmds),
2f57579e618f tuned msg;
wenzelm
parents: 18064
diff changeset
   156
      Pretty.big_list "interactive-only commands:" (map pretty_cmd int_cmds)]
9223
eb752c2fac22 removed help;
wenzelm
parents: 9213
diff changeset
   157
    |> Pretty.chunks |> Pretty.writeln
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   158
  end;
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   159
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   160
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   161
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   162
(** toplevel parsing **)
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   163
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   164
(* basic sources *)
6860
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
   165
26620
722cf4fdd4dd eliminated unused trace, read;
wenzelm
parents: 26611
diff changeset
   166
fun toplevel_source term do_recover cmd src =
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   167
  let
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   168
    val no_terminator =
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36955
diff changeset
   169
      Scan.unless Parse.semicolon (Scan.one (Token.not_sync andf Token.not_eof));
23682
cf4773532006 nested source: explicit interactive flag for recover avoids duplicate errors;
wenzelm
parents: 23679
diff changeset
   170
    fun recover int =
cf4773532006 nested source: explicit interactive flag for recover avoids duplicate errors;
wenzelm
parents: 23679
diff changeset
   171
      (int, fn _ => Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [NONE]);
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   172
  in
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   173
    src
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36955
diff changeset
   174
    |> Token.source_proper
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36955
diff changeset
   175
    |> Source.source Token.stopper
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
   176
      (Scan.bulk (Parse.$$$ "--" -- Parse.!!! Parse.doc_source >> K NONE || Parse.not_eof >> SOME))
23682
cf4773532006 nested source: explicit interactive flag for recover avoids duplicate errors;
wenzelm
parents: 23679
diff changeset
   177
        (Option.map recover do_recover)
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19060
diff changeset
   178
    |> Source.map_filter I
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36955
diff changeset
   179
    |> Source.source Token.stopper
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33223
diff changeset
   180
        (Scan.bulk (fn xs => Parse.!!! (parse_command term (cmd ())) xs))
23682
cf4773532006 nested source: explicit interactive flag for recover avoids duplicate errors;
wenzelm
parents: 23679
diff changeset
   181
        (Option.map recover do_recover)
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19060
diff changeset
   182
    |> Source.map_filter I
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   183
  end;
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   184
7746
b52c1a632e6a simplified;
wenzelm
parents: 7735
diff changeset
   185
25580
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   186
(* off-line scanning/parsing *)
14925
0f86a8a694f8 added read (provides transition names and sources);
wenzelm
parents: 14687
diff changeset
   187
27839
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   188
fun scan pos str =
16195
wenzelm
parents: 15989
diff changeset
   189
  Source.of_string str
27839
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   190
  |> Symbol.source {do_recover = false}
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36955
diff changeset
   191
  |> Token.source {do_recover = SOME false} Keyword.get_lexicons pos
16195
wenzelm
parents: 15989
diff changeset
   192
  |> Source.exhaust;
wenzelm
parents: 15989
diff changeset
   193
25580
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   194
fun parse pos str =
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   195
  Source.of_string str
27839
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   196
  |> Symbol.source {do_recover = false}
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36955
diff changeset
   197
  |> Token.source {do_recover = SOME false} Keyword.get_lexicons pos
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   198
  |> toplevel_source false NONE get_command
25580
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   199
  |> Source.exhaust;
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   200
14091
ad6ba9c55190 A patch by david aspinall
nipkow
parents: 12943
diff changeset
   201
26431
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   202
(* process file *)
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   203
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   204
fun process_file path thy =
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   205
  let
26881
bb68f50644a9 renamed Position.path to Path.position;
wenzelm
parents: 26700
diff changeset
   206
    val trs = parse (Path.position path) (File.read path);
37953
ddc3b72f9a42 simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents: 37907
diff changeset
   207
    val init = Toplevel.init_theory "" (K thy) Toplevel.empty;
28424
fc6ce1c4d5b7 simplified process_file, eliminated Toplevel.excursion;
wenzelm
parents: 27872
diff changeset
   208
    val result = fold Toplevel.command (init :: trs) Toplevel.toplevel;
fc6ce1c4d5b7 simplified process_file, eliminated Toplevel.excursion;
wenzelm
parents: 27872
diff changeset
   209
  in
fc6ce1c4d5b7 simplified process_file, eliminated Toplevel.excursion;
wenzelm
parents: 27872
diff changeset
   210
    (case (Toplevel.is_theory result, Toplevel.generic_theory_of result) of
fc6ce1c4d5b7 simplified process_file, eliminated Toplevel.excursion;
wenzelm
parents: 27872
diff changeset
   211
      (true, Context.Theory thy') => thy'
fc6ce1c4d5b7 simplified process_file, eliminated Toplevel.excursion;
wenzelm
parents: 27872
diff changeset
   212
    | _ => error "Bad result state: global theory expected")
fc6ce1c4d5b7 simplified process_file, eliminated Toplevel.excursion;
wenzelm
parents: 27872
diff changeset
   213
  end;
26431
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   214
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   215
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   216
(* interactive source of toplevel transformers *)
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   217
26600
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
   218
type isar =
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
   219
  (Toplevel.transition, (Toplevel.transition option,
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36955
diff changeset
   220
    (Token.T, (Token.T option, (Token.T, (Token.T,
30573
49899f26fbd1 de-camelized Symbol_Pos;
wenzelm
parents: 30513
diff changeset
   221
      (Symbol_Pos.T, Position.T * (Symbol.symbol, (string, unit) Source.source)
27770
10d84e124a2f updated type of nested sources;
wenzelm
parents: 27720
diff changeset
   222
  Source.source) Source.source) Source.source) Source.source)
10d84e124a2f updated type of nested sources;
wenzelm
parents: 27720
diff changeset
   223
  Source.source) Source.source) Source.source) Source.source;
26600
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
   224
38253
3d4e521014f7 Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
wenzelm
parents: 38236
diff changeset
   225
fun isar in_stream term : isar =
3d4e521014f7 Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
wenzelm
parents: 38236
diff changeset
   226
  Source.tty in_stream
27839
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   227
  |> Symbol.source {do_recover = true}
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36955
diff changeset
   228
  |> Token.source {do_recover = SOME true} Keyword.get_lexicons Position.none
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   229
  |> toplevel_source term (SOME true) get_command;
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   230
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   231
27839
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   232
(* prepare toplevel commands -- fail-safe *)
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   233
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   234
val not_singleton = "Exactly one command expected";
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   235
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   236
fun prepare_span commands span =
27839
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   237
  let
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36959
diff changeset
   238
    val range_pos = Position.encode_range (Thy_Syntax.span_range span);
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36959
diff changeset
   239
    val toks = Thy_Syntax.span_content span;
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36959
diff changeset
   240
    val _ = List.app Thy_Syntax.report_token toks;
27839
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   241
  in
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   242
    (case Source.exhaust (toplevel_source false NONE (K commands) (Source.of_list toks)) of
37713
c82cf6e11669 Outer_Syntax.prepare_command: disallow control commands here, and consequently in Isar documents;
wenzelm
parents: 37216
diff changeset
   243
      [tr] =>
c82cf6e11669 Outer_Syntax.prepare_command: disallow control commands here, and consequently in Isar documents;
wenzelm
parents: 37216
diff changeset
   244
        if Keyword.is_control (Toplevel.name_of tr) then
c82cf6e11669 Outer_Syntax.prepare_command: disallow control commands here, and consequently in Isar documents;
wenzelm
parents: 37216
diff changeset
   245
          (Toplevel.malformed range_pos "Illegal control command", true)
c82cf6e11669 Outer_Syntax.prepare_command: disallow control commands here, and consequently in Isar documents;
wenzelm
parents: 37216
diff changeset
   246
        else (tr, true)
27839
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   247
    | [] => (Toplevel.ignored range_pos, false)
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   248
    | _ => (Toplevel.malformed range_pos not_singleton, true))
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   249
    handle ERROR msg => (Toplevel.malformed range_pos msg, true)
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   250
  end;
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   251
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37953
diff changeset
   252
fun prepare_unit commands init (cmd, proof, proper_proof) =
28436
4faf705a177d load_thy: more precise treatment of improper cmd or proof (notably 'oops');
wenzelm
parents: 28432
diff changeset
   253
  let
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37953
diff changeset
   254
    val (tr, proper_cmd) = prepare_span commands cmd |>> Toplevel.modify_init init;
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   255
    val proof_trs = map (prepare_span commands) proof |> filter #2 |> map #1;
28436
4faf705a177d load_thy: more precise treatment of improper cmd or proof (notably 'oops');
wenzelm
parents: 28432
diff changeset
   256
  in
4faf705a177d load_thy: more precise treatment of improper cmd or proof (notably 'oops');
wenzelm
parents: 28432
diff changeset
   257
    if proper_cmd andalso proper_proof then [(tr, proof_trs)]
4faf705a177d load_thy: more precise treatment of improper cmd or proof (notably 'oops');
wenzelm
parents: 28432
diff changeset
   258
    else map (rpair []) (if proper_cmd then tr :: proof_trs else proof_trs)
4faf705a177d load_thy: more precise treatment of improper cmd or proof (notably 'oops');
wenzelm
parents: 28432
diff changeset
   259
  end;
28432
944cb67f809a load_thy: Toplevel.excursion based on units of command/proof pairs;
wenzelm
parents: 28424
diff changeset
   260
27839
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   261
fun prepare_command pos str =
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   262
  let val (lexs, commands) = get_syntax () in
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36959
diff changeset
   263
    (case Thy_Syntax.parse_spans lexs pos str of
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   264
      [span] => #1 (prepare_span commands span)
27839
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   265
    | _ => Toplevel.malformed pos not_singleton)
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   266
  end;
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   267
0be8644c45eb Symbol.source/OuterLex.source: more explicit do_recover argument;
wenzelm
parents: 27831
diff changeset
   268
26611
03455add4801 export load_thy -- no backpatching;
wenzelm
parents: 26600
diff changeset
   269
(* load_thy *)
7746
b52c1a632e6a simplified;
wenzelm
parents: 7735
diff changeset
   270
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37953
diff changeset
   271
fun load_thy name init pos text =
7683
e74f43f1d8a3 export token_source;
wenzelm
parents: 7676
diff changeset
   272
  let
29311
4c830711e6f1 added type 'a parser, simplified signature;
wenzelm
parents: 28436
diff changeset
   273
    val (lexs, commands) = get_syntax ();
37873
66d90b2b87bc eliminated old time_use/time_use_thy variants -- timing is implicitly controlled via Output.timing;
wenzelm
parents: 37852
diff changeset
   274
    val time = ! Output.timing;
23866
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   275
17932
677f7bec354e removed old-style theory format;
wenzelm
parents: 17412
diff changeset
   276
    val _ = Present.init_theory name;
27855
b1bf607e06c2 load_thy: no untabify (preserve position information!), present spans instead of verbatim source;
wenzelm
parents: 27839
diff changeset
   277
37902
4e7864f3643d deps_thy/load_thy: store compact text to reduce space by factor 12;
wenzelm
parents: 37873
diff changeset
   278
    val toks = Source.exhausted (Thy_Syntax.token_source lexs pos (Source.of_string text));
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36959
diff changeset
   279
    val spans = Source.exhaust (Thy_Syntax.span_source toks);
38236
d8c7be27e01d explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
wenzelm
parents: 37977
diff changeset
   280
    val _ = List.app Thy_Syntax.report_span spans;  (* FIXME ?? *)
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36959
diff changeset
   281
    val units = Source.exhaust (Thy_Syntax.unit_source (Source.of_list spans))
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37953
diff changeset
   282
      |> Par_List.map (prepare_unit commands init) |> flat;
23866
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   283
27855
b1bf607e06c2 load_thy: no untabify (preserve position information!), present spans instead of verbatim source;
wenzelm
parents: 27839
diff changeset
   284
    val _ = Present.theory_source name
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36959
diff changeset
   285
      (fn () => HTML.html_mode (implode o map Thy_Syntax.present_span) spans);
27855
b1bf607e06c2 load_thy: no untabify (preserve position information!), present spans instead of verbatim source;
wenzelm
parents: 27839
diff changeset
   286
23866
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   287
    val _ = if time then writeln ("\n**** Starting theory " ^ quote name ^ " ****") else ();
37953
ddc3b72f9a42 simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents: 37907
diff changeset
   288
    val (results, thy) = cond_timeit time "" (fn () => Toplevel.excursion units);
23866
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   289
    val _ = if time then writeln ("**** Finished theory " ^ quote name ^ " ****\n") else ();
29428
3ab54b42ded8 load_thy: explicit after_load phase for presentation;
wenzelm
parents: 29380
diff changeset
   290
3ab54b42ded8 load_thy: explicit after_load phase for presentation;
wenzelm
parents: 29380
diff changeset
   291
    fun after_load () =
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36959
diff changeset
   292
      Thy_Output.present_thy (#1 lexs) Keyword.command_tags is_markup (Lazy.force results) toks
29428
3ab54b42ded8 load_thy: explicit after_load phase for presentation;
wenzelm
parents: 29380
diff changeset
   293
      |> Buffer.content
3ab54b42ded8 load_thy: explicit after_load phase for presentation;
wenzelm
parents: 29380
diff changeset
   294
      |> Present.theory_output name;
37953
ddc3b72f9a42 simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents: 37907
diff changeset
   295
  in (after_load, thy) end;
23866
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   296
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   297
end;
36953
2af1ad9aa1a3 renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents: 36950
diff changeset
   298