| author | wenzelm | 
| Tue, 28 Oct 2014 09:57:12 +0100 | |
| changeset 58797 | 6d71f19a9fd6 | 
| parent 58017 | 5bdb58845eab | 
| child 59064 | a8bcb5a446c8 | 
| 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 | |
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57944diff
changeset | 4 | Quasi-inner syntax based on outer tokens: 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 | 
| 30513 | 10 | val context: Proof.context context_parser | 
| 11 | val theory: theory context_parser | |
| 12 | val $$$ : string -> string parser | |
| 13 | val add: string parser | |
| 14 | val del: string parser | |
| 15 | val colon: string parser | |
| 16 | val query: string parser | |
| 17 | val bang: string parser | |
| 18 | val query_colon: string parser | |
| 19 | val bang_colon: string parser | |
| 53168 | 20 | val parens: 'a parser -> 'a parser | 
| 21 | val bracks: 'a parser -> 'a parser | |
| 22 | val mode: string -> bool parser | |
| 30513 | 23 | val maybe: 'a parser -> 'a option parser | 
| 55111 | 24 | val cartouche_inner_syntax: string parser | 
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55745diff
changeset | 25 | val cartouche_source_position: Symbol_Pos.source parser | 
| 56500 
90f17a04567d
proper Args.name vs. Args.text as documented (in contrast to adhoc union in 75aaee32893d, which had to cope with more limited Args.T);
 wenzelm parents: 
56499diff
changeset | 26 | val text_source_position: Symbol_Pos.source parser | 
| 
90f17a04567d
proper Args.name vs. Args.text as documented (in contrast to adhoc union in 75aaee32893d, which had to cope with more limited Args.T);
 wenzelm parents: 
56499diff
changeset | 27 | val text: string parser | 
| 55111 | 28 | val name_inner_syntax: string parser | 
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55745diff
changeset | 29 | val name_source_position: Symbol_Pos.source parser | 
| 30513 | 30 | val name: string parser | 
| 31 | val binding: binding parser | |
| 32 | val alt_name: string parser | |
| 33 | val symbol: string parser | |
| 34 | val liberal_name: string parser | |
| 35 | val var: indexname parser | |
| 58012 | 36 | val internal_source: Token.src parser | 
| 57944 
fff8d328da56
more informative Token.Name with history of morphisms;
 wenzelm parents: 
57942diff
changeset | 37 | val internal_name: (string * morphism) parser | 
| 30513 | 38 | val internal_typ: typ parser | 
| 39 | val internal_term: term parser | |
| 40 | val internal_fact: thm list parser | |
| 41 | val internal_attribute: (morphism -> attribute) parser | |
| 58017 | 42 | val internal_declaration: declaration parser | 
| 58012 | 43 | val named_source: (Token.T -> Token.src) -> Token.src parser | 
| 30513 | 44 | val named_typ: (string -> typ) -> typ parser | 
| 45 | val named_term: (string -> term) -> term parser | |
| 57942 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 wenzelm parents: 
56500diff
changeset | 46 | val named_fact: (string -> string option * thm list) -> thm list parser | 
| 58017 | 47 | val named_attribute: (string * Position.T -> morphism -> attribute) -> | 
| 48 | (morphism -> attribute) parser | |
| 49 | val text_declaration: (Symbol_Pos.source -> declaration) -> declaration parser | |
| 30513 | 50 | val typ_abbrev: typ context_parser | 
| 51 | val typ: typ context_parser | |
| 52 | val term: term context_parser | |
| 55155 | 53 | val term_pattern: term context_parser | 
| 30513 | 54 | val term_abbrev: term context_parser | 
| 55 | val prop: term context_parser | |
| 55951 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 wenzelm parents: 
55915diff
changeset | 56 |   val type_name: {proper: bool, strict: bool} -> string context_parser
 | 
| 55954 | 57 |   val const: {proper: bool, strict: bool} -> string context_parser
 | 
