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