src/Pure/Isar/outer_syntax.ML
changeset 24868 2990c327d8c6
parent 24188 d5960310c4d5
child 24872 7fd1aa6671a4
--- a/src/Pure/Isar/outer_syntax.ML	Sat Oct 06 16:50:04 2007 +0200
+++ b/src/Pure/Isar/outer_syntax.ML	Sat Oct 06 16:50:08 2007 +0200
@@ -2,7 +2,9 @@
     ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
-The global Isabelle/Isar outer syntax.
+The global Isabelle/Isar outer syntax. Note: the syntax for files is
+statically determined at the very beginning; for interactive processing
+it may change dynamically.
 *)
 
 signature BASIC_OUTER_SYNTAX =
@@ -24,27 +26,23 @@
 signature OUTER_SYNTAX =
 sig
   include BASIC_OUTER_SYNTAX
-  type token
-  type parser
+  type parser_fn = OuterLex.token list ->
+    (Toplevel.transition -> Toplevel.transition) * OuterLex.token list
   val get_lexicons: unit -> Scan.lexicon * Scan.lexicon
   val command_keyword: string -> OuterKeyword.T option
-  val command: string -> string -> OuterKeyword.T ->
-    (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
-  val markup_command: ThyOutput.markup -> string -> string -> OuterKeyword.T ->
-    (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
-  val improper_command: string -> string -> OuterKeyword.T ->
-    (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
   val is_keyword: string -> bool
+  val keywords: string list -> unit
+  val command: string -> string -> OuterKeyword.T -> parser_fn -> unit
+  val markup_command: ThyOutput.markup -> string -> string -> OuterKeyword.T -> parser_fn -> unit
+  val improper_command: string -> string -> OuterKeyword.T -> parser_fn -> unit
   val dest_keywords: unit -> string list
   val dest_parsers: unit -> (string * string * string * bool) list
   val print_outer_syntax: unit -> unit
   val print_commands: Toplevel.transition -> Toplevel.transition
-  val add_keywords: string list -> unit
-  val add_parsers: parser list -> unit
   val check_text: string * Position.T -> Toplevel.node option -> unit
-  val isar: bool -> unit Toplevel.isar
   val scan: string -> OuterLex.token list
   val read: OuterLex.token list -> (string * OuterLex.token list * Toplevel.transition) list
+  val isar: bool -> unit Toplevel.isar
 end;
 
 structure OuterSyntax : OUTER_SYNTAX  =
@@ -58,14 +56,17 @@
 
 (* parsers *)
 
-type token = T.token;
-type parser_fn = token list -> (Toplevel.transition -> Toplevel.transition) * token list;
+type parser_fn = T.token list -> (Toplevel.transition -> Toplevel.transition) * T.token list;
 
-datatype parser =
-  Parser of string * (string * OuterKeyword.T * ThyOutput.markup option) * bool * parser_fn;
+datatype parser = Parser of
+ {comment: string,
+  kind: OuterKeyword.T,
+  markup: ThyOutput.markup option,
+  int_only: bool,
+  parse: parser_fn};
 
-fun parser int_only markup name comment kind parse =
-  Parser (name, (comment, kind, markup), int_only, parse);
+fun make_parser comment kind markup int_only parse =
+  Parser {comment = comment, kind = kind, markup = markup, int_only = int_only, parse = parse};
 
 
 (* parse command *)
@@ -80,13 +81,13 @@
 
 fun body cmd do_trace (name, _) =
   (case cmd name of
-    SOME (int_only, parse) =>
+    SOME (Parser {int_only, parse, ...}) =>
       P.!!! (Scan.prompt (name ^ "# ") (trace do_trace (P.tags |-- parse) >> pair int_only))
   | NONE => sys_error ("no parser for outer syntax command " ^ quote name));
 
 in
 
-fun command do_terminate do_trace cmd =
+fun parse_command do_terminate do_trace cmd =
   P.semicolon >> K NONE ||
   P.sync >> K NONE ||
   (P.position P.command :-- body cmd do_trace) --| terminate do_terminate
@@ -103,9 +104,7 @@
 local
 
 val global_lexicons = ref (Scan.empty_lexicon, Scan.empty_lexicon);
-val global_parsers =
-  ref (Symtab.empty: (((string * OuterKeyword.T) * (bool * parser_fn)) * ThyOutput.markup option)
-    Symtab.table);
+val global_parsers = ref (Symtab.empty: parser Symtab.table);
 val global_markups = ref ([]: (string * ThyOutput.markup) list);
 
 fun change_lexicons f = CRITICAL (fn () =>
@@ -118,41 +117,48 @@
 fun change_parsers f = CRITICAL (fn () =>
  (change global_parsers f;
   global_markups :=
-    Symtab.fold (fn (name, (_, SOME m)) => cons (name, m) | _ => I) (! global_parsers) []));
+    Symtab.fold (fn (name, Parser {markup = SOME m, ...}) => cons (name, m) | _ => I)
+      (! global_parsers) []));
 
 in
 
-
 (* access current syntax *)
 
-(*Note: the syntax for files is statically determined at the very
-  beginning; for interactive processing it may change dynamically.*)
+fun get_lexicons () = CRITICAL (fn () => ! global_lexicons);
+fun get_parsers () = CRITICAL (fn () => ! global_parsers);
+fun get_markups () = CRITICAL (fn () => ! global_markups);
 
-fun get_lexicons () = ! global_lexicons;
-fun get_parsers () = ! global_parsers;
-fun get_parser () = Option.map (#2 o #1) o Symtab.lookup (get_parsers ());
+fun get_parser () = Symtab.lookup (get_parsers ());
 
 fun command_keyword name =
-  Option.map (fn (((_, k), _), _) => k) (Symtab.lookup (get_parsers ()) name);
+  (case Symtab.lookup (get_parsers ()) name of
+    SOME (Parser {kind, ...}) => SOME kind
+  | NONE => NONE);
+
 fun command_tags name = these ((Option.map OuterKeyword.tags_of) (command_keyword name));
 
-fun is_markup kind name = (AList.lookup (op =) (! global_markups) name = SOME kind);
+fun is_markup kind name = AList.lookup (op =) (get_markups ()) name = SOME kind;
 
 
 (* augment syntax *)
 
-fun add_keywords keywords =
-  change_lexicons (apfst (Scan.extend_lexicon (map Symbol.explode keywords)));
+val keywords = change_lexicons o apfst o Scan.extend_lexicon o map Symbol.explode;
+
+
+fun add_parser (name, parser) =
+ (if not (Symtab.defined (get_parsers ()) name) then ()
+  else warning ("Redefining outer syntax command " ^ quote name);
+  change_parsers (Symtab.update (name, parser));
+  change_lexicons (apsnd (Scan.extend_lexicon [Symbol.explode name])));
 
-fun add_parser (Parser (name, (comment, kind, markup), int_only, parse)) tab =
- (if not (Symtab.defined tab name) then ()
-  else warning ("Redefined outer syntax command " ^ quote name);
-  Symtab.update (name, (((comment, kind), (int_only, parse)), markup)) tab);
+fun command name comment kind parse =
+  add_parser (name, make_parser comment kind NONE false parse);
 
-fun add_parsers parsers = CRITICAL (fn () =>
-  (change_parsers (fold add_parser parsers);
-    change_lexicons (apsnd (Scan.extend_lexicon
-      (map (fn Parser (name, _, _, _) => Symbol.explode name) parsers)))));
+fun markup_command markup name comment kind parse =
+  add_parser (name, make_parser comment kind (SOME markup) false parse);
+
+fun improper_command name comment kind parse =
+  add_parser (name, make_parser comment kind NONE true parse);
 
 end;
 
@@ -164,8 +170,8 @@
 
 fun dest_parsers () =
   get_parsers () |> Symtab.dest |> sort_wrt #1
-  |> map (fn (name, (((cmt, kind), (int_only, _)), _)) =>
-    (name, cmt, OuterKeyword.kind_of kind, int_only));
+  |> map (fn (name, Parser {comment, kind, int_only, ...}) =>
+    (name, comment, OuterKeyword.kind_of kind, int_only));
 
 fun print_outer_syntax () =
   let
@@ -200,21 +206,13 @@
       (Scan.bulk (P.$$$ "--" -- P.!!! P.text >> K NONE || P.not_eof >> SOME))
         (Option.map recover do_recover)
     |> Source.map_filter I
-    |> Source.source T.stopper (Scan.bulk (fn xs => P.!!! (command term do_trace (cmd ())) xs))
+    |> Source.source T.stopper
+        (Scan.bulk (fn xs => P.!!! (parse_command term do_trace (cmd ())) xs))
         (Option.map recover do_recover)
     |> Source.map_filter I
   end;
 
 
-(* interactive source of toplevel transformers *)
-
-fun isar term =
-  Source.tty
-  |> Symbol.source true
-  |> T.source (SOME true) get_lexicons Position.none
-  |> toplevel_source term false (SOME true) get_parser;
-
-
 (* scan text *)
 
 fun scan str =
@@ -233,6 +231,15 @@
   |> map (fn tr => (Toplevel.name_of tr, the (Toplevel.source_of tr), tr));
 
 
+(* interactive source of toplevel transformers *)
+
+fun isar term =
+  Source.tty
+  |> Symbol.source true
+  |> T.source (SOME true) get_lexicons Position.none
+  |> toplevel_source term false (SOME true) get_parser;
+
+
 
 (** read theory **)
 
@@ -324,12 +331,6 @@
   val toplevel = Toplevel.program;
 end;
 
-
-(*final declarations of this structure!*)
-val command = parser false NONE;
-val markup_command = parser false o SOME;
-val improper_command = parser true NONE;
-
 end;
 
 structure ThyLoad: THY_LOAD = ThyLoad;