| author | wenzelm | 
| Thu, 14 Nov 2013 17:17:57 +0100 | |
| changeset 54440 | 2c4940d2edf7 | 
| parent 53168 | d998de7f0efc | 
| child 55033 | 8e8243975860 | 
| 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 | Author: Markus Wenzel, TU Muenchen | 
| 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 3 | |
| 35013 | 4 | Parsing with implicit value assignment. Concrete argument syntax of | 
| 27811 | 5 | attributes, methods etc. | 
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 6 | *) | 
| 
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 | signature ARGS = | 
| 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 9 | sig | 
| 15703 | 10 | type src | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 11 | val src: (string * Token.T list) * Position.T -> src | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 12 | val dest_src: src -> (string * Token.T list) * Position.T | 
| 21030 | 13 | val pretty_src: Proof.context -> src -> Pretty.T | 
| 15703 | 14 | val map_name: (string -> string) -> src -> src | 
| 45290 | 15 | val transform_values: morphism -> src -> src | 
| 15703 | 16 | val assignable: src -> src | 
| 17 | val closure: src -> src | |
| 30513 | 18 | val context: Proof.context context_parser | 
| 19 | val theory: theory context_parser | |
| 20 | val $$$ : string -> string parser | |
| 21 | val add: string parser | |
| 22 | val del: string parser | |
| 23 | val colon: string parser | |
| 24 | val query: string parser | |
| 25 | val bang: string parser | |
| 26 | val query_colon: string parser | |
| 27 | val bang_colon: string parser | |
| 53168 | 28 | val parens: 'a parser -> 'a parser | 
| 29 | val bracks: 'a parser -> 'a parser | |
| 30 | val mode: string -> bool parser | |
| 30513 | 31 | val maybe: 'a parser -> 'a option parser | 
| 32 | val name_source: string parser | |
| 30573 | 33 | val name_source_position: (Symbol_Pos.text * Position.T) parser | 
| 30513 | 34 | val name: string parser | 
| 35 | val binding: binding parser | |
| 36 | val alt_name: string parser | |
| 37 | val symbol: string parser | |
| 38 | val liberal_name: string parser | |
| 39 | val var: indexname parser | |
| 40 | val internal_text: string parser | |
| 41 | val internal_typ: typ parser | |
| 42 | val internal_term: term parser | |
| 43 | val internal_fact: thm list parser | |
| 44 | val internal_attribute: (morphism -> attribute) parser | |
| 45 | val named_text: (string -> string) -> string parser | |
| 46 | val named_typ: (string -> typ) -> typ parser | |
| 47 | val named_term: (string -> term) -> term parser | |
| 48 | val named_fact: (string -> thm list) -> thm list parser | |
| 42464 | 49 | val named_attribute: | 
| 50 | (string * Position.T -> morphism -> attribute) -> (morphism -> attribute) parser | |
| 30513 | 51 | val typ_abbrev: typ context_parser | 
| 52 | val typ: typ context_parser | |
| 53 | val term: term context_parser | |
| 54 | val term_abbrev: term context_parser | |
| 55 | val prop: term context_parser | |
| 35360 
df2b2168e43a
clarified ProofContext.read_type_name/Args.type_name wrt strict logical constructors;
 wenzelm parents: 
35013diff
changeset | 56 | val type_name: bool -> string context_parser | 
| 35399 
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
 wenzelm parents: 
35360diff
changeset | 57 | val const: bool -> string context_parser | 
| 
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
 wenzelm parents: 
35360diff
changeset | 58 | val const_proper: bool -> string context_parser | 
| 30514 | 59 | val goal_spec: ((int -> tactic) -> tactic) context_parser | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 60 | val parse: Token.T list parser | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 61 | val parse1: (string -> bool) -> Token.T list parser | 
| 30513 | 62 | val attribs: (string -> string) -> src list parser | 
| 63 | val opt_attribs: (string -> string) -> src list parser | |
| 64 | val thm_name: (string -> string) -> string -> (binding * src list) parser | |
| 65 | val opt_thm_name: (string -> string) -> string -> (binding * src list) parser | |
| 66 | val syntax: string -> 'a context_parser -> src -> Context.generic -> 'a * Context.generic | |
| 67 | val context_syntax: string -> 'a context_parser -> src -> Proof.context -> 'a * Proof.context | |
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 68 | end; | 
| 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 69 | |
| 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 70 | structure Args: ARGS = | 
| 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 71 | struct | 
| 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 72 | |
| 15703 | 73 | (** datatype src **) | 
| 74 | ||
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 75 | datatype src = Src of (string * Token.T list) * Position.T; | 
| 15703 | 76 | |
| 77 | val src = Src; | |
| 78 | fun dest_src (Src src) = src; | |
| 79 | ||
| 21030 | 80 | fun pretty_src ctxt src = | 
| 81 | let | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
30573diff
changeset | 82 | val prt_thm = Pretty.backquote o Display.pretty_thm ctxt; | 
| 27811 | 83 | fun prt arg = | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 84 | (case Token.get_value arg of | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 85 | SOME (Token.Text s) => Pretty.str (quote s) | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 86 | | SOME (Token.Typ T) => Syntax.pretty_typ ctxt T | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 87 | | SOME (Token.Term t) => Syntax.pretty_term ctxt t | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 88 |       | SOME (Token.Fact ths) => Pretty.enclose "(" ")" (Pretty.breaks (map prt_thm ths))
 | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 89 | | _ => Pretty.str (Token.unparse arg)); | 
