| author | wenzelm |
| Thu, 14 Aug 2008 16:52:51 +0200 | |
| changeset 27866 | c721ea6e0eb4 |
| parent 27819 | 6398f7aabdfc |
| child 27882 | eaa9fef9f4c1 |
| 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 |
|
| 27811 | 5 |
Parsing with implicit value assigment. Concrete argument syntax of |
6 |
attributes, methods etc. |
|
|
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 |
| 27811 | 11 |
type T = OuterLex.token |
| 15703 | 12 |
type src |
13 |
val src: (string * T list) * Position.T -> src |
|
14 |
val dest_src: src -> (string * T list) * Position.T |
|
| 21030 | 15 |
val pretty_src: Proof.context -> src -> Pretty.T |
| 15703 | 16 |
val map_name: (string -> string) -> src -> src |
| 21480 | 17 |
val morph_values: morphism -> src -> src |
| 20263 | 18 |
val maxidx_values: src -> int -> int |
| 15703 | 19 |
val assignable: src -> src |
20 |
val closure: src -> src |
|
| 27371 | 21 |
val context: Context.generic * T list -> Context.proof * (Context.generic * T list) |
22 |
val theory: Context.generic * T list -> Context.theory * (Context.generic * T list) |
|
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
23 |
val $$$ : string -> T list -> string * T list |
| 10035 | 24 |
val add: T list -> string * T list |
25 |
val del: T list -> string * T list |
|
| 8803 | 26 |
val colon: T list -> string * T list |
| 10035 | 27 |
val query: T list -> string * T list |
28 |
val bang: T list -> string * T list |
|
29 |
val query_colon: T list -> string * T list |
|
30 |
val bang_colon: T list -> string * T list |
|
| 8803 | 31 |
val parens: (T list -> 'a * T list) -> T list -> 'a * T list |
| 10150 | 32 |
val bracks: (T list -> 'a * T list) -> T list -> 'a * T list |
| 9809 | 33 |
val mode: string -> 'a * T list -> bool * ('a * T list)
|
| 15703 | 34 |
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
|
35 |
val name: T list -> string * T list |
| 18037 | 36 |
val alt_name: T list -> string * T list |
| 16140 | 37 |
val symbol: T list -> string * T list |
| 17064 | 38 |
val liberal_name: T list -> string * T list |
| 5878 | 39 |
val var: T list -> indexname * T list |
|
21496
a3ac0c55393f
renamed Name value to Text, which is *not* a name in terms of morphisms;
wenzelm
parents:
21480
diff
changeset
|
40 |
val internal_text: T list -> string * T list |
| 15703 | 41 |
val internal_typ: T list -> typ * T list |
42 |
val internal_term: T list -> term * T list |
|
43 |
val internal_fact: T list -> thm list * T list |
|
| 21662 | 44 |
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
|
45 |
val named_text: (string -> string) -> T list -> string * T list |
| 15703 | 46 |
val named_typ: (string -> typ) -> T list -> typ * T list |
47 |
val named_term: (string -> term) -> T list -> term * T list |
|
48 |
val named_fact: (string -> thm list) -> T list -> thm list * T list |
|
| 24002 | 49 |
val named_attribute: (string -> morphism -> attribute) -> T list -> |
50 |
(morphism -> attribute) * T list |
|
| 18635 | 51 |
val typ_abbrev: Context.generic * T list -> typ * (Context.generic * T list) |
52 |
val typ: Context.generic * T list -> typ * (Context.generic * T list) |
|
53 |
val term: Context.generic * T list -> term * (Context.generic * T list) |
|
| 21724 | 54 |
val term_abbrev: Context.generic * T list -> term * (Context.generic * T list) |
| 18635 | 55 |
val prop: Context.generic * T list -> term * (Context.generic * T list) |
56 |
val tyname: Context.generic * T list -> string * (Context.generic * T list) |
|
57 |
val const: Context.generic * T list -> string * (Context.generic * T list) |
|
| 25343 | 58 |
val const_proper: Context.generic * T list -> string * (Context.generic * T list) |
| 18998 | 59 |
val bang_facts: Context.generic * T list -> thm list * (Context.generic * T list) |
| 8536 | 60 |
val goal_spec: ((int -> tactic) -> tactic) -> ('a * T list)
|
61 |
-> ((int -> tactic) -> tactic) * ('a * T list)
|
|
| 27811 | 62 |
val parse: OuterLex.token list -> T list * OuterLex.token list |
63 |
val parse1: (string -> bool) -> OuterLex.token list -> T list * OuterLex.token list |
|
| 15703 | 64 |
val attribs: (string -> string) -> T list -> src list * T list |
65 |
val opt_attribs: (string -> string) -> T list -> src list * T list |
|
| 27377 | 66 |
val thm_name: (string -> string) -> string -> T list -> (string * src list) * T list |
67 |
val opt_thm_name: (string -> string) -> string -> T list -> (string * src list) * T list |
|
| 21879 | 68 |
val syntax: string -> ('b * T list -> 'a * ('b * T list)) -> src -> 'b -> 'a * 'b
|
| 18998 | 69 |
val context_syntax: string -> (Context.generic * T list -> 'a * (Context.generic * T list)) -> |
| 21879 | 70 |
src -> Proof.context -> 'a * Proof.context |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
71 |
end; |
|
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
72 |
|
|
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
73 |
structure Args: ARGS = |
|
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
74 |
struct |
|
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
75 |
|
| 27811 | 76 |
structure T = OuterLex; |
77 |
structure P = OuterParse; |
|
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
78 |
|
|
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
79 |
|
|
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
80 |
|
| 15703 | 81 |
(** datatype src **) |
82 |
||
| 27811 | 83 |
type T = T.token; |
84 |
||
| 15703 | 85 |
datatype src = Src of (string * T list) * Position.T; |
86 |
||
87 |
val src = Src; |
|
88 |
fun dest_src (Src src) = src; |
|
89 |
||
| 21030 | 90 |
fun pretty_src ctxt src = |
91 |
let |
|
| 21697 | 92 |
val prt_thm = Pretty.backquote o ProofContext.pretty_thm ctxt; |
| 27811 | 93 |
fun prt arg = |
94 |
(case T.get_value arg of |
|
95 |
SOME (T.Text s) => Pretty.str (quote s) |
|
96 |
| SOME (T.Typ T) => Syntax.pretty_typ ctxt T |
|
97 |
| SOME (T.Term t) => Syntax.pretty_term ctxt t |
|
98 |
| SOME (T.Fact ths) => Pretty.enclose "(" ")" (Pretty.breaks (map prt_thm ths))
|
|
99 |
| _ => Pretty.str (T.unparse arg)); |
|
| 21030 | 100 |
val (s, args) = #1 (dest_src src); |
101 |
in Pretty.block (Pretty.breaks (Pretty.str s :: map prt args)) end; |
|
102 |
||
| 15703 | 103 |
fun map_name f (Src ((s, args), pos)) = Src ((f s, args), pos); |
104 |
fun map_args f (Src ((s, args), pos)) = Src ((s, map f args), pos); |
|
105 |
||
106 |
||
107 |
(* values *) |
|
108 |
||
| 27811 | 109 |
fun morph_values phi = map_args (T.map_value |
110 |
(fn T.Text s => T.Text s |
|
111 |
| T.Typ T => T.Typ (Morphism.typ phi T) |
|
112 |
| T.Term t => T.Term (Morphism.term phi t) |
|
113 |
| T.Fact ths => T.Fact (Morphism.fact phi ths) |
|
114 |
| T.Attribute att => T.Attribute (Morphism.transform phi att))); |
|
| 15703 | 115 |
|
| 27811 | 116 |
fun maxidx_values (Src ((_, args), _)) = args |> fold (fn arg => |
117 |
(case T.get_value arg of |
|
118 |
SOME (T.Typ T) => Term.maxidx_typ T |
|
119 |
| SOME (T.Term t) => Term.maxidx_term t |
|
120 |
| SOME (T.Fact ths) => fold Thm.maxidx_thm ths |
|
121 |
| _ => I)); |
|
| 15703 | 122 |
|
| 27811 | 123 |
val assignable = map_args T.assignable; |
124 |
val closure = map_args T.closure; |
|
| 15703 | 125 |
|
126 |
||
127 |
||
| 27811 | 128 |
(** argument scanners **) |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
129 |
|
| 27371 | 130 |
(* context *) |
131 |
||
132 |
fun context x = (Scan.state >> Context.proof_of) x; |
|
133 |
fun theory x = (Scan.state >> Context.theory_of) x; |
|
134 |
||
135 |
||
| 27811 | 136 |
(* basic *) |
| 5878 | 137 |
|
| 27811 | 138 |
fun token atom = Scan.ahead P.not_eof --| atom; |
| 5878 | 139 |
|
| 27811 | 140 |
val ident = token |
141 |
(P.short_ident || P.long_ident || P.sym_ident || P.term_var || |
|
142 |
P.type_ident || P.type_var || P.number); |
|
| 5878 | 143 |
|
| 27811 | 144 |
val string = token (P.string || P.verbatim); |
145 |
val alt_string = token P.alt_string; |
|
146 |
val symbolic = token P.keyword_ident_or_symbolic; |
|
147 |
||
148 |
fun $$$ x = (ident >> T.content_of || P.keyword) |
|
149 |
:|-- (fn y => if x = y then Scan.succeed x else Scan.fail); |
|
| 5878 | 150 |
|
151 |
||
| 27811 | 152 |
val named = ident || string; |
| 5878 | 153 |
|
| 10035 | 154 |
val add = $$$ "add"; |
155 |
val del = $$$ "del"; |
|
| 8803 | 156 |
val colon = $$$ ":"; |
| 10035 | 157 |
val query = $$$ "?"; |
158 |
val bang = $$$ "!"; |
|
| 20111 | 159 |
val query_colon = $$$ "?" ^^ $$$ ":"; |
160 |
val bang_colon = $$$ "!" ^^ $$$ ":"; |
|
| 10035 | 161 |
|
| 8803 | 162 |
fun parens scan = $$$ "(" |-- scan --| $$$ ")";
|
| 10150 | 163 |
fun bracks scan = $$$ "[" |-- scan --| $$$ "]"; |
| 15703 | 164 |
fun mode s = Scan.lift (Scan.optional (parens ($$$ s) >> K true) false); |
165 |
fun maybe scan = $$$ "_" >> K NONE || scan >> SOME; |
|
| 5878 | 166 |
|
| 27811 | 167 |
val name = named >> T.content_of; |
168 |
val alt_name = alt_string >> T.content_of; |
|
169 |
val symbol = symbolic >> T.content_of; |
|
| 17064 | 170 |
val liberal_name = symbol || name; |
| 8233 | 171 |
|
| 27811 | 172 |
val var = (ident >> T.content_of) :|-- (fn x => |
173 |
(case Lexicon.read_variable x of SOME v => Scan.succeed v | NONE => Scan.fail)); |
|
| 5878 | 174 |
|
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
175 |
|
| 15703 | 176 |
(* values *) |
177 |
||
| 27811 | 178 |
fun value dest = Scan.some (fn arg => |
179 |
(case T.get_value arg of SOME v => (SOME (dest v) handle Match => NONE) | NONE => NONE)); |
|
| 15703 | 180 |
|
181 |
fun evaluate mk eval arg = |
|
| 27819 | 182 |
let val x = eval arg in (T.assign (SOME (mk x)) arg; x) end; |
| 15703 | 183 |
|
| 27811 | 184 |
val internal_text = value (fn T.Text s => s); |
185 |
val internal_typ = value (fn T.Typ T => T); |
|
186 |
val internal_term = value (fn T.Term t => t); |
|
187 |
val internal_fact = value (fn T.Fact ths => ths); |
|
188 |
val internal_attribute = value (fn T.Attribute att => att); |
|
| 15703 | 189 |
|
| 27819 | 190 |
fun named_text intern = internal_text || named >> evaluate T.Text (intern o T.content_of); |
191 |
fun named_typ readT = internal_typ || named >> evaluate T.Typ (readT o T.source_of); |
|
192 |
fun named_term read = internal_term || named >> evaluate T.Term (read o T.source_of); |
|
193 |
fun named_fact get = internal_fact || named >> evaluate T.Fact (get o T.content_of) || |
|
194 |
alt_string >> evaluate T.Fact (get o T.source_of); |
|
195 |
fun named_attribute att = internal_attribute || named >> evaluate T.Attribute (att o T.content_of); |
|
| 15703 | 196 |
|
197 |
||
| 5878 | 198 |
(* terms and types *) |
199 |
||
| 18635 | 200 |
val typ_abbrev = Scan.peek (named_typ o ProofContext.read_typ_abbrev o Context.proof_of); |
| 25331 | 201 |
val typ = Scan.peek (named_typ o Syntax.read_typ o Context.proof_of); |
|
24508
c8b82fec6447
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
wenzelm
parents:
24244
diff
changeset
|
202 |
val term = Scan.peek (named_term o Syntax.read_term o Context.proof_of); |
| 21724 | 203 |
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
|
204 |
val prop = Scan.peek (named_term o Syntax.read_prop o Context.proof_of); |
| 18635 | 205 |
|
| 5878 | 206 |
|
| 15703 | 207 |
(* type and constant names *) |
208 |
||
|
25323
50d4c8257d06
removed obsolete Sign.read_tyname/const (cf. ProofContext);
wenzelm
parents:
24920
diff
changeset
|
209 |
val tyname = Scan.peek (named_typ o ProofContext.read_tyname o Context.proof_of) |
| 18998 | 210 |
>> (fn Type (c, _) => c | TFree (a, _) => a | _ => ""); |
| 15703 | 211 |
|
|
25323
50d4c8257d06
removed obsolete Sign.read_tyname/const (cf. ProofContext);
wenzelm
parents:
24920
diff
changeset
|
212 |
val const = Scan.peek (named_term o ProofContext.read_const o Context.proof_of) |
| 18998 | 213 |
>> (fn Const (c, _) => c | Free (x, _) => x | _ => ""); |
| 7553 | 214 |
|
| 25343 | 215 |
val const_proper = Scan.peek (named_term o ProofContext.read_const_proper o Context.proof_of) |
216 |
>> (fn Const (c, _) => c | _ => ""); |
|
217 |
||
| 15703 | 218 |
|
| 27811 | 219 |
(* improper method arguments *) |
| 15703 | 220 |
|
| 18998 | 221 |
val bang_facts = Scan.peek (fn context => |
| 27811 | 222 |
P.position ($$$ "!") >> (fn (_, pos) => |
223 |
(warning ("use of prems in proof method" ^ Position.str_of pos);
|
|
224 |
Assumption.prems_of (Context.proof_of context))) || Scan.succeed []); |
|
| 8536 | 225 |
|
226 |
val from_to = |
|
| 27811 | 227 |
P.nat -- ($$$ "-" |-- P.nat) >> (fn (i, j) => fn tac => Seq.INTERVAL tac i j) || |
228 |
P.nat --| $$$ "-" >> (fn i => fn tac => fn st => Seq.INTERVAL tac i (Thm.nprems_of st) st) || |
|
229 |
P.nat >> (fn i => fn tac => tac i) || |
|
| 15703 | 230 |
$$$ "!" >> K ALLGOALS; |
| 8536 | 231 |
|
| 27811 | 232 |
val goal = $$$ "[" |-- P.!!! (from_to --| $$$ "]"); |
| 8536 | 233 |
fun goal_spec def = Scan.lift (Scan.optional goal def); |
| 8233 | 234 |
|
235 |
||
| 27811 | 236 |
(* arguments within outer syntax *) |
| 5878 | 237 |
|
| 27382 | 238 |
fun parse_args is_symid = |
239 |
let |
|
| 27811 | 240 |
val keyword_symid = token (P.keyword_with is_symid); |
241 |
fun atom blk = P.group "argument" |
|
242 |
(ident || keyword_symid || string || alt_string || |
|
243 |
(if blk then token (P.$$$ ",") else Scan.fail)); |
|
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
244 |
|
| 27382 | 245 |
fun args blk x = Scan.optional (args1 blk) [] x |
246 |
and args1 blk x = |
|
247 |
((Scan.repeat1 |
|
248 |
(Scan.repeat1 (atom blk) || |
|
249 |
argsp "(" ")" ||
|
|
250 |
argsp "[" "]")) >> flat) x |
|
| 27811 | 251 |
and argsp l r x = (token (P.$$$ l) ::: P.!!! (args true @@@ (token (P.$$$ r) >> single))) x; |
| 27382 | 252 |
in (args, args1) end; |
| 15703 | 253 |
|
| 27811 | 254 |
val parse = #1 (parse_args T.ident_or_symbolic) false; |
255 |
fun parse1 is_symid = #2 (parse_args is_symid) false; |
|
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
256 |
|
| 27811 | 257 |
|
258 |
(* attributes *) |
|
| 27382 | 259 |
|
| 15703 | 260 |
fun attribs intern = |
261 |
let |
|
| 27819 | 262 |
val attrib_name = internal_text || (symbolic || named) |
263 |
>> evaluate T.Text (intern o T.content_of); |
|
| 27811 | 264 |
val attrib = P.position (attrib_name -- P.!!! parse) >> src; |
265 |
in $$$ "[" |-- P.!!! (P.list attrib --| $$$ "]") end; |
|
| 15703 | 266 |
|
267 |
fun opt_attribs intern = Scan.optional (attribs intern) []; |
|
268 |
||
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
269 |
|
| 27377 | 270 |
(* theorem specifications *) |
271 |
||
272 |
fun thm_name intern s = name -- opt_attribs intern --| $$$ s; |
|
| 27811 | 273 |
|
| 27377 | 274 |
fun opt_thm_name intern s = |
275 |
Scan.optional ((name -- opt_attribs intern || attribs intern >> pair "") --| $$$ s) ("", []);
|
|
276 |
||
277 |
||
| 27382 | 278 |
|
279 |
(** syntax wrapper **) |
|
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
280 |
|
| 8282 | 281 |
fun syntax kind scan (src as Src ((s, args), pos)) st = |
| 27811 | 282 |
(case Scan.error (Scan.finite' T.stopper (Scan.option scan)) (st, args) of |
| 21879 | 283 |
(SOME x, (st', [])) => (x, st') |
| 15703 | 284 |
| (_, (_, args')) => |
285 |
error (kind ^ " " ^ quote s ^ Position.str_of pos ^ ": bad arguments\n " ^ |
|
| 27811 | 286 |
space_implode " " (map T.unparse args'))); |
| 15703 | 287 |
|
| 21879 | 288 |
fun context_syntax kind scan src = apsnd Context.the_proof o syntax kind scan src o Context.Proof; |
| 18998 | 289 |
|
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
290 |
end; |