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