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