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