src/Pure/Isar/outer_syntax.ML
author wenzelm
Thu, 10 Apr 2008 14:53:27 +0200
changeset 26611 03455add4801
parent 26600 f11515535c83
child 26620 722cf4fdd4dd
permissions -rw-r--r--
export load_thy -- no backpatching;
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
    ID:         $Id$
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
     4
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
     5
The global Isabelle/Isar outer syntax. Note: the syntax for files is
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
     6
statically determined at the very beginning; for interactive processing
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
     7
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
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    12
  type parser_fn = OuterLex.token list ->
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    13
    (Toplevel.transition -> Toplevel.transition) * OuterLex.token list
23722
a4af559708ab export get_lexicons;
wenzelm
parents: 23682
diff changeset
    14
  val get_lexicons: unit -> Scan.lexicon * Scan.lexicon
23796
f37ff1f39fe9 exported command_keyword;
wenzelm
parents: 23722
diff changeset
    15
  val command_keyword: string -> OuterKeyword.T option
14687
e089757b952a added is_keyword;
wenzelm
parents: 14091
diff changeset
    16
  val is_keyword: string -> bool
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    17
  val keywords: string list -> unit
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    18
  val command: string -> string -> OuterKeyword.T -> parser_fn -> unit
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    19
  val markup_command: ThyOutput.markup -> string -> string -> OuterKeyword.T -> parser_fn -> unit
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    20
  val improper_command: string -> string -> OuterKeyword.T -> parser_fn -> unit
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
    21
  val dest_keywords: unit -> string list
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
    22
  val dest_parsers: unit -> (string * string * string * bool) list
5883
975c984526b0 tuned names;
wenzelm
parents: 5829
diff changeset
    23
  val print_outer_syntax: unit -> unit
24872
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
    24
  val report: unit -> unit
19060
c814a7856121 check_text: Toplevel.node option;
wenzelm
parents: 18717
diff changeset
    25
  val check_text: string * Position.T -> Toplevel.node option -> unit
16195
wenzelm
parents: 15989
diff changeset
    26
  val scan: string -> OuterLex.token list
wenzelm
parents: 15989
diff changeset
    27
  val read: OuterLex.token list -> (string * OuterLex.token list * Toplevel.transition) list
25580
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
    28
  val parse: Position.T -> string -> Toplevel.transition list
26431
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
    29
  val process_file: Path.T -> theory -> theory
26600
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
    30
  type isar
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
    31
  val isar: bool -> isar
26611
03455add4801 export load_thy -- no backpatching;
wenzelm
parents: 26600
diff changeset
    32
  val load_thy: Path.T -> string -> Position.T -> string list -> bool -> unit
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    33
end;
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    34
26600
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
    35
structure OuterSyntax: OUTER_SYNTAX =
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    36
struct
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    37
7750
e63416829538 added markup_command;
wenzelm
parents: 7746
diff changeset
    38
structure T = OuterLex;
6860
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
    39
structure P = OuterParse;
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
    40
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    41
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    42
(** outer syntax **)
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    43
24872
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
    44
(* diagnostics *)
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
    45
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
    46
fun report_keyword name =
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
    47
  Pretty.markup (Markup.keyword_decl name)
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
    48
    [Pretty.str ("Outer syntax keyword: " ^ quote name)];
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
    49
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
    50
fun report_command name kind =
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
    51
  Pretty.markup (Markup.command_decl name kind)
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
    52
    [Pretty.str ("Outer syntax command: " ^ quote name ^ " (" ^ kind ^ ")")];
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
    53
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
    54
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    55
(* parsers *)
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    56
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    57
type parser_fn = T.token list -> (Toplevel.transition -> Toplevel.transition) * T.token list;
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    58
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    59
datatype parser = Parser of
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    60
 {comment: string,
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    61
  kind: OuterKeyword.T,
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    62
  markup: ThyOutput.markup option,
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    63
  int_only: bool,
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    64
  parse: parser_fn};
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    65
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    66
fun make_parser comment kind markup int_only parse =
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    67
  Parser {comment = comment, kind = kind, markup = markup, int_only = int_only, parse = parse};
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    68
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    69
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    70
(* parse command *)
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    71
6860
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
    72
