author | wenzelm |
Sat, 15 Apr 2023 14:14:30 +0200 | |
changeset 77855 | ff801186ff66 |
parent 74887 | 56247fdb8bbb |
child 78072 | 001739cb8d08 |
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:
57944
diff
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 |
|
61814
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
60577
diff
changeset
|
12 |
val symbolic: Token.T parser |
30513 | 13 |
val $$$ : string -> string parser |
14 |
val add: string parser |
|
15 |
val del: string parser |
|
16 |
val colon: string parser |
|
17 |
val query: string parser |
|
18 |
val bang: string parser |
|
19 |
val query_colon: string parser |
|
20 |
val bang_colon: string parser |
|
53168 | 21 |
val parens: 'a parser -> 'a parser |
22 |
val bracks: 'a parser -> 'a parser |
|
23 |
val mode: string -> bool parser |
|
30513 | 24 |
val maybe: 'a parser -> 'a option parser |
60577 | 25 |
val name_token: Token.T parser |
26 |
val name: string parser |
|
69349
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
wenzelm
parents:
63120
diff
changeset
|
27 |
val name_position: (string * Position.T) parser |
55111 | 28 |
val cartouche_inner_syntax: string parser |
59809 | 29 |
val cartouche_input: Input.source parser |
30513 | 30 |
val binding: binding parser |
31 |
val alt_name: string parser |
|
32 |
val liberal_name: string parser |
|
33 |
val var: indexname parser |
|
58012 | 34 |
val internal_source: Token.src parser |
61814
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
60577
diff
changeset
|
35 |
val internal_name: Token.name_value parser |
30513 | 36 |
val internal_typ: typ parser |
37 |
val internal_term: term parser |
|
38 |
val internal_fact: thm list parser |
|
39 |
val internal_attribute: (morphism -> attribute) parser |
|
58017 | 40 |
val internal_declaration: declaration parser |
58012 | 41 |
val named_source: (Token.T -> Token.src) -> Token.src parser |
30513 | 42 |
val named_typ: (string -> typ) -> typ parser |
43 |
val named_term: (string -> term) -> term parser |
|
57942
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
wenzelm
parents:
56500
diff
changeset
|
44 |
val named_fact: (string -> string option * thm list) -> thm list parser |
58017 | 45 |
val named_attribute: (string * Position.T -> morphism -> attribute) -> |
46 |
(morphism -> attribute) parser |
|
74887 | 47 |
val embedded_declaration: (Input.source -> declaration) -> declaration parser |
30513 | 48 |
val typ_abbrev: typ context_parser |
49 |
val typ: typ context_parser |
|
50 |
val term: term context_parser |
|
55155 | 51 |
val term_pattern: term context_parser |
30513 | 52 |
val term_abbrev: term context_parser |
53 |
val prop: term context_parser |
|
55951
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents:
55915
diff
changeset
|
54 |
val type_name: {proper: bool, strict: bool} -> string context_parser |
55954 | 55 |
val const: {proper: bool, strict: bool} -> string context_parser |
30514 | 56 |
val goal_spec: ((int -> tactic) -> tactic) context_parser |
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
57 |
end; |
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
58 |
|
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
59 |
structure Args: ARGS = |
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
60 |
struct |
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
61 |
|
27811 | 62 |
(** argument scanners **) |
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
63 |
|
27371 | 64 |
(* context *) |
65 |
||
66 |
fun context x = (Scan.state >> Context.proof_of) x; |
|
67 |
fun theory x = (Scan.state >> Context.theory_of) x; |
|
68 |
||
69 |
||
27811 | 70 |
(* basic *) |
5878 | 71 |
|
56201 | 72 |
val ident = Parse.token |
36950 | 73 |
(Parse.short_ident || Parse.long_ident || Parse.sym_ident || Parse.term_var || |
74 |
Parse.type_ident || Parse.type_var || Parse.number); |
|
5878 | 75 |
|
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:
56499
diff
changeset
|
76 |
val string = Parse.token Parse.string; |
56499
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents:
56333
diff
changeset
|
77 |
val alt_string = Parse.token (Parse.alt_string || Parse.cartouche); |
56201 | 78 |
val symbolic = Parse.token (Parse.keyword_with Token.ident_or_symbolic); |
27811 | 79 |
|
55914
c5b752d549e3
clarified init_assignable: make double-sure that initial values are reset;
wenzelm
parents:
55828
diff
changeset
|
80 |
fun $$$ x = |
56201 | 81 |
(ident || Parse.token Parse.keyword) :|-- (fn tok => |
56063 | 82 |
let val y = Token.content_of tok in |
83 |
if x = y |
|
56202 | 84 |
then (Token.assign (SOME (Token.Literal (false, Markup.quasi_keyword))) tok; Scan.succeed x) |
56063 | 85 |
else Scan.fail |
86 |
end); |
|
5878 | 87 |
|
10035 | 88 |
val add = $$$ "add"; |
89 |
val del = $$$ "del"; |
|
8803 | 90 |
val colon = $$$ ":"; |
10035 | 91 |
val query = $$$ "?"; |
92 |
val bang = $$$ "!"; |
|
20111 | 93 |
val query_colon = $$$ "?" ^^ $$$ ":"; |
94 |
val bang_colon = $$$ "!" ^^ $$$ ":"; |
|
10035 | 95 |
|
8803 | 96 |
fun parens scan = $$$ "(" |-- scan --| $$$ ")"; |
10150 | 97 |
fun bracks scan = $$$ "[" |-- scan --| $$$ "]"; |
53168 | 98 |
fun mode s = Scan.optional (parens ($$$ s) >> K true) false; |
15703 | 99 |
fun maybe scan = $$$ "_" >> K NONE || scan >> SOME; |
5878 | 100 |
|
60577 | 101 |
val name_token = ident || string; |
102 |
val name = name_token >> Token.content_of; |
|
69349
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
wenzelm
parents:
63120
diff
changeset
|
103 |
val name_position = name_token >> (Input.source_content o Token.input_of); |
60577 | 104 |
|
56201 | 105 |
val cartouche = Parse.token Parse.cartouche; |
55111 | 106 |
val cartouche_inner_syntax = cartouche >> Token.inner_syntax_of; |
59809 | 107 |
val cartouche_input = cartouche >> Token.input_of; |
55045
99056d23e05b
cartouche within nested args (attributes, methods, etc.);
wenzelm
parents:
55033
diff
changeset
|
108 |
|
69349
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
wenzelm
parents:
63120
diff
changeset
|
109 |
val binding = Parse.input name >> (Binding.make o Input.source_content); |
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
110 |
val alt_name = alt_string >> Token.content_of; |
61814
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
60577
diff
changeset
|
111 |
val liberal_name = (symbolic >> Token.content_of) || name; |
8233 | 112 |
|
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
61814
diff
changeset
|
113 |
val var = |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
61814
diff
changeset
|
114 |
(ident >> Token.content_of) :|-- (fn x => |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
61814
diff
changeset
|
115 |
(case Lexicon.read_variable x of SOME v => Scan.succeed v | NONE => Scan.fail)); |
5878 | 116 |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
117 |
|
15703 | 118 |
(* values *) |
119 |
||
27811 | 120 |
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:
36950
diff
changeset
|
121 |
(case Token.get_value arg of SOME v => (SOME (dest v) handle Match => NONE) | NONE => NONE)); |
15703 | 122 |
|
58012 | 123 |
val internal_source = value (fn Token.Source src => src); |
61814
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
60577
diff
changeset
|
124 |
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:
36950
diff
changeset
|
125 |
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:
36950
diff
changeset
|
126 |
val internal_term = value (fn Token.Term t => t); |
57942
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
wenzelm
parents:
56500
diff
changeset
|
127 |
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:
36950
diff
changeset
|
128 |
val internal_attribute = value (fn Token.Attribute att => att); |
58017 | 129 |
val internal_declaration = value (fn Token.Declaration decl => decl); |
15703 | 130 |
|
61814
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
60577
diff
changeset
|
131 |
fun named_source read = |
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
60577
diff
changeset
|
132 |
internal_source || name_token >> Token.evaluate Token.Source read; |
58012 | 133 |
|
60577 | 134 |
fun named_typ read = |
74563 | 135 |
internal_typ || |
136 |
Parse.token Parse.embedded >> Token.evaluate Token.Typ (read o Token.inner_syntax_of); |
|
60577 | 137 |
|
138 |
fun named_term read = |
|
74563 | 139 |
internal_term || |
140 |
Parse.token Parse.embedded >> Token.evaluate Token.Term (read o Token.inner_syntax_of); |
|
42464 | 141 |
|
57942
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
wenzelm
parents:
56500
diff
changeset
|
142 |
fun named_fact get = |
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
wenzelm
parents:
56500
diff
changeset
|
143 |
internal_fact || |
61814
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
60577
diff
changeset
|
144 |
name_token >> Token.evaluate Token.Fact (get o Token.content_of) >> #2 || |
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
60577
diff
changeset
|
145 |
alt_string >> Token.evaluate Token.Fact (get o Token.inner_syntax_of) >> #2; |
42464 | 146 |
|
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
147 |
fun named_attribute att = |
42464 | 148 |
internal_attribute || |
61814
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
60577
diff
changeset
|
149 |
name_token >> |
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
60577
diff
changeset
|
150 |
Token.evaluate Token.Attribute (fn tok => att (Token.content_of tok, Token.pos_of tok)); |
15703 | 151 |
|
74887 | 152 |
fun embedded_declaration read = |
153 |
internal_declaration || |
|
154 |
Parse.token Parse.embedded >> Token.evaluate Token.Declaration (read o Token.input_of); |
|
59649 | 155 |
|
15703 | 156 |
|
5878 | 157 |
(* terms and types *) |
158 |
||
42360 | 159 |
val typ_abbrev = Scan.peek (named_typ o Proof_Context.read_typ_abbrev o Context.proof_of); |
25331 | 160 |
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
|
161 |
val term = Scan.peek (named_term o Syntax.read_term o Context.proof_of); |
55155 | 162 |
val term_pattern = Scan.peek (named_term o Proof_Context.read_term_pattern o Context.proof_of); |
42360 | 163 |
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:
24244
diff
changeset
|
164 |
val prop = Scan.peek (named_term o Syntax.read_prop o Context.proof_of); |
18635 | 165 |
|
5878 | 166 |
|
15703 | 167 |
(* type and constant names *) |
168 |
||
55951
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents:
55915
diff
changeset
|
169 |
fun type_name flags = |
56002 | 170 |
Scan.peek (named_typ o Proof_Context.read_type_name flags o Context.proof_of) |
18998 | 171 |
>> (fn Type (c, _) => c | TFree (a, _) => a | _ => ""); |
15703 | 172 |
|
55954 | 173 |
fun const flags = |
56002 | 174 |
Scan.peek (named_term o Proof_Context.read_const flags o Context.proof_of) |
18998 | 175 |
>> (fn Const (c, _) => c | Free (x, _) => x | _ => ""); |
7553 | 176 |
|
15703 | 177 |
|
27811 | 178 |
(* improper method arguments *) |
15703 | 179 |
|
8536 | 180 |
val from_to = |
36950 | 181 |
Parse.nat -- ($$$ "-" |-- Parse.nat) >> (fn (i, j) => fn tac => Seq.INTERVAL tac i j) || |
182 |
Parse.nat --| $$$ "-" >> (fn i => fn tac => fn st => Seq.INTERVAL tac i (Thm.nprems_of st) st) || |
|
183 |
Parse.nat >> (fn i => fn tac => tac i) || |
|
15703 | 184 |
$$$ "!" >> K ALLGOALS; |
8536 | 185 |
|
56202 | 186 |
val goal = Parse.keyword_improper "[" |-- Parse.!!! (from_to --| Parse.keyword_improper "]"); |
30514 | 187 |
fun goal_spec x = Scan.lift (Scan.optional goal (fn tac => tac 1)) x; |
8233 | 188 |
|
5822
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
wenzelm
parents:
diff
changeset
|
189 |
end; |