src/Pure/Isar/args.ML
author wenzelm
Wed, 19 Mar 2008 22:27:57 +0100
changeset 26336 a0e2b706ce73
parent 25999 f8bcd311d501
child 27234 e60cdbc5e8e1
permissions -rw-r--r--
renamed datatype thmref to Facts.ref, tuned interfaces;
     1 (*  Title:      Pure/Isar/args.ML
     2     ID:         $Id$
     3     Author:     Markus Wenzel, TU Muenchen
     4 
     5 Concrete argument syntax of attributes, methods etc. -- with special
     6 support for embedded values and static binding.
     7 *)
     8 
     9 signature ARGS =
    10 sig
    11   datatype value =
    12     Text of string | Typ of typ | Term of term | Fact of thm list |
    13     Attribute of morphism -> attribute
    14   type T
    15   val string_of: T -> string
    16   val mk_ident: string * Position.T -> T
    17   val mk_string: string * Position.T -> T
    18   val mk_alt_string: string * Position.T -> T
    19   val mk_keyword: string * Position.T -> T
    20   val mk_text: string -> T
    21   val mk_typ: typ -> T
    22   val mk_term: term -> T
    23   val mk_fact: thm list -> T
    24   val mk_attribute: (morphism -> attribute) -> T
    25   val stopper: T * (T -> bool)
    26   val not_eof: T -> bool
    27   type src
    28   val src: (string * T list) * Position.T -> src
    29   val dest_src: src -> (string * T list) * Position.T
    30   val pretty_src: Proof.context -> src -> Pretty.T
    31   val map_name: (string -> string) -> src -> src
    32   val morph_values: morphism -> src -> src
    33   val maxidx_values: src -> int -> int
    34   val assignable: src -> src
    35   val assign: value option -> T -> unit
    36   val closure: src -> src
    37   val position: (T list -> 'a * 'b) -> T list -> ('a * Position.T) * 'b
    38   val !!! : (T list -> 'a) -> T list -> 'a
    39   val $$$ : string -> T list -> string * T list
    40   val add: T list -> string * T list
    41   val del: T list -> string * T list
    42   val colon: T list -> string * T list
    43   val query: T list -> string * T list
    44   val bang: T list -> string * T list
    45   val query_colon: T list -> string * T list
    46   val bang_colon: T list -> string * T list
    47   val parens: (T list -> 'a * T list) -> T list -> 'a * T list
    48   val bracks: (T list -> 'a * T list) -> T list -> 'a * T list
    49   val mode: string -> 'a * T list -> bool * ('a * T list)
    50   val maybe: (T list -> 'a * T list) -> T list -> 'a option * T list
    51   val terminator: T list -> T * T list
    52   val name: T list -> string * T list
    53   val alt_name: T list -> string * T list
    54   val symbol: T list -> string * T list
    55   val liberal_name: T list -> string * T list
    56   val nat: T list -> int * T list
    57   val int: T list -> int * T list
    58   val var: T list -> indexname * T list
    59   val list: (T list -> 'a * T list) -> T list -> 'a list * T list
    60   val list1: (T list -> 'a * T list) -> T list -> 'a list * T list
    61   val enum: string -> ('a * T list -> 'b * ('a * T list)) -> 'a * T list -> 'b list * ('a * T list)
    62   val enum1: string -> ('a * T list -> 'b * ('a * T list)) -> 'a * T list -> 'b list * ('a * T list)
    63   val and_list: ('a * T list -> 'b * ('a * T list)) -> 'a * T list -> 'b list * ('a * T list)
    64   val and_list1: ('a * T list -> 'b * ('a * T list)) -> 'a * T list -> 'b list * ('a * T list)
    65   val ahead: T list -> T * T list
    66   val internal_text: T list -> string * T list
    67   val internal_typ: T list -> typ * T list
    68   val internal_term: T list -> term * T list
    69   val internal_fact: T list -> thm list * T list
    70   val internal_attribute: T list -> (morphism -> attribute) * T list
    71   val named_text: (string -> string) -> T list -> string * T list
    72   val named_typ: (string -> typ) -> T list -> typ * T list
    73   val named_term: (string -> term) -> T list -> term * T list
    74   val named_fact: (string -> thm list) -> T list -> thm list * T list
    75   val named_attribute: (string -> morphism -> attribute) -> T list ->
    76     (morphism -> attribute) * T list
    77   val typ_abbrev: Context.generic * T list -> typ * (Context.generic * T list)
    78   val typ: Context.generic * T list -> typ * (Context.generic * T list)
    79   val term: Context.generic * T list -> term * (Context.generic * T list)
    80   val term_abbrev: Context.generic * T list -> term * (Context.generic * T list)
    81   val prop: Context.generic * T list -> term * (Context.generic * T list)
    82   val tyname: Context.generic * T list -> string * (Context.generic * T list)
    83   val const: Context.generic * T list -> string * (Context.generic * T list)
    84   val const_proper: Context.generic * T list -> string * (Context.generic * T list)
    85   val thm_sel: T list -> Facts.interval list * T list
    86   val bang_facts: Context.generic * T list -> thm list * (Context.generic * T list)
    87   val goal_spec: ((int -> tactic) -> tactic) -> ('a * T list)
    88     -> ((int -> tactic) -> tactic) * ('a * T list)
    89   val attribs: (string -> string) -> T list -> src list * T list
    90   val opt_attribs: (string -> string) -> T list -> src list * T list
    91   val syntax: string -> ('b * T list -> 'a * ('b * T list)) -> src -> 'b -> 'a * 'b
    92   val context_syntax: string -> (Context.generic * T list -> 'a * (Context.generic * T list)) ->
    93     src -> Proof.context -> 'a * Proof.context
    94 end;
    95 
    96 structure Args: ARGS =
    97 struct
    98 
    99 
   100 (** datatype T **)
   101 
   102 (*An argument token is a symbol (with raw string value), together with
   103   an optional assigned internal value.  Note that an assignable ref
   104   designates an intermediate state of internalization -- it is NOT
   105   meant to persist.*)
   106 
   107 datatype kind = Ident | String | AltString | Keyword | EOF;
   108 
   109 type symbol = kind * string * Position.T;
   110 
   111 datatype value =
   112   Text of string |
   113   Typ of typ |
   114   Term of term |
   115   Fact of thm list |
   116   Attribute of morphism -> attribute;
   117 
   118 datatype slot =
   119   Empty |
   120   Value of value option |
   121   Assignable of value option ref;
   122 
   123 datatype T = Arg of symbol * slot;
   124 
   125 fun string_of (Arg ((Ident, x, _), _)) = x
   126   | string_of (Arg ((String, x, _), _)) = quote x
   127   | string_of (Arg ((AltString, x, _), _)) = enclose "`" "`" x
   128   | string_of (Arg ((Keyword, x, _), _)) = x
   129   | string_of (Arg ((EOF, _, _), _)) = "";
   130 
   131 fun sym_of (Arg ((_, x, _), _)) = x;
   132 fun pos_of (Arg ((_, _, pos), _)) = pos;
   133 
   134 
   135 (* basic constructors *)
   136 
   137 fun mk_symbol k (x, p) = Arg ((k, x, p), Empty);
   138 fun mk_value k v = Arg ((Keyword, k, Position.none), Value (SOME v));
   139 
   140 val mk_ident = mk_symbol Ident;
   141 val mk_string = mk_symbol String;
   142 val mk_alt_string = mk_symbol AltString;
   143 val mk_keyword = mk_symbol Keyword;
   144 val mk_text = mk_value "<text>" o Text;
   145 val mk_typ = mk_value "<typ>" o Typ;
   146 val mk_term = mk_value "<term>" o Term;
   147 val mk_fact = mk_value "<fact>" o Fact;
   148 val mk_attribute = mk_value "<attribute>" o Attribute;
   149 
   150 
   151 (* eof *)
   152 
   153 val eof = mk_symbol EOF ("", Position.none);
   154 
   155 fun is_eof (Arg ((EOF, _, _), _)) = true
   156   | is_eof _ = false;
   157 
   158 val stopper = (eof, is_eof);
   159 val not_eof = not o is_eof;
   160 
   161 
   162 
   163 (** datatype src **)
   164 
   165 datatype src = Src of (string * T list) * Position.T;
   166 
   167 val src = Src;
   168 fun dest_src (Src src) = src;
   169 
   170 fun pretty_src ctxt src =
   171   let
   172     val prt_thm = Pretty.backquote o ProofContext.pretty_thm ctxt;
   173     fun prt (Arg (_, Value (SOME (Text s)))) = Pretty.str (quote s)
   174       | prt (Arg (_, Value (SOME (Typ T)))) = Syntax.pretty_typ ctxt T
   175       | prt (Arg (_, Value (SOME (Term t)))) = Syntax.pretty_term ctxt t
   176       | prt (Arg (_, Value (SOME (Fact ths)))) =
   177           Pretty.enclose "(" ")" (Pretty.breaks (map prt_thm ths))
   178       | prt arg = Pretty.str (string_of arg);
   179     val (s, args) = #1 (dest_src src);
   180   in Pretty.block (Pretty.breaks (Pretty.str s :: map prt args)) end;
   181 
   182 fun map_name f (Src ((s, args), pos)) = Src ((f s, args), pos);
   183 fun map_args f (Src ((s, args), pos)) = Src ((s, map f args), pos);
   184 
   185 
   186 (* values *)
   187 
   188 fun map_value f (Arg (s, Value (SOME v))) = Arg (s, Value (SOME (f v)))
   189   | map_value _ arg = arg;
   190 
   191 fun morph_values phi = map_args (map_value
   192   (fn Text s => Text s
   193     | Typ T => Typ (Morphism.typ phi T)
   194     | Term t => Term (Morphism.term phi t)
   195     | Fact ths => Fact (Morphism.fact phi ths)
   196     | Attribute att => Attribute (Morphism.transform phi att)));
   197 
   198 fun maxidx_values (Src ((_, args), _)) = args |> fold
   199   (fn (Arg (_, Value (SOME (Typ T)))) => Term.maxidx_typ T
   200     | (Arg (_, Value (SOME (Term t)))) => Term.maxidx_term t
   201     | (Arg (_, Value (SOME (Fact ths)))) => fold Thm.maxidx_thm ths
   202     | _ => I);
   203 
   204 
   205 (* static binding *)
   206 
   207 (*1st stage: make empty slots assignable*)
   208 val assignable =
   209   map_args (fn Arg (s, Empty) => Arg (s, Assignable (ref NONE)) | arg => arg);
   210 
   211 (*2nd stage: assign values as side-effect of scanning*)
   212 fun assign v (arg as Arg (_, Assignable r)) = r := v
   213   | assign _ _ = ();
   214 
   215 val ahead = Scan.ahead (Scan.one not_eof);
   216 fun touch scan = ahead -- scan >> (fn (arg, y) => (assign NONE arg; y));
   217 
   218 (*3rd stage: static closure of final values*)
   219 val closure =
   220   map_args (fn Arg (s, Assignable (ref v)) => Arg (s, Value v) | arg => arg);
   221 
   222 
   223 
   224 (** scanners **)
   225 
   226 (* position *)
   227 
   228 fun position scan =
   229   (Scan.ahead (Scan.one not_eof) >> pos_of) -- scan >> Library.swap;
   230 
   231 
   232 (* cut *)
   233 
   234 fun !!! scan =
   235   let
   236     fun get_pos [] = " (past end-of-text!)"
   237       | get_pos (arg :: _) = Position.str_of (pos_of arg);
   238 
   239     fun err (args, NONE) = "Argument syntax error" ^ get_pos args
   240       | err (args, SOME msg) = "Argument syntax error" ^ get_pos args ^ ": " ^ msg;
   241   in Scan.!! err scan end;
   242 
   243 
   244 (* basic *)
   245 
   246 fun some_ident f =
   247   touch (Scan.some (fn Arg ((Ident, x, _), _) => f x | _ => NONE));
   248 
   249 val named =
   250   touch (Scan.one (fn Arg ((k, _, _), _) => k = Ident orelse k = String));
   251 
   252 val alt_string =
   253   touch (Scan.one (fn Arg ((k, _, _), _) => k = AltString));
   254 
   255 val symbolic =
   256   touch (Scan.one (fn Arg ((k, x, _), _) => k = Keyword andalso OuterLex.is_sid x));
   257 
   258 fun &&& x =
   259   touch (Scan.one (fn Arg ((k, y, _), _) => k = Keyword andalso x = y));
   260 
   261 fun $$$ x =
   262   touch (Scan.one (fn Arg ((k, y, _), _) => (k = Ident orelse k = Keyword) andalso x = y))
   263   >> sym_of;
   264 
   265 val add = $$$ "add";
   266 val del = $$$ "del";
   267 val colon = $$$ ":";
   268 val query = $$$ "?";
   269 val bang = $$$ "!";
   270 val query_colon = $$$ "?" ^^ $$$ ":";
   271 val bang_colon = $$$ "!" ^^ $$$ ":";
   272 
   273 fun parens scan = $$$ "(" |-- scan --| $$$ ")";
   274 fun bracks scan = $$$ "[" |-- scan --| $$$ "]";
   275 fun mode s = Scan.lift (Scan.optional (parens ($$$ s) >> K true) false);
   276 fun maybe scan = $$$ "_" >> K NONE || scan >> SOME;
   277 val terminator = Scan.ahead (Scan.one is_eof);
   278 
   279 val name = named >> sym_of;
   280 val alt_name = alt_string >> sym_of;
   281 val symbol = symbolic >> sym_of;
   282 val liberal_name = symbol || name;
   283 
   284 val nat = some_ident Lexicon.read_nat;
   285 val int = Scan.optional ($$$ "-" >> K ~1) 1 -- nat >> op *;
   286 val var = some_ident Lexicon.read_variable;
   287 
   288 
   289 (* enumerations *)
   290 
   291 fun list1 scan = scan ::: Scan.repeat ($$$ "," |-- scan);
   292 fun list scan = list1 scan || Scan.succeed [];
   293 
   294 fun enum1 sep scan = scan ::: Scan.repeat (Scan.lift ($$$ sep) |-- scan);
   295 fun enum sep scan = enum1 sep scan || Scan.succeed [];
   296 
   297 fun and_list1 scan = enum1 "and" scan;
   298 fun and_list scan = enum "and" scan;
   299 
   300 
   301 (* values *)
   302 
   303 fun value dest = Scan.some  (*no touch here*)
   304   (fn Arg (_, Value (SOME v)) => (SOME (dest v) handle Match => NONE)
   305     | Arg _ => NONE);
   306 
   307 fun evaluate mk eval arg =
   308   let val x = eval (sym_of arg) in (assign (SOME (mk x)) arg; x) end;
   309 
   310 val internal_text = value (fn Text s => s);
   311 val internal_typ = value (fn Typ T => T);
   312 val internal_term = value (fn Term t => t);
   313 val internal_fact = value (fn Fact ths => ths);
   314 val internal_attribute = value (fn Attribute att => att);
   315 
   316 fun named_text intern = internal_text || named >> evaluate Text intern;
   317 fun named_typ readT = internal_typ || named >> evaluate Typ readT;
   318 fun named_term read = internal_term || named >> evaluate Term read;
   319 fun named_fact get = internal_fact || (alt_string || named) >> evaluate Fact get;
   320 fun named_attribute att = internal_attribute || named >> evaluate Attribute att;
   321 
   322 
   323 (* terms and types *)
   324 
   325 val typ_abbrev = Scan.peek (named_typ o ProofContext.read_typ_abbrev o Context.proof_of);
   326 val typ = Scan.peek (named_typ o Syntax.read_typ o Context.proof_of);
   327 val term = Scan.peek (named_term o Syntax.read_term o Context.proof_of);
   328 val term_abbrev = Scan.peek (named_term o ProofContext.read_term_abbrev o Context.proof_of);
   329 val prop = Scan.peek (named_term o Syntax.read_prop o Context.proof_of);
   330 
   331 
   332 (* type and constant names *)
   333 
   334 val tyname = Scan.peek (named_typ o ProofContext.read_tyname o Context.proof_of)
   335   >> (fn Type (c, _) => c | TFree (a, _) => a | _ => "");
   336 
   337 val const = Scan.peek (named_term o ProofContext.read_const o Context.proof_of)
   338   >> (fn Const (c, _) => c | Free (x, _) => x | _ => "");
   339 
   340 val const_proper = Scan.peek (named_term o ProofContext.read_const_proper o Context.proof_of)
   341   >> (fn Const (c, _) => c | _ => "");
   342 
   343 
   344 (* misc *)
   345 
   346 val thm_sel = parens (list1
   347  (nat --| $$$ "-" -- nat >> Facts.FromTo ||
   348   nat --| $$$ "-" >> Facts.From ||
   349   nat >> Facts.Single));
   350 
   351 val bang_facts = Scan.peek (fn context =>
   352   $$$ "!" >> (fn _ => (warning "use of prems in proof method";
   353     Assumption.prems_of (Context.proof_of context))) || Scan.succeed []);
   354 
   355 
   356 (* goal specification *)
   357 
   358 (* range *)
   359 
   360 val from_to =
   361   nat -- ($$$ "-" |-- nat) >> (fn (i, j) => fn tac => Seq.INTERVAL tac i j) ||
   362   nat --| $$$ "-" >> (fn i => fn tac => fn st => Seq.INTERVAL tac i (Thm.nprems_of st) st) ||
   363   nat >> (fn i => fn tac => tac i) ||
   364   $$$ "!" >> K ALLGOALS;
   365 
   366 val goal = $$$ "[" |-- !!! (from_to --| $$$ "]");
   367 fun goal_spec def = Scan.lift (Scan.optional goal def);
   368 
   369 
   370 
   371 (* attribs *)
   372 
   373 local
   374 
   375 val exclude = member (op =) (explode "()[],");
   376 
   377 fun atomic blk = touch (Scan.one (fn Arg ((k, x, _), _) =>
   378     k <> Keyword orelse not (exclude x) orelse blk andalso x = ","));
   379 
   380 fun args blk x = Scan.optional (args1 blk) [] x
   381 and args1 blk x =
   382   ((Scan.repeat1
   383     (Scan.repeat1 (atomic blk) ||
   384       argsp "(" ")" ||
   385       argsp "[" "]")) >> flat) x
   386 and argsp l r x =
   387   (Scan.trace (&&& l -- !!! (args true -- &&& r)) >> #2) x;
   388 
   389 in
   390 
   391 fun attribs intern =
   392   let
   393     val attrib_name = internal_text || (symbolic || named) >> evaluate Text intern;
   394     val attrib = position (attrib_name -- !!! (args false)) >> src;
   395   in $$$ "[" |-- !!! (list attrib --| $$$ "]") end;
   396 
   397 fun opt_attribs intern = Scan.optional (attribs intern) [];
   398 
   399 end;
   400 
   401 
   402 (* syntax interface *)
   403 
   404 fun syntax kind scan (src as Src ((s, args), pos)) st =
   405   (case Scan.error (Scan.finite' stopper (Scan.option scan)) (st, args) of
   406     (SOME x, (st', [])) => (x, st')
   407   | (_, (_, args')) =>
   408       error (kind ^ " " ^ quote s ^ Position.str_of pos ^ ": bad arguments\n  " ^
   409         space_implode " " (map string_of args')));
   410 
   411 fun context_syntax kind scan src = apsnd Context.the_proof o syntax kind scan src o Context.Proof;
   412 
   413 end;