local
6199
9b1be867e21a removed load;
wenzelm
parents: 6107
diff changeset
    73
14925
0f86a8a694f8 added read (provides transition names and sources);
wenzelm
parents: 14687
diff changeset
    74
fun terminate false = Scan.succeed ()
0f86a8a694f8 added read (provides transition names and sources);
wenzelm
parents: 14687
diff changeset
    75
  | terminate true = P.group "end of input" (Scan.option P.sync -- P.semicolon >> K ());
0f86a8a694f8 added read (provides transition names and sources);
wenzelm
parents: 14687
diff changeset
    76
0f86a8a694f8 added read (provides transition names and sources);
wenzelm
parents: 14687
diff changeset
    77
fun trace false parse = parse
0f86a8a694f8 added read (provides transition names and sources);
wenzelm
parents: 14687
diff changeset
    78
  | trace true parse = Scan.trace parse >> (fn (f, toks) => f o Toplevel.source toks);
0f86a8a694f8 added read (provides transition names and sources);
wenzelm
parents: 14687
diff changeset
    79
17071
f753d6dd9bd0 moved structure Keyword to OuterKeyword (Isar/outer_keyword.ML);
wenzelm
parents: 16894
diff changeset
    80
fun body cmd do_trace (name, _) =
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
    81
  (case cmd name of
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    82
    SOME (Parser {int_only, parse, ...}) =>
17118
1ff59b7b35b7 fixed command prompt (was broken due to P.tags);
wenzelm
parents: 17071
diff changeset
    83
      P.!!! (Scan.prompt (name ^ "# ") (trace do_trace (P.tags |-- parse) >> pair int_only))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15224
diff changeset
    84
  | NONE => sys_error ("no parser for outer syntax command " ^ quote name));
6860
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
    85
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
    86
in
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    87
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
    88
fun parse_command do_terminate do_trace cmd =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15224
diff changeset
    89
  P.semicolon >> K NONE ||
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15224
diff changeset
    90
  P.sync >> K NONE ||
17118
1ff59b7b35b7 fixed command prompt (was broken due to P.tags);
wenzelm
parents: 17071
diff changeset
    91
  (P.position P.command :-- body cmd do_trace) --| terminate do_terminate
6860
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
    92
    >> (fn ((name, pos), (int_only, f)) =>
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15224
diff changeset
    93
      SOME (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>
6860
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
    94
        Toplevel.interactive int_only |> f));
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    95
6199
9b1be867e21a removed load;
wenzelm
parents: 6107
diff changeset
    96
end;
9b1be867e21a removed load;
wenzelm
parents: 6107
diff changeset
    97
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    98
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
    99
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   100
(** global outer syntax **)
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   101
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   102
local
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   103
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   104
val global_lexicons = ref (Scan.empty_lexicon, Scan.empty_lexicon);
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   105
val global_parsers = ref (Symtab.empty: parser Symtab.table);
22120
8424ef945cb5 renamed IsarOutput to ThyOutput;
wenzelm
parents: 22086
diff changeset
   106
val global_markups = ref ([]: (string * ThyOutput.markup) list);
5952
7d4ec8992b23 added commands;
wenzelm
parents: 5923
diff changeset
   107
23939
e543359fe8b6 marked some CRITICAL sections;
wenzelm
parents: 23884
diff changeset
   108