| 30514 | 58 | val goal_spec: ((int -> tactic) -> tactic) context_parser | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57944diff
changeset | 59 | val attribs: (xstring * Position.T -> string) -> Token.src list parser | 
| 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57944diff
changeset | 60 | val opt_attribs: (xstring * Position.T -> string) -> Token.src list parser | 
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 61 | end; | 
| 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 62 | |
| 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 63 | structure Args: ARGS = | 
| 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 64 | struct | 
| 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 65 | |
| 27811 | 66 | (** argument scanners **) | 
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 67 | |
| 27371 | 68 | (* context *) | 
| 69 | ||
| 70 | fun context x = (Scan.state >> Context.proof_of) x; | |
| 71 | fun theory x = (Scan.state >> Context.theory_of) x; | |
| 72 | ||
| 73 | ||
| 27811 | 74 | (* basic *) | 
| 5878 | 75 | |
| 56201 | 76 | val ident = Parse.token | 
| 36950 | 77 | (Parse.short_ident || Parse.long_ident || Parse.sym_ident || Parse.term_var || | 
| 78 | Parse.type_ident || Parse.type_var || Parse.number); | |
| 5878 | 79 | |
| 56500 
90f17a04567d
proper Args.name vs. Args.text as documented (in contrast to adhoc union in 75aaee32893d, which had to cope with more limited Args.T);
 wenzelm parents: 
56499diff
changeset | 80 | val string = Parse.token Parse.string; | 
| 56499 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56333diff
changeset | 81 | val alt_string = Parse.token (Parse.alt_string || Parse.cartouche); | 
| 56201 | 82 | val symbolic = Parse.token (Parse.keyword_with Token.ident_or_symbolic); | 
| 27811 | 83 | |
| 55914 
c5b752d549e3
clarified init_assignable: make double-sure that initial values are reset;
 wenzelm parents: 
55828diff
changeset | 84 | fun $$$ x = | 
| 56201 | 85 | (ident || Parse.token Parse.keyword) :|-- (fn tok => | 
| 56063 | 86 | let val y = Token.content_of tok in | 
| 87 | if x = y | |
| 56202 | 88 | then (Token.assign (SOME (Token.Literal (false, Markup.quasi_keyword))) tok; Scan.succeed x) | 
| 56063 | 89 | else Scan.fail | 
| 90 | end); | |
| 5878 | 91 | |
| 27811 | 92 | val named = ident || string; | 
| 5878 | 93 | |
| 10035 | 94 | val add = $$$ "add"; | 
| 95 | val del = $$$ "del"; | |
| 8803 | 96 | val colon = $$$ ":"; | 
| 10035 | 97 | val query = $$$ "?"; | 
| 98 | val bang = $$$ "!"; | |
| 20111 | 99 | val query_colon = $$$ "?" ^^ $$$ ":"; | 
| 100 | val bang_colon = $$$ "!" ^^ $$$ ":"; | |
| 10035 | 101 | |
| 8803 | 102 | fun parens scan = $$$ "(" |-- scan --| $$$ ")";
 | 
| 10150 | 103 | fun bracks scan = $$$ "[" |-- scan --| $$$ "]"; | 
| 53168 | 104 | fun mode s = Scan.optional (parens ($$$ s) >> K true) false; | 
| 15703 | 105 | fun maybe scan = $$$ "_" >> K NONE || scan >> SOME; | 
| 5878 | 106 | |
| 56201 | 107 | val cartouche = Parse.token Parse.cartouche; | 
| 55111 | 108 | val cartouche_inner_syntax = cartouche >> Token.inner_syntax_of; | 
| 55045 
99056d23e05b
cartouche within nested args (attributes, methods, etc.);
 wenzelm parents: 
55033diff
changeset | 109 | val cartouche_source_position = cartouche >> Token.source_position_of; | 
| 
99056d23e05b
cartouche within nested args (attributes, methods, etc.);
 wenzelm parents: 
