author | wenzelm |
Wed, 12 Jul 2006 19:59:13 +0200 | |
changeset 20111 | ba1676dd3546 |
parent 19482 | 9f11af8f7ef9 |
child 20224 | 9c40a144ee0e |
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 |
18728 | 11 |
datatype value = |
12 |
Name of string | Typ of typ | Term of term | Fact of thm list | Attribute of 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 |
|
18037 | 17 |
val mk_alt_string: string * Position.T -> T |
15703 | 18 |
val mk_keyword: string * Position.T -> T |
19 |
val mk_name: string -> T |
|
20 |
val mk_typ: typ -> T |
|
21 |
val mk_term: term -> T |
|
22 |
val mk_fact: thm list -> T |
|
18728 | 23 |
val mk_attribute: attribute -> T |
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
24 |
val stopper: T * (T -> bool) |
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
25 |
val not_eof: T -> bool |
15703 | 26 |
type src |
27 |
val src: (string * T list) * Position.T -> src |
|
28 |
val dest_src: src -> (string * T list) * Position.T |
|
29 |
val map_name: (string -> string) -> src -> src |
|
30 |
val map_values: (string -> string) -> (typ -> typ) -> (term -> term) -> (thm -> thm) |
|
31 |
-> src -> src |
|
32 |
val assignable: src -> src |
|
33 |
val assign: value option -> T -> unit |
|
34 |
val closure: src -> src |
|
5878 | 35 |
val position: (T list -> 'a * 'b) -> T list -> ('a * Position.T) * 'b |
36 |
val !!! : (T list -> 'a) -> T list -> 'a |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
37 |
val $$$ : string -> T list -> string * T list |
10035 | 38 |
val add: T list -> string * T list |
39 |
val del: T list -> string * T list |
|
8803 | 40 |
val colon: T list -> string * T list |
10035 | 41 |
val query: T list -> string * T list |
42 |
val bang: T list -> string * T list |
|
43 |
val query_colon: T list -> string * T list |
|
44 |
val bang_colon: T list -> string * T list |
|
8803 | 45 |
val parens: (T list -> 'a * T list) -> T list -> 'a * T list |
10150 | 46 |
val bracks: (T list -> 'a * T list) -> T list -> 'a * T list |
9809 | 47 |
val mode: string -> 'a * T list -> bool * ('a * T list) |
15703 | 48 |
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
|
49 |
val name: T list -> string * T list |
18037 | 50 |
val alt_name: T list -> string * T list |
16140 | 51 |
val symbol: T list -> string * T list |
17064 | 52 |
val liberal_name: T list -> string * T list |
5878 | 53 |
val nat: T list -> int * T list |
9538 | 54 |
val int: T list -> int * T list |
5878 | 55 |
val var: T list -> indexname * T list |
15703 | 56 |
val list: (T list -> 'a * T list) -> T list -> 'a list * T list |
57 |
val list1: (T list -> 'a * T list) -> T list -> 'a list * T list |
|
6447 | 58 |
val enum: string -> ('a * T list -> 'b * ('a * T list)) -> 'a * T list -> 'b list * ('a * T list) |
59 |
val enum1: string -> ('a * T list -> 'b * ('a * T list)) -> 'a * T list -> 'b list * ('a * T list) |
|
60 |
val and_list: ('a * T list -> 'b * ('a * T list)) -> 'a * T list -> 'b list * ('a * T list) |
|
61 |
val and_list1: ('a * T list -> 'b * ('a * T list)) -> 'a * T list -> 'b list * ('a * T list) |
|
15703 | 62 |
val ahead: T list -> T * T list |
63 |
val internal_name: T list -> string * T list |
|
64 |
val internal_typ: T list -> typ * T list |
|
65 |
val internal_term: T list -> term * T list |
|
66 |
val internal_fact: T list -> thm list * T list |
|
18728 | 67 |
val internal_attribute: T list -> attribute * T list |
15703 | 68 |
val named_name: (string -> string) -> T list -> string * T list |
69 |
val named_typ: (string -> typ) -> T list -> typ * T list |
|
70 |
val named_term: (string -> term) -> T list -> term * T list |
|
71 |
val named_fact: (string -> thm list) -> T list -> thm list * T list |
|
18635 | 72 |
val typ_abbrev: Context.generic * T list -> typ * (Context.generic * T list) |
73 |
val typ: Context.generic * T list -> typ * (Context.generic * T list) |
|
74 |
val term: Context.generic * T list -> term * (Context.generic * T list) |
|
75 |
val prop: Context.generic * T list -> term * (Context.generic * T list) |
|
76 |
val tyname: Context.generic * T list -> string * (Context.generic * T list) |
|
77 |
val const: Context.generic * T list -> string * (Context.generic * T list) |
|
15703 | 78 |
val thm_sel: T list -> PureThy.interval list * T list |
18998 | 79 |
val bang_facts: Context.generic * T list -> thm list * (Context.generic * T list) |
8536 | 80 |
val goal_spec: ((int -> tactic) -> tactic) -> ('a * T list) |
81 |
-> ((int -> tactic) -> tactic) * ('a * T list) |
|
15703 | 82 |
val attribs: (string -> string) -> T list -> src list * T list |
83 |
val opt_attribs: (string -> string) -> T list -> src list * T list |
|
18998 | 84 |
val syntax: string -> ('b * T list -> 'a * ('b * T list)) -> src -> 'b -> 'b * 'a |
85 |
val context_syntax: string -> (Context.generic * T list -> 'a * (Context.generic * T list)) -> |
|
86 |
src -> ProofContext.context -> ProofContext.context * 'a |
|
15703 | 87 |
val pretty_src: ProofContext.context -> src -> Pretty.T |
88 |
val pretty_attribs: ProofContext.context -> src list -> Pretty.T list |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
89 |
end; |
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
90 |
|
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
91 |
structure Args: ARGS = |
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
92 |
struct |
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
93 |
|
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
94 |
|
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
95 |
(** datatype T **) |
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
96 |
|
15703 | 97 |
(*An argument token is a symbol (with raw string value), together with |
98 |
an optional assigned internal value. Note that an assignable ref |
|
99 |
designates an intermediate state of internalization -- it is NOT |
|
100 |
meant to persist.*) |
|
101 |
||
18037 | 102 |
datatype kind = Ident | String | AltString | Keyword | EOF; |
15703 | 103 |
|
104 |
type symbol = kind * string * Position.T; |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
105 |
|
15703 | 106 |
datatype value = |
107 |
Name of string | |
|
108 |
Typ of typ | |
|
109 |
Term of term | |
|
110 |
Fact of thm list | |
|
18728 | 111 |
Attribute of attribute; |
15703 | 112 |
|
113 |
datatype slot = |
|
114 |
Empty | |
|
115 |
Value of value option | |
|
116 |
Assignable of value option ref; |
|
117 |
||
118 |
datatype T = Arg of symbol * slot; |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
119 |
|
15703 | 120 |
fun string_of (Arg ((Ident, x, _), _)) = x |
18037 | 121 |
| string_of (Arg ((String, x, _), _)) = quote x |
122 |
| string_of (Arg ((AltString, x, _), _)) = enclose "`" "`" x |
|
15703 | 123 |
| string_of (Arg ((Keyword, x, _), _)) = x |
124 |
| string_of (Arg ((EOF, _, _), _)) = ""; |
|
125 |
||
126 |
fun sym_of (Arg ((_, x, _), _)) = x; |
|
127 |
fun pos_of (Arg ((_, _, pos), _)) = pos; |
|
128 |
||
129 |
||
130 |
(* basic constructors *) |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
131 |
|
15703 | 132 |
fun mk_symbol k (x, p) = Arg ((k, x, p), Empty); |
133 |
fun mk_value k v = Arg ((Keyword, k, Position.none), Value (SOME v)); |
|
9748 | 134 |
|
15703 | 135 |
val mk_ident = mk_symbol Ident; |
136 |
val mk_string = mk_symbol String; |
|
18037 | 137 |
val mk_alt_string = mk_symbol AltString; |
15703 | 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 |
||
18037 | 229 |
val alt_string = |
230 |
touch (Scan.one (fn Arg ((k, _, _), _) => k = AltString)); |
|
231 |
||
15703 | 232 |
val symbolic = |
233 |
touch (Scan.one (fn Arg ((k, x, _), _) => k = Keyword andalso OuterLex.is_sid x)); |
|
234 |
||
235 |
fun &&& x = |
|
236 |
touch (Scan.one (fn Arg ((k, y, _), _) => k = Keyword andalso x = y)); |
|
237 |
||
238 |
fun $$$ x = |
|
16140 | 239 |
touch (Scan.one (fn Arg ((k, y, _), _) => (k = Ident orelse k = Keyword) andalso x = y)) |
240 |
>> sym_of; |
|
5878 | 241 |
|
10035 | 242 |
val add = $$$ "add"; |
243 |
val del = $$$ "del"; |
|
8803 | 244 |
val colon = $$$ ":"; |
10035 | 245 |
val query = $$$ "?"; |
246 |
val bang = $$$ "!"; |
|
20111 | 247 |
val query_colon = $$$ "?" ^^ $$$ ":"; |
248 |
val bang_colon = $$$ "!" ^^ $$$ ":"; |
|
10035 | 249 |
|
8803 | 250 |
fun parens scan = $$$ "(" |-- scan --| $$$ ")"; |
10150 | 251 |
fun bracks scan = $$$ "[" |-- scan --| $$$ "]"; |
15703 | 252 |
fun mode s = Scan.lift (Scan.optional (parens ($$$ s) >> K true) false); |
253 |
fun maybe scan = $$$ "_" >> K NONE || scan >> SOME; |
|
5878 | 254 |
|
15703 | 255 |
val name = named >> sym_of; |
18037 | 256 |
val alt_name = alt_string >> sym_of; |
16140 | 257 |
val symbol = symbolic >> sym_of; |
17064 | 258 |
val liberal_name = symbol || name; |
8233 | 259 |
|
15703 | 260 |
val nat = some_ident Syntax.read_nat; |
9538 | 261 |
val int = Scan.optional ($$$ "-" >> K ~1) 1 -- nat >> op *; |
15987 | 262 |
val var = some_ident Syntax.read_variable; |
14563 | 263 |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
264 |
|
5878 | 265 |
(* enumerations *) |
266 |
||
15703 | 267 |
fun list1 scan = scan -- Scan.repeat ($$$ "," |-- scan) >> op ::; |
268 |
fun list scan = list1 scan || Scan.succeed []; |
|
269 |
||
6447 | 270 |
fun enum1 sep scan = scan -- Scan.repeat (Scan.lift ($$$ sep) |-- scan) >> op ::; |
5878 | 271 |
fun enum sep scan = enum1 sep scan || Scan.succeed []; |
272 |
||
6447 | 273 |
fun and_list1 scan = enum1 "and" scan; |
274 |
fun and_list scan = enum "and" scan; |
|
5878 | 275 |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
276 |
|
15703 | 277 |
(* values *) |
278 |
||
279 |
fun value dest = Scan.some (*no touch here*) |
|
280 |
(fn Arg (_, Value (SOME v)) => (SOME (dest v) handle Match => NONE) |
|
281 |
| Arg _ => NONE); |
|
282 |
||
283 |
fun evaluate mk eval arg = |
|
284 |
let val x = eval (sym_of arg) in (assign (SOME (mk x)) arg; x) end; |
|
285 |
||
286 |
val internal_name = value (fn Name s => s); |
|
287 |
val internal_typ = value (fn Typ T => T); |
|
288 |
val internal_term = value (fn Term t => t); |
|
289 |
val internal_fact = value (fn Fact ths => ths); |
|
290 |
val internal_attribute = value (fn Attribute att => att); |
|
291 |
||
292 |
fun named_name intern = internal_name || named >> evaluate Name intern; |
|
293 |
fun named_typ readT = internal_typ || named >> evaluate Typ readT; |
|
294 |
fun named_term read = internal_term || named >> evaluate Term read; |
|
18037 | 295 |
fun named_fact get = internal_fact || (alt_string || named) >> evaluate Fact get; |
15703 | 296 |
|
297 |
||
5878 | 298 |
(* terms and types *) |
299 |
||
18635 | 300 |
val typ_abbrev = Scan.peek (named_typ o ProofContext.read_typ_abbrev o Context.proof_of); |
301 |
val typ = Scan.peek (named_typ o ProofContext.read_typ o Context.proof_of); |
|
302 |
val term = Scan.peek (named_term o ProofContext.read_term o Context.proof_of); |
|
303 |
val prop = Scan.peek (named_term o ProofContext.read_prop o Context.proof_of); |
|
304 |
||
5878 | 305 |
|
15703 | 306 |
(* type and constant names *) |
307 |
||
18998 | 308 |
val tyname = Scan.peek (named_typ o Context.cases Sign.read_tyname ProofContext.read_tyname) |
309 |
>> (fn Type (c, _) => c | TFree (a, _) => a | _ => ""); |
|
15703 | 310 |
|
18998 | 311 |
val const = Scan.peek (named_term o Context.cases Sign.read_const ProofContext.read_const) |
312 |
>> (fn Const (c, _) => c | Free (x, _) => x | _ => ""); |
|
7553 | 313 |
|
15703 | 314 |
|
315 |
(* misc *) |
|
316 |
||
317 |
val thm_sel = parens (list1 |
|
318 |
(nat --| $$$ "-" -- nat >> PureThy.FromTo || |
|
319 |
nat --| $$$ "-" >> PureThy.From || |
|
320 |
nat >> PureThy.Single)); |
|
321 |
||
18998 | 322 |
val bang_facts = Scan.peek (fn context => |
323 |
$$$ "!" >> K (ProofContext.prems_of (Context.proof_of context)) || Scan.succeed []); |
|
7553 | 324 |
|
325 |
||
8536 | 326 |
(* goal specification *) |
327 |
||
328 |
(* range *) |
|
8233 | 329 |
|
8536 | 330 |
val from_to = |
15703 | 331 |
nat -- ($$$ "-" |-- nat) >> (fn (i, j) => fn tac => Seq.INTERVAL tac i j) || |
332 |
nat --| $$$ "-" >> (fn i => fn tac => fn st => Seq.INTERVAL tac i (Thm.nprems_of st) st) || |
|
8549 | 333 |
nat >> (fn i => fn tac => tac i) || |
15703 | 334 |
$$$ "!" >> K ALLGOALS; |
8536 | 335 |
|
15703 | 336 |
val goal = $$$ "[" |-- !!! (from_to --| $$$ "]"); |
8536 | 337 |
fun goal_spec def = Scan.lift (Scan.optional goal def); |
8233 | 338 |
|
339 |
||
15703 | 340 |
|
341 |
(* attribs *) |
|
342 |
||
343 |
local |
|
5878 | 344 |
|
9126 | 345 |
val exclude = explode "()[],"; |
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
346 |
|
15703 | 347 |
fun atomic blk = touch (Scan.one (fn Arg ((k, x, _), _) => |
348 |
k <> Keyword orelse not (x mem exclude) orelse blk andalso x = ",")); |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
349 |
|
5878 | 350 |
fun args blk x = Scan.optional (args1 blk) [] x |
351 |
and args1 blk x = |
|
352 |
((Scan.repeat1 |
|
15703 | 353 |
(Scan.repeat1 (atomic blk) || |
354 |
argsp "(" ")" || |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
18998
diff
changeset
|
355 |
argsp "[" "]")) >> flat) x |
15703 | 356 |
and argsp l r x = |
357 |
(Scan.trace (&&& l -- !!! (args true -- &&& r)) >> #2) x; |
|
358 |
||
359 |
in |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
360 |
|
15703 | 361 |
fun attribs intern = |
362 |
let |
|
363 |
val attrib_name = internal_name || (symbolic || named) >> evaluate Name intern; |
|
364 |
val attrib = position (attrib_name -- !!! (args false)) >> src; |
|
365 |
in $$$ "[" |-- !!! (list attrib --| $$$ "]") end; |
|
366 |
||
367 |
fun opt_attribs intern = Scan.optional (attribs intern) []; |
|
368 |
||
369 |
end; |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
370 |
|
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
371 |
|
15703 | 372 |
(* syntax interface *) |
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
373 |
|
8282 | 374 |
fun syntax kind scan (src as Src ((s, args), pos)) st = |
9901 | 375 |
(case Scan.error (Scan.finite' stopper (Scan.option scan)) (st, args) of |
15531 | 376 |
(SOME x, (st', [])) => (st', x) |
15703 | 377 |
| (_, (_, args')) => |
378 |
error (kind ^ " " ^ quote s ^ Position.str_of pos ^ ": bad arguments\n " ^ |
|
379 |
space_implode " " (map string_of args'))); |
|
380 |
||
18998 | 381 |
fun context_syntax kind scan src = apfst Context.the_proof o syntax kind scan src o Context.Proof; |
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; |