fun change_lexicons f = CRITICAL (fn () =>
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   109
  let val lexs = f (! global_lexicons) in
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   110
    (case (op inter_string) (pairself Scan.dest_lexicon lexs) of
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   111
      [] => global_lexicons := lexs
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   112
    | bads => error ("Clash of outer syntax commands and keywords: " ^ commas_quote bads))
23939
e543359fe8b6 marked some CRITICAL sections;
wenzelm
parents: 23884
diff changeset
   113
  end);
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   114
23939
e543359fe8b6 marked some CRITICAL sections;
wenzelm
parents: 23884
diff changeset
   115
fun change_parsers f = CRITICAL (fn () =>
e543359fe8b6 marked some CRITICAL sections;
wenzelm
parents: 23884
diff changeset
   116
 (change global_parsers f;
e543359fe8b6 marked some CRITICAL sections;
wenzelm
parents: 23884
diff changeset
   117
  global_markups :=
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   118
    Symtab.fold (fn (name, Parser {markup = SOME m, ...}) => cons (name, m) | _ => I)
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   119
      (! global_parsers) []));
6722
5e82c7196789 added keyword classification;
wenzelm
parents: 6685
diff changeset
   120
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   121
in
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   122
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   123
(* access current syntax *)
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   124
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   125
fun get_lexicons () = CRITICAL (fn () => ! global_lexicons);
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   126
fun get_parsers () = CRITICAL (fn () => ! global_parsers);
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   127
fun get_markups () = CRITICAL (fn () => ! global_markups);
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   128
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   129
fun get_parser () = Symtab.lookup (get_parsers ());
7789
57d20133224e verbatim markup tokens;
wenzelm
parents: 7774
diff changeset
   130
23796
f37ff1f39fe9 exported command_keyword;
wenzelm
parents: 23722
diff changeset
   131
fun command_keyword name =
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   132
  (case Symtab.lookup (get_parsers ()) name of
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   133
    SOME (Parser {kind, ...}) => SOME kind
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   134
  | NONE => NONE);
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   135
26600
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
   136
fun command_tags name = these (Option.map OuterKeyword.tags_of (command_keyword name));
17071
f753d6dd9bd0 moved structure Keyword to OuterKeyword (Isar/outer_keyword.ML);
wenzelm
parents: 16894
diff changeset
   137
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   138
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
   139
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   140
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   141
(* augment syntax *)
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   142
24872
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
   143
fun keywords names =
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
   144
 (change_lexicons (apfst (Scan.extend_lexicon (map Symbol.explode names)));
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
   145
  List.app (Pretty.writeln o report_keyword) names);
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   146
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   147
24872
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
   148
fun add_parser (name, parser as Parser {kind, ...}) =
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   149
 (if not (Symtab.defined (get_parsers ()) name) then ()
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   150
  else warning ("Redefining outer syntax command " ^ quote name);
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   151
  change_parsers (Symtab.update (name, parser));
24872
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
   152
  change_lexicons (apsnd (Scan.extend_lexicon [Symbol.explode name]));
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
   153
  Pretty.writeln (report_command name (OuterKeyword.kind_of kind)));
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   154
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   155
fun command name comment kind parse =
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   156
  add_parser (name, make_parser comment kind NONE false parse);
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   157
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   158
fun markup_command markup name comment kind parse =
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   159
  add_parser (name, make_parser comment kind (SOME markup) false parse);
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   160
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   161
fun improper_command name comment kind parse =
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   162
  add_parser (name, make_parser comment kind NONE true parse);
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   163
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   164
end;
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   165
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   166
24872
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
   167
(* inspect syntax *)
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   168
14687
e089757b952a added is_keyword;
wenzelm
parents: 14091
diff changeset
   169