55033diff
changeset | 110 | |
| 56500 
90f17a04567d
proper Args.name vs. Args.text as documented (in contrast to adhoc union in 75aaee32893d, which had to cope with more limited Args.T);
 wenzelm parents: 
56499diff
changeset | 111 | val text_token = named || Parse.token (Parse.verbatim || Parse.cartouche); | 
| 
90f17a04567d
proper Args.name vs. Args.text as documented (in contrast to adhoc union in 75aaee32893d, which had to cope with more limited Args.T);
 wenzelm parents: 
56499diff
changeset | 112 | val text_source_position = text_token >> Token.source_position_of; | 
| 
90f17a04567d
proper Args.name vs. Args.text as documented (in contrast to adhoc union in 75aaee32893d, which had to cope with more limited Args.T);
 wenzelm parents: 
56499diff
changeset | 113 | val text = text_token >> Token.content_of; | 
| 
90f17a04567d
proper Args.name vs. Args.text as documented (in contrast to adhoc union in 75aaee32893d, which had to cope with more limited Args.T);
 wenzelm parents: 
56499diff
changeset | 114 | |
| 55111 | 115 | val name_inner_syntax = named >> Token.inner_syntax_of; | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 116 | val name_source_position = named >> Token.source_position_of; | 
| 27882 
eaa9fef9f4c1
Args.name_source(_position) for proper position information;
 wenzelm parents: 
27819diff
changeset | 117 | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 118 | val name = named >> Token.content_of; | 
| 36950 | 119 | 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 | 120 | 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 | 121 | val symbol = symbolic >> Token.content_of; | 
| 17064 | 122 | val liberal_name = symbol || name; | 
| 8233 | 123 | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 124 | val var = (ident >> Token.content_of) :|-- (fn x => | 
| 27811 | 125 | (case Lexicon.read_variable x of SOME v => Scan.succeed v | NONE => Scan.fail)); | 
| 5878 | 126 | |
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 127 | |
| 15703 | 128 | (* values *) | 
| 129 | ||
| 27811 | 130 | 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 | 131 | (case Token.get_value arg of SOME v => (SOME (dest v) handle Match => NONE) | NONE => NONE)); | 
| 15703 | 132 | |
| 133 | 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 | 134 | 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 | 135 | |
| 58012 | 136 | val internal_source = value (fn Token.Source src => src); | 
| 57944 
fff8d328da56
more informative Token.Name with history of morphisms;
 wenzelm parents: 
