src/Pure/Thy/thy_syntax.ML
author wenzelm
Thu, 06 Nov 2014 15:47:04 +0100
changeset 58923 cb9b69cca999
parent 58864 505a8150368a
child 58978 e42da880c61e
permissions -rw-r--r--
more explicit Keyword.keywords;

(*  Title:      Pure/Thy/thy_syntax.ML
    Author:     Makarius

Superficial theory syntax: tokens and spans.
*)

signature THY_SYNTAX =
sig
  val reports_of_tokens: Token.T list -> bool * Position.report_text list
  val present_token: Token.T -> Output.output
  val present_span: Command_Span.span -> Output.output
  datatype 'a element = Element of 'a * ('a element list * 'a) option
  val atom: 'a -> 'a element
  val map_element: ('a -> 'b) -> 'a element -> 'b element
  val flat_element: 'a element -> 'a list
  val last_element: 'a element -> 'a
  val parse_elements: Keyword.keywords -> Command_Span.span list -> Command_Span.span element list
end;

structure Thy_Syntax: THY_SYNTAX =
struct

(** presentation **)

local

fun reports_of_token tok =
  let
    val {text, pos, ...} = Token.source_position_of tok;
    val malformed_symbols =
      Symbol_Pos.explode (text, pos)
      |> map_filter (fn (sym, pos) =>
          if Symbol.is_malformed sym
          then SOME ((pos, Markup.bad), "Malformed symbolic character") else NONE);
    val is_malformed = Token.is_error tok orelse not (null malformed_symbols);
    val reports = Token.report tok :: Token.completion_report tok @ malformed_symbols;
  in (is_malformed, reports) end;

in

fun reports_of_tokens toks =
  let val results = map reports_of_token toks
  in (exists fst results, maps snd results) end;

end;

fun present_token tok =
  Markup.enclose (Token.markup tok) (Output.output (Token.unparse tok));

val present_span = implode o map present_token o Command_Span.content;



(** specification elements: commands with optional proof **)

datatype 'a element = Element of 'a * ('a element list * 'a) option;

fun element (a, b) = Element (a, SOME b);
fun atom a = Element (a, NONE);

fun map_element f (Element (a, NONE)) = Element (f a, NONE)
  | map_element f (Element (a, SOME (elems, b))) =
      Element (f a, SOME ((map o map_element) f elems, f b));

fun flat_element (Element (a, NONE)) = [a]
  | flat_element (Element (a, SOME (elems, b))) = a :: maps flat_element elems @ [b];

fun last_element (Element (a, NONE)) = a
  | last_element (Element (_, SOME (_, b))) = b;


(* scanning spans *)

val eof = Command_Span.Span (Command_Span.Command_Span ("", Position.none), []);

fun is_eof (Command_Span.Span (Command_Span.Command_Span ("", _), _)) = true
  | is_eof _ = false;

val not_eof = not o is_eof;

val stopper = Scan.stopper (K eof) is_eof;


(* parse *)

local

fun command_with pred =
  Scan.one
    (fn (Command_Span.Span (Command_Span.Command_Span (name, _), _)) => pred name | _ => false);

fun parse_element keywords =
  let
    val proof_atom =
      Scan.one
        (fn (Command_Span.Span (Command_Span.Command_Span (name, _), _)) =>
              Keyword.is_proof_body keywords name
          | _ => true) >> atom;
    fun proof_element x =
      (command_with (Keyword.is_proof_goal keywords) -- proof_rest >> element || proof_atom) x
    and proof_rest x =
      (Scan.repeat proof_element -- command_with (Keyword.is_qed keywords)) x;
  in
    command_with (Keyword.is_theory_goal keywords) -- proof_rest >> element ||
    Scan.one not_eof >> atom
  end;

in

fun parse_elements keywords =
  Source.of_list #>
  Source.source stopper (Scan.bulk (parse_element keywords)) #>
  Source.exhaust;

end;

end;