src/Pure/Isar/outer_parse.ML
author wenzelm
Tue, 13 Sep 2005 22:19:43 +0200
changeset 17358 5746c9bd4356
parent 17165 9b498019723f
child 17756 d4a35f82fbb4
permissions -rw-r--r--
added name_facts;

(*  Title:      Pure/Isar/outer_parse.ML
    ID:         $Id$
    Author:     Markus Wenzel, TU Muenchen

Generic parsers for Isabelle/Isar outer syntax.
*)

signature OUTER_PARSE =
sig
  type token
  val group: string -> (token list -> 'a) -> token list -> 'a
  val !!! : (token list -> 'a) -> token list -> 'a
  val !!!! : (token list -> 'a) -> token list -> 'a
  val triple1: ('a * 'b) * 'c -> 'a * 'b * 'c
  val triple2: 'a * ('b * 'c) -> 'a * 'b * 'c
  val triple_swap: ('a * 'b) * 'c -> ('a * 'c) * 'b
  val not_eof: token list -> token * token list
  val position: (token list -> 'a * 'b) -> token list -> ('a * Position.T) * 'b
  val command: token list -> string * token list
  val keyword: token list -> string * token list
  val short_ident: token list -> string * token list
  val long_ident: token list -> string * token list
  val sym_ident: token list -> string * token list
  val minus: token list -> string * token list
  val term_var: token list -> string * token list
  val type_ident: token list -> string * token list
  val type_var: token list -> string * token list
  val number: token list -> string * token list
  val string: token list -> string * token list
  val verbatim: token list -> string * token list
  val sync: token list -> string * token list
  val eof: token list -> string * token list
  val $$$ : string -> token list -> string * token list
  val reserved : string -> token list -> string * token list
  val semicolon: token list -> string * token list
  val underscore: token list -> string * token list
  val maybe: (token list -> 'a * token list) -> token list -> 'a option * token list
  val tag_name: token list -> string * token list
  val tags: token list -> string list * token list
  val opt_unit: token list -> unit * token list
  val opt_keyword: string -> token list -> bool * token list
  val nat: token list -> int * token list
  val enum: string -> (token list -> 'a * token list) -> token list -> 'a list * token list
  val enum1: string -> (token list -> 'a * token list) -> token list -> 'a list * token list
  val list: (token list -> 'a * token list) -> token list -> 'a list * token list
  val list1: (token list -> 'a * token list) -> token list -> 'a list * token list
  val and_list: (token list -> 'a * token list) -> token list -> 'a list * token list
  val and_list1: (token list -> 'a * token list) -> token list -> 'a list * token list
  val name: token list -> bstring * token list
  val xname: token list -> xstring * token list
  val text: token list -> string * token list
  val path: token list -> Path.T * token list
  val sort: token list -> string * token list
  val arity: token list -> (string list * string) * token list
  val type_args: token list -> string list * token list
  val typ: token list -> string * token list
  val opt_infix: token list -> Syntax.mixfix * token list
  val opt_mixfix: token list -> Syntax.mixfix * token list
  val opt_infix': token list -> Syntax.mixfix * token list
  val opt_mixfix': token list -> Syntax.mixfix * token list
  val mixfix': token list -> Syntax.mixfix * token list
  val const: token list -> (bstring * string * Syntax.mixfix) * token list
  val term: token list -> string * token list
  val prop: token list -> string * token list
  val propp: token list -> (string * (string list * string list)) * token list
  val termp: token list -> (string * string list) * token list
  val arguments: token list -> Args.T list * token list
  val attribs: token list -> Attrib.src list * token list
  val opt_attribs: token list -> Attrib.src list * token list
  val thm_name: string -> token list -> (bstring * Attrib.src list) * token list
  val opt_thm_name: string -> token list -> (bstring * Attrib.src list) * token list
  val spec_name: token list -> ((bstring * string) * Attrib.src list) * token list
  val spec_opt_name: token list -> ((bstring * string) * Attrib.src list) * token list
  val xthm: token list -> (thmref * Attrib.src list) * token list
  val xthms1: token list -> (thmref * Attrib.src list) list * token list
  val name_facts: token list ->
    ((bstring * Attrib.src list) * (thmref * Attrib.src list) list) list * token list
  val locale_target: token list -> xstring * token list
  val opt_locale_target: token list -> xstring option * token list
  val locale_expr: token list -> Locale.expr * token list
  val locale_keyword: token list -> string * token list
  val locale_element: token list -> Locale.element * token list
  val locale_elem_or_expr: token list -> Locale.element Locale.elem_expr * token list
  val method: token list -> Method.text * token list
end;

structure OuterParse: OUTER_PARSE =
struct

structure T = OuterLex;
type token = T.token;


(** error handling **)

(* group atomic parsers (no cuts!) *)

fun fail_with s = Scan.fail_with
  (fn [] => s ^ " expected (past end-of-file!)"
    | (tok :: _) => s ^ " expected,\nbut " ^ T.name_of tok ^ T.pos_of tok ^ " was found");

fun group s scan = scan || fail_with s;


(* cut *)

fun cut kind scan =
  let
    fun get_pos [] = " (past end-of-file!)"
      | get_pos (tok :: _) = T.pos_of tok;

    fun err (toks, NONE) = kind ^ get_pos toks
      | err (toks, SOME msg) = kind ^ get_pos toks ^ ": " ^ msg;
  in Scan.!! err scan end;

fun !!! scan = cut "Outer syntax error" scan;
fun !!!! scan = cut "Corrupted outer syntax in presentation" scan;



(** basic parsers **)

(* utils *)

fun triple1 ((x, y), z) = (x, y, z);
fun triple2 (x, (y, z)) = (x, y, z);
fun triple_swap ((x, y), z) = ((x, z), y);


(* tokens *)

val not_eof = Scan.one T.not_eof;

fun position scan = (Scan.ahead not_eof >> T.position_of) -- scan >> Library.swap;

fun kind k =
  group (T.str_of_kind k) (Scan.one (T.is_kind k) >> T.val_of);

val command = kind T.Command;
val keyword = kind T.Keyword;
val short_ident = kind T.Ident;
val long_ident = kind T.LongIdent;
val sym_ident = kind T.SymIdent;
val term_var = kind T.Var;
val type_ident = kind T.TypeIdent;
val type_var = kind T.TypeVar;
val number = kind T.Nat;
val string = kind T.String;
val alt_string = kind T.AltString;
val verbatim = kind T.Verbatim;
val sync = kind T.Sync;
val eof = kind T.EOF;

fun $$$ x =
  group (T.str_of_kind T.Keyword ^ " " ^ quote x)
    (Scan.one (T.keyword_with (equal x)) >> T.val_of);

fun reserved x =
  group ("Reserved identifier " ^ quote x)
    (Scan.one (T.ident_with (equal x)) >> T.val_of);

val semicolon = $$$ ";";

val minus = sym_ident :-- (fn "-" => Scan.succeed () | _ => Scan.fail) >> #1;
val underscore = sym_ident :-- (fn "_" => Scan.succeed () | _ => Scan.fail) >> #1;
fun maybe scan = underscore >> K NONE || scan >> SOME;

val nat = number >> (#1 o Library.read_int o Symbol.explode);

val tag_name = group "tag name" (short_ident || string);
val tags = Scan.repeat ($$$ "%" |-- !!! tag_name);

val opt_unit = Scan.optional ($$$ "(" -- $$$ ")" >> (K ())) ();
fun opt_keyword s = Scan.optional ($$$ "(" |-- !!! (($$$ s >> K true) --| $$$ ")")) false;


(* enumerations *)

fun enum1 sep scan = scan -- Scan.repeat ($$$ sep |-- !!! scan) >> op ::;
fun enum sep scan = enum1 sep scan || Scan.succeed [];

fun list1 scan = enum1 "," scan;
fun list scan = enum "," scan;

fun and_list1 scan = enum1 "and" scan;
fun and_list scan = enum "and" scan;


(* names and text *)

val name = group "name declaration" (short_ident || sym_ident || string || number);
val xname = group "name reference" (short_ident || long_ident || sym_ident || string || number);
val text = group "text" (short_ident || long_ident || sym_ident || string || number || verbatim);
val path = group "file name/path specification" name >> Path.unpack;


(* sorts and arities *)

val sort = group "sort" xname;

fun gen_arity cod =
  Scan.optional ($$$ "(" |-- !!! (list1 sort --| $$$ ")")) [] -- cod;

val arity = gen_arity sort;


(* types *)

val typ = group "type"
  (short_ident || long_ident || sym_ident || type_ident || type_var || string || number);

val type_args =
  type_ident >> single ||
  $$$ "(" |-- !!! (list1 type_ident --| $$$ ")") ||
  Scan.succeed [];


(* mixfix annotations *)

val infx = $$$ "infix" |-- !!! (nat >> Syntax.Infix || string -- nat >> Syntax.InfixName);
val infxl = $$$ "infixl" |-- !!! (nat >> Syntax.Infixl || string -- nat >> Syntax.InfixlName);
val infxr = $$$ "infixr" |-- !!! (nat >> Syntax.Infixr || string -- nat >> Syntax.InfixrName);

val binder =
  $$$ "binder" |--
    !!! (string -- ($$$ "[" |-- nat --| $$$ "]" -- nat || nat >> (fn n => (n, n))))
  >> (Syntax.Binder o triple2);


val opt_pris = Scan.optional ($$$ "[" |-- !!! (list nat --| $$$ "]")) [];

val mixfix =
  (string -- !!! (opt_pris -- Scan.optional nat Syntax.max_pri))
  >> (Syntax.Mixfix o triple2);

fun fix_decl guard fix = $$$ "(" |-- guard (fix --| $$$ ")");
fun opt_fix guard fix = Scan.optional (fix_decl guard fix) Syntax.NoSyn;

val opt_infix = opt_fix !!! (infxl || infxr || infx);
val opt_mixfix = opt_fix !!! (mixfix || binder || infxl || infxr || infx);
val opt_infix' = opt_fix I (infxl || infxr || infx);
val opt_mixfix' = opt_fix I (mixfix || binder || infxl || infxr || infx);
val mixfix' = fix_decl !!! (mixfix || binder || infxl || infxr || infx);


(* consts *)

val const =
  name -- ($$$ "::" |-- !!! (typ -- opt_mixfix)) >> triple2;


(* terms *)

val trm = short_ident || long_ident || sym_ident || term_var || number || string;

val term = group "term" trm;
val prop = group "proposition" trm;


(* patterns *)

val is_terms = Scan.repeat1 ($$$ "is" |-- term);
val is_props = Scan.repeat1 ($$$ "is" |-- prop);
val concl_props = $$$ "concl" |-- !!! is_props;
val any_props = concl_props >> pair [] || is_props -- Scan.optional concl_props [];

val propp = prop -- Scan.optional ($$$ "(" |-- !!! (any_props --| $$$ ")")) ([], []);
val propp' = prop -- Scan.optional ($$$ "(" |-- !!! (is_props --| $$$ ")")) [];
val termp = term -- Scan.optional ($$$ "(" |-- !!! (is_terms --| $$$ ")")) [];


(* arguments *)

fun keyword_symid is_symid = Scan.one (T.keyword_with is_symid) >> T.val_of;
val keyword_sid = keyword_symid T.is_sid;

fun atom_arg is_symid blk =
  group "argument"
    (position (short_ident || long_ident || sym_ident || term_var ||
        type_ident || type_var || number) >> Args.mk_ident ||
      position (keyword_symid is_symid) >> Args.mk_keyword ||
      position (string || verbatim) >> Args.mk_string ||
      position (if blk then $$$ "," else Scan.fail) >> Args.mk_keyword);

fun paren_args l r scan = position ($$$ l) -- !!! (scan true -- position ($$$ r))
  >> (fn (x, (ys, z)) => Args.mk_keyword x :: ys @ [Args.mk_keyword z]);

fun args is_symid blk x = Scan.optional (args1 is_symid blk) [] x
and args1 is_symid blk x =
  ((Scan.repeat1
    (Scan.repeat1 (atom_arg is_symid blk) ||
      paren_args "(" ")" (args is_symid) ||
      paren_args "[" "]" (args is_symid))) >> List.concat) x;

val arguments = args T.is_sid false;


(* theorem specifications *)

val attrib = position ((keyword_sid || xname) -- !!! arguments) >> Args.src;
val attribs = $$$ "[" |-- !!! (list attrib --| $$$ "]");
val opt_attribs = Scan.optional attribs [];

fun thm_name s = name -- opt_attribs --| $$$ s;
fun opt_thm_name s =
  Scan.optional ((name -- opt_attribs || (attribs >> pair "")) --| $$$ s) ("", []);;

val spec_name = thm_name ":" -- prop >> (fn ((x, y), z) => ((x, z), y));
val spec_opt_name = opt_thm_name ":" -- prop >> (fn ((x, y), z) => ((x, z), y));

val thm_sel = $$$ "(" |-- list1
 (nat --| minus -- nat >> PureThy.FromTo ||
  nat --| minus >> PureThy.From ||
  nat >> PureThy.Single) --| $$$ ")";

val xthm = (xname -- thm_sel >> NameSelection || xname >> Name) -- opt_attribs;
val xthms1 = Scan.repeat1 xthm;

val name_facts = and_list1 (opt_thm_name "=" -- xthms1);


(* locale elements *)

local

val loc_mixfix = $$$ "(" -- $$$ "structure" -- !!! ($$$ ")") >> K NONE || opt_mixfix >> SOME;
val loc_keyword = $$$ "fixes" || $$$ "constrains" || $$$ "assumes" ||
   $$$ "defines" || $$$ "notes" || $$$ "includes";

val loc_element =
  $$$ "fixes" |-- !!! (and_list1 (name -- Scan.option ($$$ "::" |-- typ) -- loc_mixfix
    >> triple1)) >> Locale.Fixes ||
  $$$ "constrains" |-- !!! (and_list1 (name -- ($$$ "::" |-- typ)))
    >> Locale.Constrains ||
  $$$ "assumes" |-- !!! (and_list1 (opt_thm_name ":" -- Scan.repeat1 propp))
    >> Locale.Assumes ||
  $$$ "defines" |-- !!! (and_list1 (opt_thm_name ":" -- propp'))
    >> Locale.Defines ||
  $$$ "notes" |-- !!! (and_list1 (opt_thm_name "=" -- xthms1))
    >> Locale.Notes;

fun plus1 scan =
  scan -- Scan.repeat ($$$ "+" |-- Scan.unless loc_keyword (!!! scan)) >> op ::;

val rename = name -- Scan.option mixfix';

fun expr2 x = (xname >> Locale.Locale || $$$ "(" |-- !!! (expr0 --| $$$ ")")) x
and expr1 x = (expr2 -- Scan.repeat1 (maybe rename) >> Locale.Rename || expr2) x
and expr0 x = (plus1 expr1 >> (fn [e] => e | es => Locale.Merge es)) x;

in

val locale_target = ($$$ "(" -- $$$ "in") |-- !!! (xname --| $$$ ")");
val opt_locale_target = Scan.option locale_target;
val locale_expr = expr0;
val locale_keyword = loc_keyword;

val locale_element = group "locale element" loc_element;

val locale_elem_or_expr = group "locale element or includes"
  (loc_element >> Locale.Elem ||
  $$$ "includes" |-- !!! locale_expr >> Locale.Expr);

end;


(* proof methods *)

fun is_symid_meth s =
  s <> "|" andalso s <> "?" andalso s <> "+" andalso T.is_sid s;

fun meth4 x =
 (position (xname >> rpair []) >> (Method.Source o Args.src) ||
  $$$ "(" |-- !!! (meth0 --| $$$ ")")) x
and meth3 x =
 (meth4 --| $$$ "?" >> Method.Try ||
  meth4 --| $$$ "+" >> Method.Repeat1 ||
  meth4) x
and meth2 x =
 (position (xname -- args1 is_symid_meth false) >> (Method.Source o Args.src) ||
  meth3) x
and meth1 x = (enum1 "," meth2 >> (fn [m] => m | ms => Method.Then ms)) x
and meth0 x = (enum1 "|" meth1 >> (fn [m] => m | ms => Method.Orelse ms)) x;

val method = meth3;


end;