57942diff
changeset | 137 | val internal_name = value (fn Token.Name a => a); | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 138 | 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 | 139 | val internal_term = value (fn Token.Term t => t); | 
| 57942 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 wenzelm parents: 
56500diff
changeset | 140 | val internal_fact = value (fn Token.Fact (_, ths) => ths); | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 141 | val internal_attribute = value (fn Token.Attribute att => att); | 
| 58017 | 142 | val internal_declaration = value (fn Token.Declaration decl => decl); | 
| 15703 | 143 | |
| 58012 | 144 | fun named_source read = internal_source || named >> evaluate Token.Source read; | 
| 145 | ||
| 146 | fun named_typ read = internal_typ || named >> evaluate Token.Typ (read o Token.inner_syntax_of); | |
| 55111 | 147 | fun named_term read = internal_term || named >> evaluate Token.Term (read o Token.inner_syntax_of); | 
| 42464 | 148 | |
| 57942 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 wenzelm parents: 
56500diff
changeset | 149 | fun named_fact get = | 
| 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 wenzelm parents: 
56500diff
changeset | 150 | internal_fact || | 
| 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 wenzelm parents: 
56500diff
changeset | 151 | named >> evaluate Token.Fact (get o Token.content_of) >> #2 || | 
| 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 wenzelm parents: 
56500diff
changeset | 152 | alt_string >> evaluate Token.Fact (get o Token.inner_syntax_of) >> #2; | 
| 42464 | 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 | fun named_attribute att = | 
| 42464 | 155 | internal_attribute || | 
| 55708 | 156 | named >> evaluate Token.Attribute (fn tok => att (Token.content_of tok, Token.pos_of tok)); | 
| 15703 | 157 | |
| 58017 | 158 | fun text_declaration read = | 
| 159 | internal_declaration || | |
| 160 | text_token >> evaluate Token.Declaration (read o Token.source_position_of); | |
| 161 | ||
| 15703 | 162 | |
| 5878 | 163 | (* terms and types *) | 
| 164 | ||
| 42360 | 165 | val typ_abbrev = Scan.peek (named_typ o Proof_Context.read_typ_abbrev o Context.proof_of); | 
| 25331 | 166 | 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 | 167 | val term = Scan.peek (named_term o Syntax.read_term o Context.proof_of); | 
| 55155 | 168 | val term_pattern = Scan.peek (named_term o Proof_Context.read_term_pattern o Context.proof_of); | 
| 42360 | 169 | 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 | 170 | val prop = Scan.peek (named_term o Syntax.read_prop o Context.proof_of); | 
| 18635 | 171 | |
| 5878 | 172 | |
| 15703 | 173 | (* type and constant names *) | 
| 174 | ||
| 55951 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 wenzelm parents: 
55915diff
changeset | 175 | fun type_name flags = | 
| 56002 | 176 | Scan.peek (named_typ o Proof_Context.read_type_name flags o Context.proof_of) | 
| 18998 | 177 | >> (fn Type (c, _) => c | TFree (a, _) => a | _ => ""); | 
| 15703 | 178 | |
| 55954 | 179 | fun const flags = | 
| 56002 | 180 | Scan.peek (named_term o Proof_Context.read_const flags o Context.proof_of) | 
| 18998 | 181 | >> (fn Const (c, _) => c | Free (x, _) => x | _ => ""); | 
| 7553 | 182 | |
| 15703 | 183 | |
| 27811 | 184 | (* improper method arguments *) | 
| 15703 | 185 | |
| 8536 | 186 | val from_to = | 
| 36950 | 187 | Parse.nat -- ($$$ "-" |-- Parse.nat) >> (fn (i, j) => fn tac => Seq.INTERVAL tac i j) || | 
| 188 | Parse.nat --| $$$ "-" >> (fn i => fn tac => fn st => Seq.INTERVAL tac i (Thm.nprems_of st) st) || | |
| 189 | Parse.nat >> (fn i => fn tac => tac i) || | |
| 15703 | 190 | $$$ "!" >> K ALLGOALS; | 
| 8536 | 191 | |
| 56202 | 192 | val goal = Parse.keyword_improper "[" |-- Parse.!!! (from_to --| Parse.keyword_improper "]"); | 
| 30514 | 193 | fun goal_spec x = Scan.lift (Scan.optional goal (fn tac => tac 1)) x; | 
| 8233 | 194 | |
| 195 | ||
| 27811 | 196 | (* attributes *) | 
| 27382 | 197 | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55956diff
changeset | 198 | fun attribs check = | 
| 15703 | 199 | let | 
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55956diff
changeset | 200 | fun intern tok = check (Token.content_of tok, Token.pos_of tok); | 
| 57944 
fff8d328da56
more informative Token.Name with history of morphisms;
 wenzelm parents: 
57942diff
changeset | 201 | val attrib_name = internal_name >> #1 || (symbolic || named) >> evaluate Token.name0 intern; | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57944diff
changeset | 202 | val attrib = Parse.position attrib_name -- Parse.!!! Parse.args >> uncurry Token.src; | 
| 36950 | 203 | in $$$ "[" |-- Parse.!!! (Parse.list attrib --| $$$ "]") end; | 
| 15703 | 204 | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55956diff
changeset | 205 | fun opt_attribs check = Scan.optional (attribs check) []; | 
| 15703 | 206 | |
| 5822 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 wenzelm parents: diff
changeset | 207 | end; |