| 21030 | 90 | val (s, args) = #1 (dest_src src); | 
| 91 | in Pretty.block (Pretty.breaks (Pretty.str s :: map prt args)) end; | |
| 92 | ||
| 15703 | 93 | fun map_name f (Src ((s, args), pos)) = Src ((f s, args), pos); | 
| 94 | fun map_args f (Src ((s, args), pos)) = Src ((s, map f args), pos); | |
| 95 | ||
| 96 | ||
| 97 | (* values *) | |
| 98 | ||
| 45290 | 99 | fun transform_values phi = map_args (Token.map_value | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 100 | (fn Token.Text s => Token.Text s | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 101 | | Token.Typ T => Token.Typ (Morphism.typ phi T) | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 102 | | Token.Term t => Token.Term (Morphism.term phi t) | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 103 | | Token.Fact ths => Token.Fact (Morphism.fact phi ths) | 
| 53112 | 104 | | Token.Attribute att => Token.Attribute (Morphism.transform phi att) | 
| 105 | | tok as Token.Files _ => tok)); | |
| 15703 | 106 | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 107 | val assignable = map_args Token.assignable; | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 108 | val closure = map_args Token.closure; | 
| 15703 | 109 | |
| 110 | ||
| 111 | ||
| 27811 | 112 | (** argument scanners **) | 
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 113 | |
| 27371 | 114 | (* context *) | 
| 115 | ||
| 116 | fun context x = (Scan.state >> Context.proof_of) x; | |
| 117 | fun theory x = (Scan.state >> Context.theory_of) x; | |
| 118 | ||
| 119 | ||
| 27811 | 120 | (* basic *) | 
| 5878 | 121 | |
| 36950 | 122 | fun token atom = Scan.ahead Parse.not_eof --| atom; | 
| 5878 | 123 | |
| 27811 | 124 | val ident = token | 
| 36950 | 125 | (Parse.short_ident || Parse.long_ident || Parse.sym_ident || Parse.term_var || | 
| 126 | Parse.type_ident || Parse.type_var || Parse.number); | |
| 5878 | 127 | |
| 36950 | 128 | val string = token (Parse.string || Parse.verbatim); | 
| 129 | val alt_string = token Parse.alt_string; | |
| 130 | val symbolic = token Parse.keyword_ident_or_symbolic; | |
| 27811 | 131 | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 132 | fun $$$ x = (ident >> Token.content_of || Parse.keyword) | 
| 27811 | 133 | :|-- (fn y => if x = y then Scan.succeed x else Scan.fail); | 
| 5878 | 134 | |
| 135 | ||
| 27811 | 136 | val named = ident || string; | 
| 5878 | 137 | |
| 10035 | 138 | val add = $$$ "add"; | 
| 139 | val del = $$$ "del"; | |
| 8803 | 140 | val colon = $$$ ":"; | 
| 10035 | 141 | val query = $$$ "?"; | 
| 142 | val bang = $$$ "!"; | |
| 20111 | 143 | val query_colon = $$$ "?" ^^ $$$ ":"; | 
| 144 | val bang_colon = $$$ "!" ^^ $$$ ":"; | |
| 10035 | 145 | |
| 8803 | 146 | fun parens scan = $$$ "(" |-- scan --| $$$ ")";
 | 
| 10150 | 147 | fun bracks scan = $$$ "[" |-- scan --| $$$ "]"; | 
| 53168 | 148 | fun mode s = Scan.optional (parens ($$$ s) >> K true) false; | 
| 15703 | 149 | fun maybe scan = $$$ "_" >> K NONE || scan >> SOME; | 
| 5878 | 150 | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 151 | val name_source = named >> Token.source_of; | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 152 | val name_source_position = named >> Token.source_position_of; | 
| 27882 
eaa9fef9f4c1
Args.name_source(_position) for proper position information;
 wenzelm parents: 
