abstract type span, tuned interfaces;
authorwenzelm
Tue, 12 Aug 2008 21:28:03 +0200
changeset 27842 6c35d50d309f
parent 27841 55b028593335
child 27843 0bd68bf0cbb8
abstract type span, tuned interfaces; added report_token, report_span; markup ident tokens;
src/Pure/Thy/thy_edit.ML
--- a/src/Pure/Thy/thy_edit.ML	Tue Aug 12 21:28:01 2008 +0200
+++ b/src/Pure/Thy/thy_edit.ML	Tue Aug 12 21:28:03 2008 +0200
@@ -2,24 +2,27 @@
     ID:         $Id$
     Author:     Makarius
 
-Basic editing of theory sources.
+Basic editor operations on theory sources.
 *)
 
 signature THY_EDIT =
 sig
-  val token_source: Position.T -> (string, 'a) Source.source ->
+  val token_source: Scan.lexicon * Scan.lexicon -> Position.T -> (string, 'a) Source.source ->
     (OuterLex.token, (SymbolPos.T, Position.T * (Symbol.symbol, (string, 'a)
       Source.source) Source.source) Source.source) Source.source
-  val parse_tokens: Position.T -> (string, 'a) Source.source -> OuterLex.token list
+  val parse_tokens: Scan.lexicon * Scan.lexicon -> Position.T -> string -> OuterLex.token list
   val present_token: OuterLex.token -> output
-  datatype span_kind = Command of string | Ignored | Unknown
-  type span = span_kind * OuterLex.token list
+  val report_token: OuterLex.token -> unit
+  datatype span_kind = Command of string | Ignored | Malformed
+  type span
+  val span_kind: span -> span_kind
+  val span_content: span -> OuterLex.token list
   val span_range: span -> Position.range
-  val span_source: Position.T -> (string, 'a) Source.source ->
-    (span, (OuterLex.token, (SymbolPos.T, Position.T * (Symbol.symbol, (string, 'a)
-      Source.source) Source.source) Source.source) Source.source) Source.source
-  val parse_spans: Position.T -> (string, 'a) Source.source -> span list
+  val span_source: (OuterLex.token, 'a) Source.source ->
+    (span, (OuterLex.token, 'a) Source.source) Source.source
+  val parse_spans: Scan.lexicon * Scan.lexicon -> Position.T -> string -> span list
   val present_span: span -> output
+  val report_span: span -> unit
   val present_html: Path.T -> Path.T -> unit
 end;
 
@@ -34,11 +37,14 @@
 
 (* parse *)
 
-fun token_source pos src =
-  Symbol.source true src
-  |> T.source (SOME false) OuterKeyword.get_lexicons pos;
+fun token_source lexs pos src =
+  Symbol.source {do_recover = true} src
+  |> T.source {do_recover = SOME false} (K lexs) pos;
 
-fun parse_tokens pos src = Source.exhaust (token_source pos src);
+fun parse_tokens lexs pos str =
+  Source.of_string str
+  |> token_source lexs pos
+  |> Source.exhaust;
 
 
 (* present *)
@@ -48,6 +54,13 @@
 val token_kind_markup =
  fn T.Command   => (Markup.commandN, [])
   | T.Keyword   => (Markup.keywordN, [])
+  | T.Ident     => Markup.ident
+  | T.LongIdent => Markup.ident
+  | T.SymIdent  => Markup.ident
+  | T.Var       => Markup.ident
+  | T.TypeIdent => Markup.ident
+  | T.TypeVar   => Markup.ident
+  | T.Nat       => Markup.ident
   | T.String    => Markup.string
   | T.AltString => Markup.altstring
   | T.Verbatim  => Markup.verbatim
@@ -63,21 +76,31 @@
 fun present_token tok =
   Markup.enclose (token_kind_markup (T.kind_of tok)) (Output.output (T.unparse tok));
 
+fun report_token tok =
+  Position.report (token_kind_markup (T.kind_of tok)) (T.position_of tok);
+
 end;
 
 
 
 (** spans **)
 
-datatype span_kind = Command of string | Ignored | Unknown;
-type span = span_kind * OuterLex.token list;
+(* type span *)
+
+datatype span_kind = Command of string | Ignored | Malformed;
+datatype span = Span of span_kind * OuterLex.token list;
 
-fun span_range ((_, []): span) = raise Fail "span_range: empty span"
-  | span_range (_, toks) =
+fun span_kind (Span (k, _)) = k;
+fun span_content (Span (_, toks)) = toks;
+
+fun span_range span =
+  (case span_content span of
+    [] => (Position.none, Position.none)
+  | toks =>
       let
         val start_pos = T.position_of (hd toks);
         val end_pos = T.end_position_of (List.last toks);
-      in (start_pos, end_pos) end;
+      in (start_pos, end_pos) end);
 
 
 (* parse *)
@@ -90,19 +113,21 @@
 
 val span =
   Scan.ahead P.command -- P.not_eof -- Scan.repeat body
-    >> (fn ((name, c), bs) => (Command name, c :: bs)) ||
-  Scan.many1 is_whitespace >> pair Ignored ||
-  Scan.repeat1 body >> pair Unknown;
+    >> (fn ((name, c), bs) => Span (Command name, c :: bs)) ||
+  Scan.many1 is_whitespace >> (fn toks => Span (Ignored, toks)) ||
+  Scan.repeat1 body >> (fn toks => Span (Malformed, toks));
 
 in
 
-fun span_source pos src =
-  token_source pos src
-  |> Source.source T.stopper (Scan.bulk span) NONE;
+fun span_source src = Source.source T.stopper (Scan.bulk span) NONE src;
 
 end;
 
-fun parse_spans pos src = Source.exhaust (span_source pos src);
+fun parse_spans lexs pos str =
+  Source.of_string str
+  |> token_source lexs pos
+  |> span_source
+  |> Source.exhaust;
 
 
 (* present *)
@@ -111,12 +136,15 @@
 
 fun kind_markup (Command name) = Markup.command_span name
   | kind_markup Ignored = Markup.ignored_span
-  | kind_markup Unknown = Markup.unknown_span;
+  | kind_markup Malformed = Markup.malformed_span;
 
 in
 
-fun present_span (kind, toks) =
-  Markup.enclose (kind_markup kind) (implode (map present_token toks));
+fun present_span span =
+  Markup.enclose (kind_markup (span_kind span)) (implode (map present_token (span_content span)));
+
+fun report_span span =
+  Position.report (kind_markup (span_kind span)) (Position.encode_range (span_range span));
 
 end;
 
@@ -124,7 +152,7 @@
 (* HTML presentation -- example *)
 
 fun present_html inpath outpath =
-  parse_spans (Path.position inpath) (Source.of_string (File.read inpath))
+  parse_spans (OuterKeyword.get_lexicons ()) (Path.position inpath) (File.read inpath)
   |> HTML.html_mode (implode o map present_span)
   |> enclose
     (HTML.begin_document (Path.implode (Path.base inpath)) ^ "<div class=\"source\"><pre>")