fun is_keyword s = Scan.is_literal (#1 (get_lexicons ())) (Symbol.explode s);
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   170
fun dest_keywords () = Scan.dest_lexicon (#1 (get_lexicons ()));
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   171
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   172
fun dest_parsers () =
16727
e264077b68a7 dest_parsers: sort result;
wenzelm
parents: 16195
diff changeset
   173
  get_parsers () |> Symtab.dest |> sort_wrt #1
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   174
  |> map (fn (name, Parser {comment, kind, int_only, ...}) =>
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   175
    (name, comment, OuterKeyword.kind_of kind, int_only));
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   176
9223
eb752c2fac22 removed help;
wenzelm
parents: 9213
diff changeset
   177
fun print_outer_syntax () =
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   178
  let
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   179
    fun pretty_cmd (name, comment, _, _) =
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   180
      Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   181
    val (int_cmds, cmds) = List.partition #4 (dest_parsers ());
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   182
  in
8720
840c75ab2a7f Pretty.chunks;
wenzelm
parents: 8660
diff changeset
   183
    [Pretty.strs ("syntax keywords:" :: map quote (dest_keywords ())),
18326
2f57579e618f tuned msg;
wenzelm
parents: 18064
diff changeset
   184
      Pretty.big_list "commands:" (map pretty_cmd cmds),
2f57579e618f tuned msg;
wenzelm
parents: 18064
diff changeset
   185
      Pretty.big_list "interactive-only commands:" (map pretty_cmd int_cmds)]
9223
eb752c2fac22 removed help;
wenzelm
parents: 9213
diff changeset
   186
    |> Pretty.chunks |> Pretty.writeln
7026
69724548fad1 separate command tokens;
wenzelm
parents: 6868
diff changeset
   187
  end;
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   188
24872
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
   189
fun report () =
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
   190
  (map report_keyword (dest_keywords ()) @
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
   191
    map (fn (name, _, kind, _) => report_command name kind) (dest_parsers ()))
7fd1aa6671a4 report on keyword/command declarations;
wenzelm
parents: 24868
diff changeset
   192
  |> Pretty.chunks |> Pretty.writeln;
7367
a79d4683fadf print_help;
wenzelm
parents: 7333
diff changeset
   193
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   194
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   195
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   196
(** toplevel parsing **)
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   197
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   198
(* basic sources *)
6860
8dc6a1e6fa13 added sync;
wenzelm
parents: 6733
diff changeset
   199
17071
f753d6dd9bd0 moved structure Keyword to OuterKeyword (Isar/outer_keyword.ML);
wenzelm
parents: 16894
diff changeset
   200
fun toplevel_source term do_trace do_recover cmd src =
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   201
  let
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   202
    val no_terminator =
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   203
      Scan.unless P.semicolon (Scan.one (T.not_sync andf T.not_eof));
23682
cf4773532006 nested source: explicit interactive flag for recover avoids duplicate errors;
wenzelm
parents: 23679
diff changeset
   204
    fun recover int =
cf4773532006 nested source: explicit interactive flag for recover avoids duplicate errors;
wenzelm
parents: 23679
diff changeset
   205
      (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
   206
  in
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   207
    src
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 10749
diff changeset
   208
    |> T.source_proper
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   209
    |> Source.source T.stopper
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15224
diff changeset
   210
      (Scan.bulk (P.$$$ "--" -- P.!!! P.text >> K NONE || P.not_eof >> SOME))
23682
cf4773532006 nested source: explicit interactive flag for recover avoids duplicate errors;
wenzelm
parents: 23679
diff changeset
   211
        (Option.map recover do_recover)
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19060
diff changeset
   212
    |> Source.map_filter I
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   213
    |> Source.source T.stopper
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   214
        (Scan.bulk (fn xs => P.!!! (parse_command term do_trace (cmd ())) xs))
23682
cf4773532006 nested source: explicit interactive flag for recover avoids duplicate errors;
wenzelm
parents: 23679
diff changeset
   215
        (Option.map recover do_recover)
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19060
diff changeset
   216
    |> Source.map_filter I
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   217
  end;
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   218
7746
b52c1a632e6a simplified;
wenzelm
parents: 7735
diff changeset
   219
25580
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   220
(* off-line scanning/parsing *)
14925
0f86a8a694f8 added read (provides transition names and sources);
wenzelm
parents: 14687
diff changeset
   221
25580
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   222
(*tokens*)
15144
85929e1b307d Remove isar_readstring. Split read into scanner and parser.
aspinall
parents: 14981
diff changeset
   223
fun scan str =
16195
wenzelm
parents: 15989
diff changeset
   224
  Source.of_string str
wenzelm
parents: 15989
diff changeset
   225
  |> Symbol.source false
23679
57dceb84d1a0 toplevel_source: interactive flag indicates intermittent error_msg;
wenzelm
parents: 22826
diff changeset
   226
  |> T.source (SOME false) get_lexicons Position.none
16195
wenzelm
parents: 15989
diff changeset
   227
  |> Source.exhaust;
wenzelm
parents: 15989
diff changeset
   228
25580
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   229
(*commands from tokens, with trace*)
15144
85929e1b307d Remove isar_readstring. Split read into scanner and parser.
aspinall
parents: 14981
diff changeset
   230
fun read toks =
85929e1b307d Remove isar_readstring. Split read into scanner and parser.
aspinall
parents: 14981
diff changeset
   231
  Source.of_list toks
25580
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   232
  |> toplevel_source false true NONE get_parser
14925
0f86a8a694f8 added read (provides transition names and sources);
wenzelm
parents: 14687
diff changeset
   233
  |> Source.exhaust
15973
wenzelm
parents: 15830
diff changeset
   234
  |> map (fn tr => (Toplevel.name_of tr, the (Toplevel.source_of tr), tr));
14091
ad6ba9c55190 A patch by david aspinall
nipkow
parents: 12943
diff changeset
   235
25580
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   236
(*commands from string, without trace*)
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   237
fun parse pos str =
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   238
  Source.of_string str
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   239
  |> Symbol.source false
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   240
  |> T.source (SOME false) get_lexicons pos
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   241
  |> toplevel_source false false NONE get_parser
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   242
  |> Source.exhaust;
6623674df897 added off-line parse;
wenzelm
parents: 25526
diff changeset
   243
14091
ad6ba9c55190 A patch by david aspinall
nipkow
parents: 12943
diff changeset
   244
26431
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   245
(* process file *)
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   246
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   247
fun process_file path thy =
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   248
  let
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   249
    val result = ref thy;
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   250
    val trs = parse (Position.path path) (File.read path);
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   251
    val init = Toplevel.init_theory (K thy) (fn thy' => result := thy') (K ());
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   252
    val _ = Toplevel.excursion (init Toplevel.empty :: trs @ [Toplevel.exit Toplevel.empty]);
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   253
  in ! result end;
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   254
f1c79c00f1e4 added process_file;
wenzelm
parents: 26415
diff changeset
   255
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   256
(* interactive source of toplevel transformers *)
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   257
26600
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
   258
type isar =
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
   259
  (Toplevel.transition, (Toplevel.transition option,
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
   260
    (OuterLex.token, (OuterLex.token option, (OuterLex.token, (OuterLex.token,
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
   261
      Position.T * (Symbol.symbol, (string, unit) Source.source) Source.source)
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
   262
          Source.source) Source.source) Source.source) Source.source) Source.source) Source.source;
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
   263
f11515535c83 moved structure Isar to isar.ML;
wenzelm
parents: 26431
diff changeset
   264
fun isar term : isar =
24868
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   265
  Source.tty
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   266
  |> Symbol.source true
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   267
  |> T.source (SOME true) get_lexicons Position.none
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   268
  |> toplevel_source term false (SOME true) get_parser;
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   269
2990c327d8c6 simplified interfaces for outer syntax;
wenzelm
parents: 24188
diff changeset
   270
16195
wenzelm
parents: 15989
diff changeset
   271
9132
52286129faa5 moved header stuff to thy_header.ML;
wenzelm
parents: 9056
diff changeset
   272
(** read theory **)
6247
e8bbe64861b8 more robust handling of theory context;
wenzelm
parents: 6229
diff changeset
   273
12943
1db24da0537b added check_text;
wenzelm
parents: 12876
diff changeset
   274
(* check_text *)
1db24da0537b added check_text;
wenzelm
parents: 12876
diff changeset
   275
22120
8424ef945cb5 renamed IsarOutput to ThyOutput;
wenzelm
parents: 22086
diff changeset
   276
fun check_text s state = (ThyOutput.eval_antiquote (#1 (get_lexicons ())) state s; ());
12943
1db24da0537b added check_text;
wenzelm
parents: 12876
diff changeset
   277
1db24da0537b added check_text;
wenzelm
parents: 12876
diff changeset
   278
26611
03455add4801 export load_thy -- no backpatching;
wenzelm
parents: 26600
diff changeset
   279
(* load_thy *)
7746
b52c1a632e6a simplified;
wenzelm
parents: 7735
diff changeset
   280
26323
73efc70edeef theory loader: discontinued *attached* ML scripts;
wenzelm
parents: 26291
diff changeset
   281
fun load_thy dir name pos text time =
7683
e74f43f1d8a3 export token_source;
wenzelm
parents: 7676
diff changeset
   282
  let
24065
21483400c2ca load_thy: avoid reloading of text;
wenzelm
parents: 23939
diff changeset
   283
    val text_src = Source.of_list (Library.untabify text);
23866
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   284
17932
677f7bec354e removed old-style theory format;
wenzelm
parents: 17412
diff changeset
   285
    val _ = Present.init_theory name;
24065
21483400c2ca load_thy: avoid reloading of text;
wenzelm
parents: 23939
diff changeset
   286
    val _ = Present.verbatim_source name (fn () => Source.exhaust (Symbol.source false text_src));
21483400c2ca load_thy: avoid reloading of text;
wenzelm
parents: 23939
diff changeset
   287
    val toks = text_src
17932
677f7bec354e removed old-style theory format;
wenzelm
parents: 17412
diff changeset
   288
      |> Symbol.source false
24065
21483400c2ca load_thy: avoid reloading of text;
wenzelm
parents: 23939
diff changeset
   289
      |> T.source NONE (K (get_lexicons ())) pos
17932
677f7bec354e removed old-style theory format;
wenzelm
parents: 17412
diff changeset
   290
      |> Source.exhausted;
677f7bec354e removed old-style theory format;
wenzelm
parents: 17412
diff changeset
   291
    val trs = toks
23679
57dceb84d1a0 toplevel_source: interactive flag indicates intermittent error_msg;
wenzelm
parents: 22826
diff changeset
   292
      |> toplevel_source false false NONE (K (get_parser ()))
17932
677f7bec354e removed old-style theory format;
wenzelm
parents: 17412
diff changeset
   293
      |> Source.exhaust;
23866
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   294
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   295
    val _ = if time then writeln ("\n**** Starting theory " ^ quote name ^ " ****") else ();
25685
c36e10812ae4 cond_timeit: added message argument;
wenzelm
parents: 25583
diff changeset
   296
    val _ = cond_timeit time "" (fn () =>
23866
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   297
      ThyOutput.process_thy (#1 (get_lexicons ())) command_tags is_markup trs toks
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   298
      |> Buffer.content
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   299
      |> Present.theory_output name);
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   300
    val _ = if time then writeln ("**** Finished theory " ^ quote name ^ " ****\n") else ();
24065
21483400c2ca load_thy: avoid reloading of text;
wenzelm
parents: 23939
diff changeset
   301
  in () end;
23866
5295671034f8 moved deps_thy to ThyLoad (independent of outer syntax);
wenzelm
parents: 23796
diff changeset
   302
5829
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   303
end;
0acb30dd92bc The global Isabelle/Isar outer syntax.
wenzelm
parents:
diff changeset
   304