27819diff
changeset | 153 | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 154 | val name = named >> Token.content_of; | 
| 36950 | 155 | val binding = Parse.position name >> Binding.make; | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 156 | val alt_name = alt_string >> Token.content_of; | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 157 | val symbol = symbolic >> Token.content_of; | 
| 17064 | 158 | val liberal_name = symbol || name; | 
| 8233 | 159 | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 160 | val var = (ident >> Token.content_of) :|-- (fn x => | 
| 27811 | 161 | (case Lexicon.read_variable x of SOME v => Scan.succeed v | NONE => Scan.fail)); | 
| 5878 | 162 | |
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 163 | |
| 15703 | 164 | (* values *) | 
| 165 | ||
| 27811 | 166 | fun value dest = Scan.some (fn arg => | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 167 | (case Token.get_value arg of SOME v => (SOME (dest v) handle Match => NONE) | NONE => NONE)); | 
| 15703 | 168 | |
| 169 | fun evaluate mk eval arg = | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 170 | let val x = eval arg in (Token.assign (SOME (mk x)) arg; x) end; | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 171 | |
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 172 | val internal_text = value (fn Token.Text s => s); | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 173 | val internal_typ = value (fn Token.Typ T => T); | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 174 | val internal_term = value (fn Token.Term t => t); | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 175 | val internal_fact = value (fn Token.Fact ths => ths); | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 176 | val internal_attribute = value (fn Token.Attribute att => att); | 
| 15703 | 177 | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 178 | fun named_text intern = internal_text || named >> evaluate Token.Text (intern o Token.content_of); | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 179 | fun named_typ readT = internal_typ || named >> evaluate Token.Typ (readT o Token.source_of); | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 180 | fun named_term read = internal_term || named >> evaluate Token.Term (read o Token.source_of); | 
| 42464 | 181 | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 182 | fun named_fact get = internal_fact || named >> evaluate Token.Fact (get o Token.content_of) || | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 183 | alt_string >> evaluate Token.Fact (get o Token.source_of); | 
| 42464 | 184 | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 185 | fun named_attribute att = | 
| 42464 | 186 | internal_attribute || | 
| 187 | named >> evaluate Token.Attribute (fn tok => att (Token.content_of tok, Token.position_of tok)); | |
| 15703 | 188 | |
| 189 | ||
| 5878 | 190 | (* terms and types *) | 
| 191 | ||
| 42360 | 192 | val typ_abbrev = Scan.peek (named_typ o Proof_Context.read_typ_abbrev o Context.proof_of); | 
| 25331 | 193 | 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: 
24244diff
changeset | 194 | val term = Scan.peek (named_term o Syntax.read_term o Context.proof_of); | 
| 42360 | 195 | val term_abbrev = Scan.peek (named_term o Proof_Context.read_term_abbrev o Context.proof_of); | 
| 24508 
c8b82fec6447
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
 wenzelm parents: 
24244diff
changeset | 196 | val prop = Scan.peek (named_term o Syntax.read_prop o Context.proof_of); | 
| 18635 | 197 | |
| 5878 | 198 | |
| 15703 | 199 | (* type and constant names *) | 
| 200 | ||
| 35360 
df2b2168e43a
clarified ProofContext.read_type_name/Args.type_name wrt strict logical constructors;
 wenzelm parents: 
35013diff
changeset | 201 | fun type_name strict = | 
| 42360 | 202 | Scan.peek (fn ctxt => named_typ (Proof_Context.read_type_name (Context.proof_of ctxt) strict)) | 
| 18998 | 203 | >> (fn Type (c, _) => c | TFree (a, _) => a | _ => ""); | 
| 15703 | 204 | |
| 35399 
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
 wenzelm parents: 
35360diff
changeset | 205 | fun const strict = | 
| 42360 | 206 | Scan.peek (fn ctxt => named_term (Proof_Context.read_const (Context.proof_of ctxt) strict dummyT)) | 
| 18998 | 207 | >> (fn Const (c, _) => c | Free (x, _) => x | _ => ""); | 
| 7553 | 208 | |
| 35399 
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
 wenzelm parents: 
35360diff
changeset | 209 | fun const_proper strict = | 
| 42360 | 210 | Scan.peek (fn ctxt => named_term (Proof_Context.read_const_proper (Context.proof_of ctxt) strict)) | 
| 25343 | 211 | >> (fn Const (c, _) => c | _ => ""); | 
| 212 | ||
| 15703 | 213 | |
| 27811 | 214 | (* improper method arguments *) | 
| 15703 | 215 | |
| 8536 | 216 | val from_to = | 
| 36950 | 217 | Parse.nat -- ($$$ "-" |-- Parse.nat) >> (fn (i, j) => fn tac => Seq.INTERVAL tac i j) || | 
| 218 | Parse.nat --| $$$ "-" >> (fn i => fn tac => fn st => Seq.INTERVAL tac i (Thm.nprems_of st) st) || | |
| 219 | Parse.nat >> (fn i => fn tac => tac i) || | |
| 15703 | 220 | $$$ "!" >> K ALLGOALS; | 
| 8536 | 221 | |
| 36950 | 222 | val goal = $$$ "[" |-- Parse.!!! (from_to --| $$$ "]"); | 
| 30514 | 223 | fun goal_spec x = Scan.lift (Scan.optional goal (fn tac => tac 1)) x; | 
| 8233 | 224 | |
| 225 | ||
| 27811 | 226 | (* arguments within outer syntax *) | 
| 5878 | 227 | |
| 53113 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 228 | val argument_kinds = | 
| 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 229 | [Token.Ident, Token.LongIdent, Token.SymIdent, Token.Var, Token.TypeIdent, Token.TypeVar, | 
| 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 230 | Token.Nat, Token.Float, Token.String, Token.AltString, Token.Verbatim]; | 
| 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 231 | |
| 27382 | 232 | fun parse_args is_symid = | 
| 233 | let | |
| 53113 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 234 | fun argument blk = | 
| 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 235 | Parse.group (fn () => "argument") | 
| 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 236 | (Scan.one (fn tok => | 
| 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 237 | let val kind = Token.kind_of tok in | 
| 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 238 | member (op =) argument_kinds kind orelse | 
| 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 239 | Token.keyword_with is_symid tok orelse | 
| 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 240 | (blk andalso Token.keyword_with (fn s => s = ",") tok) | 
| 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 241 | end)); | 
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 242 | |
| 27382 | 243 | fun args blk x = Scan.optional (args1 blk) [] x | 
| 244 | and args1 blk x = | |
| 245 | ((Scan.repeat1 | |
| 53113 
d9ba3417cb41
more careful parsing of nested argument tokens -- avoid Parse.RESET_VALUE on still uninterpreted material;
 wenzelm parents: 
53112diff
changeset | 246 | (Scan.repeat1 (argument blk) || | 
| 27382 | 247 |           argsp "(" ")" ||
 | 
| 248 | argsp "[" "]")) >> flat) x | |
| 36950 | 249 | and argsp l r x = | 
| 250 | (token (Parse.$$$ l) ::: Parse.!!! (args true @@@ (token (Parse.$$$ r) >> single))) x; | |
| 27382 | 251 | in (args, args1) end; | 
| 15703 | 252 | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 253 | val parse = #1 (parse_args Token.ident_or_symbolic) false; | 
| 27811 | 254 | fun parse1 is_symid = #2 (parse_args is_symid) false; | 
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 255 | |
| 27811 | 256 | |
| 257 | (* attributes *) | |
| 27382 | 258 | |
| 15703 | 259 | fun attribs intern = | 
| 260 | let | |
| 27819 | 261 | val attrib_name = internal_text || (symbolic || named) | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 262 | >> evaluate Token.Text (intern o Token.content_of); | 
| 36950 | 263 | val attrib = Parse.position (attrib_name -- Parse.!!! parse) >> src; | 
| 264 | in $$$ "[" |-- Parse.!!! (Parse.list attrib --| $$$ "]") end; | |
| 15703 | 265 | |
| 266 | fun opt_attribs intern = Scan.optional (attribs intern) []; | |
| 267 | ||
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 268 | |
| 27377 | 269 | (* theorem specifications *) | 
| 270 | ||
| 28078 | 271 | fun thm_name intern s = binding -- opt_attribs intern --| $$$ s; | 
| 27811 | 272 | |
| 27377 | 273 | fun opt_thm_name intern s = | 
| 28078 | 274 | Scan.optional | 
| 28965 | 275 | ((binding -- opt_attribs intern || attribs intern >> pair Binding.empty) --| $$$ s) | 
| 276 | (Binding.empty, []); | |
| 27377 | 277 | |
| 278 | ||
| 27382 | 279 | |
| 280 | (** syntax wrapper **) | |
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 281 | |
| 32784 | 282 | fun syntax kind scan (Src ((s, args), pos)) st = | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 283 | (case Scan.error (Scan.finite' Token.stopper (Scan.option scan)) (st, args) of | 
| 21879 | 284 | (SOME x, (st', [])) => (x, st') | 
| 15703 | 285 | | (_, (_, args')) => | 
| 48992 | 286 | error (kind ^ " " ^ quote s ^ Position.here pos ^ ": bad arguments\n " ^ | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 287 | space_implode " " (map Token.unparse args'))); | 
| 15703 | 288 | |
| 21879 | 289 | fun context_syntax kind scan src = apsnd Context.the_proof o syntax kind scan src o Context.Proof; | 
| 18998 | 290 | |
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 291 | end; |