author | wenzelm |
Wed, 10 Jan 2024 13:37:29 +0100 | |
changeset 79469 | deb50d396ff7 |
parent 79463 | 7d10708bbc32 |
permissions | -rw-r--r-- |
74559 | 1 |
(* Title: Pure/ML/ml_antiquotations.ML |
56205 | 2 |
Author: Makarius |
3 |
||
74559 | 4 |
Miscellaneous ML antiquotations. |
56205 | 5 |
*) |
6 |
||
74559 | 7 |
signature ML_ANTIQUOTATIONS = |
74341 | 8 |
sig |
9 |
val make_judgment: Proof.context -> term -> term |
|
10 |
val dest_judgment: Proof.context -> term -> term |
|
11 |
end; |
|
12 |
||
74559 | 13 |
structure ML_Antiquotations: ML_ANTIQUOTATIONS = |
56205 | 14 |
struct |
15 |
||
58632 | 16 |
(* ML support *) |
17 |
||
18 |
val _ = Theory.setup |
|
67147 | 19 |
(ML_Antiquotation.inline \<^binding>\<open>undefined\<close> |
61597
53e32a9b66b8
added @{undefined} with somewhat undefined symbol;
wenzelm
parents:
61596
diff
changeset
|
20 |
(Scan.succeed "(raise General.Match)") #> |
53e32a9b66b8
added @{undefined} with somewhat undefined symbol;
wenzelm
parents:
61596
diff
changeset
|
21 |
|
67147 | 22 |
ML_Antiquotation.inline \<^binding>\<open>assert\<close> |
58632 | 23 |
(Scan.succeed "(fn b => if b then () else raise General.Fail \"Assertion failed\")") #> |
24 |
||
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
25 |
ML_Antiquotation.declaration_embedded \<^binding>\<open>print\<close> |
74563 | 26 |
(Scan.lift (Scan.optional Parse.embedded "Output.writeln")) |
58632 | 27 |
(fn src => fn output => fn ctxt => |
28 |
let |
|
59127
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents:
59112
diff
changeset
|
29 |
val struct_name = ML_Context.struct_name ctxt; |
58632 | 30 |
val (_, pos) = Token.name_of_src src; |
59112 | 31 |
val (a, ctxt') = ML_Context.variant "output" ctxt; |
58632 | 32 |
val env = |
33 |
"val " ^ a ^ ": string -> unit =\n\ |
|
34 |
\ (" ^ output ^ ") o (fn s => s ^ Position.here (" ^ |
|
35 |
ML_Syntax.print_position pos ^ "));\n"; |
|
36 |
val body = |
|
62900
c641bf9402fd
simplified default print_depth: context is usually available, in contrast to 0d295e339f52;
wenzelm
parents:
62899
diff
changeset
|
37 |
"(fn x => (" ^ struct_name ^ "." ^ a ^ " (" ^ ML_Pretty.make_string_fn ^ " x); x))"; |
63204
921a5be54132
support rat numerals via special antiquotation syntax;
wenzelm
parents:
63120
diff
changeset
|
38 |
in (K (env, body), ctxt') end) #> |
921a5be54132
support rat numerals via special antiquotation syntax;
wenzelm
parents:
63120
diff
changeset
|
39 |
|
67147 | 40 |
ML_Antiquotation.value \<^binding>\<open>rat\<close> |
63204
921a5be54132
support rat numerals via special antiquotation syntax;
wenzelm
parents:
63120
diff
changeset
|
41 |
(Scan.lift (Scan.optional (Args.$$$ "~" >> K ~1) 1 -- Parse.nat -- |
921a5be54132
support rat numerals via special antiquotation syntax;
wenzelm
parents:
63120
diff
changeset
|
42 |
Scan.optional (Args.$$$ "/" |-- Parse.nat) 1) >> (fn ((sign, a), b) => |
73586 | 43 |
"Rat.make " ^ ML_Syntax.print_pair ML_Syntax.print_int ML_Syntax.print_int (sign * a, b))) #> |
44 |
||
45 |
ML_Antiquotation.conditional \<^binding>\<open>if_linux\<close> (fn _ => ML_System.platform_is_linux) #> |
|
46 |
ML_Antiquotation.conditional \<^binding>\<open>if_macos\<close> (fn _ => ML_System.platform_is_macos) #> |
|
47 |
ML_Antiquotation.conditional \<^binding>\<open>if_windows\<close> (fn _ => ML_System.platform_is_windows) #> |
|
48 |
ML_Antiquotation.conditional \<^binding>\<open>if_unix\<close> (fn _ => ML_System.platform_is_unix)); |
|
58632 | 49 |
|
50 |
||
51 |
(* formal entities *) |
|
52 |
||
56205 | 53 |
val _ = Theory.setup |
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
54 |
(ML_Antiquotation.value_embedded \<^binding>\<open>system_option\<close> |
74563 | 55 |
(Args.context -- Scan.lift Parse.embedded_position >> (fn (ctxt, (name, pos)) => |
67208 | 56 |
(Completion.check_option (Options.default ()) ctxt (name, pos) |> ML_Syntax.print_string))) #> |
56205 | 57 |
|
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
58 |
ML_Antiquotation.value_embedded \<^binding>\<open>theory\<close> |
74563 | 59 |
(Args.context -- Scan.lift Parse.embedded_position >> (fn (ctxt, (name, pos)) => |
68482 | 60 |
(Theory.check {long = false} ctxt (name, pos); |
61 |
"Context.get_theory {long = false} (Proof_Context.theory_of ML_context) " ^ |
|
62 |
ML_Syntax.print_string name)) |
|
56205 | 63 |
|| Scan.succeed "Proof_Context.theory_of ML_context") #> |
64 |
||
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
65 |
ML_Antiquotation.value_embedded \<^binding>\<open>theory_context\<close> |
74563 | 66 |
(Args.context -- Scan.lift Parse.embedded_position >> (fn (ctxt, (name, pos)) => |
68482 | 67 |
(Theory.check {long = false} ctxt (name, pos); |
77889
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
wenzelm
parents:
74880
diff
changeset
|
68 |
"Proof_Context.get_global {long = false} (Proof_Context.theory_of ML_context) " ^ |
56205 | 69 |
ML_Syntax.print_string name))) #> |
70 |
||
67147 | 71 |
ML_Antiquotation.inline \<^binding>\<open>context\<close> |
59127
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents:
59112
diff
changeset
|
72 |
(Args.context >> (fn ctxt => ML_Context.struct_name ctxt ^ ".ML_context")) #> |
56205 | 73 |
|
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
74 |
ML_Antiquotation.inline_embedded \<^binding>\<open>typ\<close> |
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
75 |
(Args.typ >> (ML_Syntax.atomic o ML_Syntax.print_typ)) #> |
56205 | 76 |
|
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
77 |
ML_Antiquotation.inline_embedded \<^binding>\<open>term\<close> |
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
78 |
(Args.term >> (ML_Syntax.atomic o ML_Syntax.print_term)) #> |
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
79 |
|
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
80 |
ML_Antiquotation.inline_embedded \<^binding>\<open>prop\<close> |
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
81 |
(Args.prop >> (ML_Syntax.atomic o ML_Syntax.print_term)) #> |
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
82 |
|
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
83 |
ML_Antiquotation.value_embedded \<^binding>\<open>ctyp\<close> (Args.typ >> (fn T => |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59573
diff
changeset
|
84 |
"Thm.ctyp_of ML_context " ^ ML_Syntax.atomic (ML_Syntax.print_typ T))) #> |
56205 | 85 |
|
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
86 |
ML_Antiquotation.value_embedded \<^binding>\<open>cterm\<close> (Args.term >> (fn t => |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59573
diff
changeset
|
87 |
"Thm.cterm_of ML_context " ^ ML_Syntax.atomic (ML_Syntax.print_term t))) #> |
56205 | 88 |
|
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
89 |
ML_Antiquotation.value_embedded \<^binding>\<open>cprop\<close> (Args.prop >> (fn t => |
62075 | 90 |
"Thm.cterm_of ML_context " ^ ML_Syntax.atomic (ML_Syntax.print_term t))) #> |
91 |
||
70565 | 92 |
ML_Antiquotation.inline_embedded \<^binding>\<open>oracle_name\<close> |
74563 | 93 |
(Args.context -- Scan.lift Parse.embedded_position >> (fn (ctxt, (name, pos)) => |
73613 | 94 |
ML_Syntax.print_string (Thm.check_oracle ctxt (name, pos))))); |
63553
4a72b37ac4b8
text antiquotation for locales (similar to classes)
haftmann
parents:
63204
diff
changeset
|
95 |
|
4a72b37ac4b8
text antiquotation for locales (similar to classes)
haftmann
parents:
63204
diff
changeset
|
96 |
|
56205 | 97 |
(* type classes *) |
98 |
||
74563 | 99 |
fun class syn = Args.context -- Scan.lift Parse.embedded_inner_syntax >> (fn (ctxt, s) => |
56205 | 100 |
Proof_Context.read_class ctxt s |
101 |
|> syn ? Lexicon.mark_class |
|
102 |
|> ML_Syntax.print_string); |
|
103 |
||
104 |
val _ = Theory.setup |
|
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
105 |
(ML_Antiquotation.inline_embedded \<^binding>\<open>class\<close> (class false) #> |
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
106 |
ML_Antiquotation.inline_embedded \<^binding>\<open>class_syntax\<close> (class true) #> |
56205 | 107 |
|
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
108 |
ML_Antiquotation.inline_embedded \<^binding>\<open>sort\<close> |
74563 | 109 |
(Args.context -- Scan.lift Parse.embedded_inner_syntax >> (fn (ctxt, s) => |
56205 | 110 |
ML_Syntax.atomic (ML_Syntax.print_sort (Syntax.read_sort ctxt s))))); |
111 |
||
112 |
||
113 |
(* type constructors *) |
|
114 |
||
74563 | 115 |
fun type_name kind check = Args.context -- Scan.lift (Parse.token Parse.embedded) |
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:
68482
diff
changeset
|
116 |
>> (fn (ctxt, tok) => |
56205 | 117 |
let |
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:
68482
diff
changeset
|
118 |
val s = Token.inner_syntax_of tok; |
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:
68482
diff
changeset
|
119 |
val (_, pos) = Input.source_content (Token.input_of tok); |
56205 | 120 |
val Type (c, _) = Proof_Context.read_type_name {proper = true, strict = false} ctxt s; |
121 |
val decl = Type.the_decl (Proof_Context.tsig_of ctxt) (c, pos); |
|
122 |
val res = |
|
123 |
(case try check (c, decl) of |
|
124 |
SOME res => res |
|
125 |
| NONE => error ("Not a " ^ kind ^ ": " ^ quote c ^ Position.here pos)); |
|
126 |
in ML_Syntax.print_string res end); |
|
127 |
||
128 |
val _ = Theory.setup |
|
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
129 |
(ML_Antiquotation.inline_embedded \<^binding>\<open>type_name\<close> |
79469 | 130 |
(type_name "logical type" (fn (c, Type.Logical_Type _) => c)) #> |
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
131 |
ML_Antiquotation.inline_embedded \<^binding>\<open>type_abbrev\<close> |
56205 | 132 |
(type_name "type abbreviation" (fn (c, Type.Abbreviation _) => c)) #> |
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
133 |
ML_Antiquotation.inline_embedded \<^binding>\<open>nonterminal\<close> |
56205 | 134 |
(type_name "nonterminal" (fn (c, Type.Nonterminal) => c)) #> |
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
135 |
ML_Antiquotation.inline_embedded \<^binding>\<open>type_syntax\<close> |
56205 | 136 |
(type_name "type" (fn (c, _) => Lexicon.mark_type c))); |
137 |
||
138 |
||
139 |
(* constants *) |
|
140 |
||
74563 | 141 |
fun const_name check = Args.context -- Scan.lift (Parse.token Parse.embedded) |
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:
68482
diff
changeset
|
142 |
>> (fn (ctxt, tok) => |
56205 | 143 |
let |
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:
68482
diff
changeset
|
144 |
val s = Token.inner_syntax_of tok; |
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:
68482
diff
changeset
|
145 |
val (_, pos) = Input.source_content (Token.input_of tok); |
56205 | 146 |
val Const (c, _) = Proof_Context.read_const {proper = true, strict = false} ctxt s; |
147 |
val res = check (Proof_Context.consts_of ctxt, c) |
|
148 |
handle TYPE (msg, _, _) => error (msg ^ Position.here pos); |
|
149 |
in ML_Syntax.print_string res end); |
|
150 |
||
151 |
val _ = Theory.setup |
|
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
152 |
(ML_Antiquotation.inline_embedded \<^binding>\<open>const_name\<close> |
79463 | 153 |
(const_name (fn (consts, c) => (Consts.the_const_type consts c; c))) #> |
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
154 |
ML_Antiquotation.inline_embedded \<^binding>\<open>const_abbrev\<close> |
56205 | 155 |
(const_name (fn (consts, c) => (Consts.the_abbreviation consts c; c))) #> |
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
156 |
ML_Antiquotation.inline_embedded \<^binding>\<open>const_syntax\<close> |
56205 | 157 |
(const_name (fn (_, c) => Lexicon.mark_const c)) #> |
158 |
||
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
159 |
ML_Antiquotation.inline_embedded \<^binding>\<open>syntax_const\<close> |
74563 | 160 |
(Args.context -- Scan.lift Parse.embedded_position >> (fn (ctxt, arg) => |
74331 | 161 |
ML_Syntax.print_string (Proof_Context.check_syntax_const ctxt arg))) #> |
56205 | 162 |
|
69595 | 163 |
ML_Antiquotation.inline_embedded \<^binding>\<open>const\<close> |
74563 | 164 |
(Args.context -- Scan.lift (Parse.position Parse.embedded_inner_syntax) -- Scan.optional |
56205 | 165 |
(Scan.lift (Args.$$$ "(") |-- Parse.enum1' "," Args.typ --| Scan.lift (Args.$$$ ")")) [] |
56251 | 166 |
>> (fn ((ctxt, (raw_c, pos)), Ts) => |
56205 | 167 |
let |
168 |
val Const (c, _) = |
|
169 |
Proof_Context.read_const {proper = true, strict = true} ctxt raw_c; |
|
170 |
val consts = Proof_Context.consts_of ctxt; |
|
171 |
val n = length (Consts.typargs consts (c, Consts.type_scheme consts c)); |
|
172 |
val _ = length Ts <> n andalso |
|
173 |
error ("Constant requires " ^ string_of_int n ^ " type argument(s): " ^ |
|
56251 | 174 |
quote c ^ enclose "(" ")" (commas (replicate n "_")) ^ Position.here pos); |
56205 | 175 |
val const = Const (c, Consts.instance consts (c, Ts)); |
176 |
in ML_Syntax.atomic (ML_Syntax.print_term const) end))); |
|
177 |
||
178 |
||
74341 | 179 |
(* object-logic judgment *) |
180 |
||
74344
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
wenzelm
parents:
74341
diff
changeset
|
181 |
fun make_judgment ctxt = |
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
wenzelm
parents:
74341
diff
changeset
|
182 |
let val const = Object_Logic.judgment_const ctxt |
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
wenzelm
parents:
74341
diff
changeset
|
183 |
in fn t => Const const $ t end; |
74341 | 184 |
|
74344
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
wenzelm
parents:
74341
diff
changeset
|
185 |
fun dest_judgment ctxt = |
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
wenzelm
parents:
74341
diff
changeset
|
186 |
let |
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
wenzelm
parents:
74341
diff
changeset
|
187 |
val is_judgment = Object_Logic.is_judgment ctxt; |
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
wenzelm
parents:
74341
diff
changeset
|
188 |
val drop_judgment = Object_Logic.drop_judgment ctxt; |
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
wenzelm
parents:
74341
diff
changeset
|
189 |
in |
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
wenzelm
parents:
74341
diff
changeset
|
190 |
fn t => |
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
wenzelm
parents:
74341
diff
changeset
|
191 |
if is_judgment t then drop_judgment t |
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
wenzelm
parents:
74341
diff
changeset
|
192 |
else raise TERM ("dest_judgment", [t]) |
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
wenzelm
parents:
74341
diff
changeset
|
193 |
end; |
74341 | 194 |
|
195 |
val _ = Theory.setup |
|
196 |
(ML_Antiquotation.value \<^binding>\<open>make_judgment\<close> |
|
74559 | 197 |
(Scan.succeed "ML_Antiquotations.make_judgment ML_context") #> |
74341 | 198 |
ML_Antiquotation.value \<^binding>\<open>dest_judgment\<close> |
74559 | 199 |
(Scan.succeed "ML_Antiquotations.dest_judgment ML_context")); |
74341 | 200 |
|
201 |
||
74580 | 202 |
(* type/term constructors *) |
203 |
||
74556 | 204 |
local |
205 |
||
206 |
val keywords = Keyword.add_minor_keywords ["for", "=>"] Keyword.empty_keywords; |
|
207 |
||
74880 | 208 |
val parse_name_args = |
209 |
Parse.input Parse.name -- Scan.repeat Parse.embedded_ml; |
|
74374
e585e5a906ba
more convenient ML arguments: avoid excessive nesting of cartouches;
wenzelm
parents:
74344
diff
changeset
|
210 |
|
74880 | 211 |
val parse_for_args = |
212 |
Scan.optional (Parse.$$$ "for" |-- Parse.!!! (Scan.repeat1 Parse.embedded_ml)) []; |
|
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
213 |
|
74317 | 214 |
fun parse_body b = |
74880 | 215 |
if b then Parse.$$$ "=>" |-- Parse.!!! (Parse.embedded_ml >> single) else Scan.succeed []; |
74317 | 216 |
|
74374
e585e5a906ba
more convenient ML arguments: avoid excessive nesting of cartouches;
wenzelm
parents:
74344
diff
changeset
|
217 |
fun is_dummy [Antiquote.Text tok] = ML_Lex.content_of tok = "_" |
e585e5a906ba
more convenient ML arguments: avoid excessive nesting of cartouches;
wenzelm
parents:
74344
diff
changeset
|
218 |
| is_dummy _ = false; |
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
219 |
|
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
220 |
val ml = ML_Lex.tokenize_no_range; |
74317 | 221 |
val ml_range = ML_Lex.tokenize_range; |
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
222 |
val ml_dummy = ml "_"; |
74387
6edb71482de6
avoid overlapping PIDE markup (amending bb25ea271b15);
wenzelm
parents:
74378
diff
changeset
|
223 |
fun ml_enclose range x = ml "(" @ x @ ml_range range ")"; |
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
224 |
fun ml_parens x = ml "(" @ x @ ml ")"; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
225 |
fun ml_bracks x = ml "[" @ x @ ml "]"; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
226 |
fun ml_commas xs = flat (separate (ml ", ") xs); |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
227 |
val ml_list = ml_bracks o ml_commas; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
228 |
val ml_string = ml o ML_Syntax.print_string; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
229 |
fun ml_pair (x, y) = ml_parens (ml_commas [x, y]); |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
230 |
|
74317 | 231 |
fun type_antiquotation binding {function} = |
78804 | 232 |
ML_Context.add_antiquotation_embedded binding |
233 |
(fn range => fn input => fn ctxt => |
|
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
234 |
let |
78804 | 235 |
val ((s, type_args), fn_body) = input |
236 |
|> Parse.read_embedded ctxt keywords (parse_name_args -- parse_body function); |
|
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
237 |
val pos = Input.pos_of s; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
238 |
|
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
239 |
val Type (c, Ts) = |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
240 |
Proof_Context.read_type_name {proper = true, strict = true} ctxt |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
241 |
(Syntax.implode_input s); |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
242 |
val n = length Ts; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
243 |
val _ = |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
244 |
length type_args = n orelse |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
245 |
error ("Type constructor " ^ quote (Proof_Context.markup_type ctxt c) ^ |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
246 |
" takes " ^ string_of_int n ^ " argument(s)" ^ Position.here pos); |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
247 |
|
74562 | 248 |
val (decls1, ctxt1) = ML_Context.expand_antiquotes_list type_args ctxt; |
249 |
val (decls2, ctxt2) = ML_Context.expand_antiquotes_list fn_body ctxt1; |
|
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
250 |
fun decl' ctxt' = |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
251 |
let |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
252 |
val (ml_args_env, ml_args_body) = split_list (decls1 ctxt'); |
74317 | 253 |
val (ml_fn_env, ml_fn_body) = split_list (decls2 ctxt'); |
74378
bb25ea271b15
clarified positions, notably for ML compiler errors;
wenzelm
parents:
74377
diff
changeset
|
254 |
val ml1 = |
bb25ea271b15
clarified positions, notably for ML compiler errors;
wenzelm
parents:
74377
diff
changeset
|
255 |
ml_enclose range (ml "Term.Type " @ ml_pair (ml_string c, ml_list ml_args_body)); |
74317 | 256 |
val ml2 = |
257 |
if function then |
|
74378
bb25ea271b15
clarified positions, notably for ML compiler errors;
wenzelm
parents:
74377
diff
changeset
|
258 |
ml_enclose range |
bb25ea271b15
clarified positions, notably for ML compiler errors;
wenzelm
parents:
74377
diff
changeset
|
259 |
(ml_range range "fn " @ ml1 @ ml "=> " @ flat ml_fn_body @ |
bb25ea271b15
clarified positions, notably for ML compiler errors;
wenzelm
parents:
74377
diff
changeset
|
260 |
ml "| T => " @ ml_range range "raise" @ |
bb25ea271b15
clarified positions, notably for ML compiler errors;
wenzelm
parents:
74377
diff
changeset
|
261 |
ml " Term.TYPE (" @ ml_string ("Type_fn " ^ quote c) @ ml ", [T], [])") |
74317 | 262 |
else ml1; |
263 |
in (flat (ml_args_env @ ml_fn_env), ml2) end; |
|
264 |
in (decl', ctxt2) end); |
|
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
265 |
|
74317 | 266 |
fun const_antiquotation binding {pattern, function} = |
78804 | 267 |
ML_Context.add_antiquotation_embedded binding |
268 |
(fn range => fn input => fn ctxt => |
|
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
269 |
let |
78804 | 270 |
val (((s, type_args), term_args), fn_body) = input |
271 |
|> Parse.read_embedded ctxt keywords |
|
74880 | 272 |
(parse_name_args -- parse_for_args -- parse_body function); |
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
273 |
|
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
274 |
val Const (c, T) = |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
275 |
Proof_Context.read_const {proper = true, strict = true} ctxt |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
276 |
(Syntax.implode_input s); |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
277 |
|
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
278 |
val consts = Proof_Context.consts_of ctxt; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
279 |
val type_paths = Consts.type_arguments consts c; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
280 |
val type_params = map Term.dest_TVar (Consts.typargs consts (c, T)); |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
281 |
|
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
282 |
val n = length type_params; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
283 |
val m = length (Term.binder_types T); |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
284 |
|
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
285 |
fun err msg = |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
286 |
error ("Constant " ^ quote (Proof_Context.markup_const ctxt c) ^ msg ^ |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
287 |
Position.here (Input.pos_of s)); |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
288 |
val _ = |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
289 |
length type_args <> n andalso err (" takes " ^ string_of_int n ^ " type argument(s)"); |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
290 |
val _ = |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
291 |
length term_args > m andalso Term.is_Type (Term.body_type T) andalso |
74377 | 292 |
err (" cannot have more than " ^ string_of_int m ^ " argument(s)"); |
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
293 |
|
74562 | 294 |
val (decls1, ctxt1) = ML_Context.expand_antiquotes_list type_args ctxt; |
295 |
val (decls2, ctxt2) = ML_Context.expand_antiquotes_list term_args ctxt1; |
|
296 |
val (decls3, ctxt3) = ML_Context.expand_antiquotes_list fn_body ctxt2; |
|
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
297 |
fun decl' ctxt' = |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
298 |
let |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
299 |
val (ml_args_env1, ml_args_body1) = split_list (decls1 ctxt'); |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
300 |
val (ml_args_env2, ml_args_body2) = split_list (decls2 ctxt'); |
74317 | 301 |
val (ml_fn_env, ml_fn_body) = split_list (decls3 ctxt'); |
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
302 |
|
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
303 |
val relevant = map is_dummy type_args ~~ type_paths; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
304 |
fun relevant_path is = |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
305 |
not pattern orelse |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
306 |
let val p = rev is |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
307 |
in relevant |> exists (fn (u, q) => not u andalso is_prefix (op =) p q) end; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
308 |
|
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
309 |
val ml_typarg = the o AList.lookup (op =) (type_params ~~ ml_args_body1); |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
310 |
fun ml_typ is (Type (d, Us)) = |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
311 |
if relevant_path is then |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
312 |
ml "Term.Type " @ |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
313 |
ml_pair (ml_string d, ml_list (map_index (fn (i, U) => ml_typ (i :: is) U) Us)) |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
314 |
else ml_dummy |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
315 |
| ml_typ is (TVar arg) = if relevant_path is then ml_typarg arg else ml_dummy |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
316 |
| ml_typ _ (TFree _) = raise Match; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
317 |
|
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
318 |
fun ml_app [] = ml "Term.Const " @ ml_pair (ml_string c, ml_typ [] T) |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
319 |
| ml_app (u :: us) = ml "Term.$ " @ ml_pair (ml_app us, u); |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
320 |
|
74317 | 321 |
val ml_env = flat (ml_args_env1 @ ml_args_env2 @ ml_fn_env); |
74378
bb25ea271b15
clarified positions, notably for ML compiler errors;
wenzelm
parents:
74377
diff
changeset
|
322 |
val ml1 = ml_enclose range (ml_app (rev ml_args_body2)); |
74317 | 323 |
val ml2 = |
324 |
if function then |
|
74378
bb25ea271b15
clarified positions, notably for ML compiler errors;
wenzelm
parents:
74377
diff
changeset
|
325 |
ml_enclose range |
bb25ea271b15
clarified positions, notably for ML compiler errors;
wenzelm
parents:
74377
diff
changeset
|
326 |
(ml_range range "fn " @ ml1 @ ml "=> " @ flat ml_fn_body @ |
bb25ea271b15
clarified positions, notably for ML compiler errors;
wenzelm
parents:
74377
diff
changeset
|
327 |
ml "| t => " @ ml_range range "raise" @ |
bb25ea271b15
clarified positions, notably for ML compiler errors;
wenzelm
parents:
74377
diff
changeset
|
328 |
ml " Term.TERM (" @ ml_string ("Const_fn " ^ quote c) @ ml ", [t])") |
74317 | 329 |
else ml1; |
330 |
in (ml_env, ml2) end; |
|
331 |
in (decl', ctxt3) end); |
|
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
332 |
|
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
333 |
val _ = Theory.setup |
74317 | 334 |
(type_antiquotation \<^binding>\<open>Type\<close> {function = false} #> |
335 |
type_antiquotation \<^binding>\<open>Type_fn\<close> {function = true} #> |
|
336 |
const_antiquotation \<^binding>\<open>Const\<close> {pattern = false, function = false} #> |
|
337 |
const_antiquotation \<^binding>\<open>Const_\<close> {pattern = true, function = false} #> |
|
338 |
const_antiquotation \<^binding>\<open>Const_fn\<close> {pattern = true, function = true}); |
|
74291
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
339 |
|
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
340 |
in end; |
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
341 |
|
b83fa8f3a271
ML antiquotations for type constructors and term constants;
wenzelm
parents:
74290
diff
changeset
|
342 |
|
78701 | 343 |
(* exception handling *) |
344 |
||
345 |
local |
|
346 |
||
347 |
val tokenize_text = map Antiquote.Text o ML_Lex.tokenize_no_range; |
|
78703
55b1664b564b
more position information, e.g. for warning about fn-pattern;
wenzelm
parents:
78702
diff
changeset
|
348 |
val tokenize_range_text = map Antiquote.Text oo ML_Lex.tokenize_range; |
78701 | 349 |
|
350 |
val bg_expr = tokenize_text "(fn () =>"; |
|
351 |
val en_expr = tokenize_text ")"; |
|
352 |
fun make_expr a = bg_expr @ a @ en_expr; |
|
353 |
||
78703
55b1664b564b
more position information, e.g. for warning about fn-pattern;
wenzelm
parents:
78702
diff
changeset
|
354 |
fun make_handler range a = |
55b1664b564b
more position information, e.g. for warning about fn-pattern;
wenzelm
parents:
78702
diff
changeset
|
355 |
tokenize_range_text range "(fn" @ a @ |
55b1664b564b
more position information, e.g. for warning about fn-pattern;
wenzelm
parents:
78702
diff
changeset
|
356 |
tokenize_range_text range "| exn => Exn.reraise exn)"; |
78701 | 357 |
|
358 |
fun report_special tok = (ML_Lex.pos_of tok, Markup.keyword3); |
|
359 |
||
360 |
fun print_special tok = |
|
361 |
let val (pos, markup) = report_special tok |
|
362 |
in Markup.markup markup (ML_Lex.content_of tok) ^ Position.here pos end; |
|
363 |
||
364 |
val is_catch = ML_Lex.is_ident_with (fn x => x = "catch"); |
|
365 |
val is_finally = ML_Lex.is_ident_with (fn x => x = "finally"); |
|
366 |
fun is_special t = is_catch t orelse is_finally t; |
|
367 |
val is_special_text = fn Antiquote.Text t => is_special t | _ => false; |
|
368 |
||
369 |
fun parse_try ctxt input = |
|
370 |
let |
|
371 |
val ants = ML_Lex.read_source input; |
|
372 |
val specials = ants |
|
373 |
|> map_filter (fn Antiquote.Text t => if is_special t then SOME t else NONE | _ => NONE); |
|
374 |
val _ = Context_Position.reports ctxt (map report_special specials); |
|
375 |
in |
|
376 |
(case specials of |
|
377 |
[] => ("Isabelle_Thread.try", [make_expr ants]) |
|
378 |
| [s] => |
|
379 |
let val (a, b) = ants |> chop_prefix (not o is_special_text) ||> tl in |
|
380 |
if is_finally s |
|
381 |
then ("Isabelle_Thread.try_finally", [make_expr a, make_expr b]) |
|
78703
55b1664b564b
more position information, e.g. for warning about fn-pattern;
wenzelm
parents:
78702
diff
changeset
|
382 |
else ("Isabelle_Thread.try_catch", [make_expr a, make_handler (ML_Lex.range_of s) b]) |
78701 | 383 |
end |
384 |
| _ => error ("Too many special keywords: " ^ commas (map print_special specials))) |
|
385 |
end; |
|
386 |
||
387 |
fun parse_can (_: Proof.context) input = |
|
388 |
("Isabelle_Thread.can", [make_expr (ML_Lex.read_source input)]); |
|
389 |
||
390 |
in |
|
73550
2f6855142a8c
support for ML special forms: modified evaluation similar to Scheme;
wenzelm
parents:
70565
diff
changeset
|
391 |
|
2f6855142a8c
support for ML special forms: modified evaluation similar to Scheme;
wenzelm
parents:
70565
diff
changeset
|
392 |
val _ = Theory.setup |
78701 | 393 |
(ML_Antiquotation.special_form \<^binding>\<open>try\<close> parse_try #> |
394 |
ML_Antiquotation.special_form \<^binding>\<open>can\<close> parse_can); |
|
395 |
||
396 |
end; |
|
397 |
||
398 |
||
399 |
||
400 |
(* special form for option type *) |
|
401 |
||
402 |
val _ = Theory.setup |
|
78804 | 403 |
(ML_Context.add_antiquotation_embedded \<^binding>\<open>if_none\<close> |
404 |
(fn _ => fn input => fn ctxt => |
|
77907
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
wenzelm
parents:
77889
diff
changeset
|
405 |
let |
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
wenzelm
parents:
77889
diff
changeset
|
406 |
val tokenize = ML_Lex.tokenize_no_range; |
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
wenzelm
parents:
77889
diff
changeset
|
407 |
|
78804 | 408 |
val (decl, ctxt') = ML_Context.read_antiquotes input ctxt; |
77907
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
wenzelm
parents:
77889
diff
changeset
|
409 |
fun decl' ctxt'' = |
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
wenzelm
parents:
77889
diff
changeset
|
410 |
let |
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
wenzelm
parents:
77889
diff
changeset
|
411 |
val (ml_env, ml_body) = decl ctxt''; |
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
wenzelm
parents:
77889
diff
changeset
|
412 |
val ml_body' = tokenize "(fn SOME x => x | NONE => " @ ml_body @ tokenize ")"; |
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
wenzelm
parents:
77889
diff
changeset
|
413 |
in (ml_env, ml_body') end; |
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
wenzelm
parents:
77889
diff
changeset
|
414 |
in (decl', ctxt') end)); |
73550
2f6855142a8c
support for ML special forms: modified evaluation similar to Scheme;
wenzelm
parents:
70565
diff
changeset
|
415 |
|
2f6855142a8c
support for ML special forms: modified evaluation similar to Scheme;
wenzelm
parents:
70565
diff
changeset
|
416 |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
417 |
(* basic combinators *) |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
418 |
|
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
419 |
local |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
420 |
|
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
421 |
val parameter = Parse.position Parse.nat >> (fn (n, pos) => |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
422 |
if n > 1 then n else error ("Bad parameter: " ^ string_of_int n ^ Position.here pos)); |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
423 |
|
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
424 |
fun indices n = map string_of_int (1 upto n); |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
425 |
|
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
426 |
fun empty n = replicate_string n " []"; |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
427 |
fun dummy n = replicate_string n " _"; |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
428 |
fun vars x n = implode (map (fn a => " " ^ x ^ a) (indices n)); |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
429 |
fun cons n = implode (map (fn a => " (x" ^ a ^ " :: xs" ^ a ^ ")") (indices n)); |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
430 |
|
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
431 |
val tuple = enclose "(" ")" o commas; |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
432 |
fun tuple_empty n = tuple (replicate n "[]"); |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
433 |
fun tuple_vars x n = tuple (map (fn a => x ^ a) (indices n)); |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
434 |
fun tuple_cons n = "(" ^ tuple_vars "x" n ^ " :: xs)" |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
435 |
fun cons_tuple n = tuple (map (fn a => "x" ^ a ^ " :: xs" ^ a) (indices n)); |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
436 |
|
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
437 |
in |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
438 |
|
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
439 |
val _ = Theory.setup |
67147 | 440 |
(ML_Antiquotation.value \<^binding>\<open>map\<close> |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
441 |
(Scan.lift parameter >> (fn n => |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
442 |
"fn f =>\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
443 |
\ let\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
444 |
\ fun map _" ^ empty n ^ " = []\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
445 |
\ | map f" ^ cons n ^ " = f" ^ vars "x" n ^ " :: map f" ^ vars "xs" n ^ "\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
446 |
\ | map _" ^ dummy n ^ " = raise ListPair.UnequalLengths\n" ^ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
447 |
" in map f end")) #> |
67147 | 448 |
ML_Antiquotation.value \<^binding>\<open>fold\<close> |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
449 |
(Scan.lift parameter >> (fn n => |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
450 |
"fn f =>\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
451 |
\ let\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
452 |
\ fun fold _" ^ empty n ^ " a = a\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
453 |
\ | fold f" ^ cons n ^ " a = fold f" ^ vars "xs" n ^ " (f" ^ vars "x" n ^ " a)\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
454 |
\ | fold _" ^ dummy n ^ " _ = raise ListPair.UnequalLengths\n" ^ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
455 |
" in fold f end")) #> |
67147 | 456 |
ML_Antiquotation.value \<^binding>\<open>fold_map\<close> |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
457 |
(Scan.lift parameter >> (fn n => |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
458 |
"fn f =>\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
459 |
\ let\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
460 |
\ fun fold_map _" ^ empty n ^ " a = ([], a)\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
461 |
\ | fold_map f" ^ cons n ^ " a =\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
462 |
\ let\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
463 |
\ val (x, a') = f" ^ vars "x" n ^ " a\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
464 |
\ val (xs, a'') = fold_map f" ^ vars "xs" n ^ " a'\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
465 |
\ in (x :: xs, a'') end\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
466 |
\ | fold_map _" ^ dummy n ^ " _ = raise ListPair.UnequalLengths\n" ^ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
467 |
" in fold_map f end")) #> |
67147 | 468 |
ML_Antiquotation.value \<^binding>\<open>split_list\<close> |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
469 |
(Scan.lift parameter >> (fn n => |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
470 |
"fn list =>\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
471 |
\ let\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
472 |
\ fun split_list [] =" ^ tuple_empty n ^ "\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
473 |
\ | split_list" ^ tuple_cons n ^ " =\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
474 |
\ let val" ^ tuple_vars "xs" n ^ " = split_list xs\n\ |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
475 |
\ in " ^ cons_tuple n ^ "end\n\ |
59057
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
476 |
\ in split_list list end")) #> |
67147 | 477 |
ML_Antiquotation.value \<^binding>\<open>apply\<close> |
59057
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
478 |
(Scan.lift (parameter -- Scan.option (Args.parens (Parse.position Parse.nat))) >> |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
479 |
(fn (n, opt_index) => |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
480 |
let |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
481 |
val cond = |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
482 |
(case opt_index of |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
483 |
NONE => K true |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
484 |
| SOME (index, index_pos) => |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
485 |
if 1 <= index andalso index <= n then equal (string_of_int index) |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
486 |
else error ("Bad index: " ^ string_of_int index ^ Position.here index_pos)); |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
487 |
in |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
488 |
"fn f => fn " ^ tuple_vars "x" n ^ " => " ^ |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
489 |
tuple (map (fn a => (if cond a then "f x" else "x") ^ a) (indices n)) |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58978
diff
changeset
|
490 |
end))); |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
491 |
|
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
492 |
end; |
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
493 |
|
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58632
diff
changeset
|
494 |
|
56205 | 495 |
(* outer syntax *) |
496 |
||
497 |
val _ = Theory.setup |
|
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
498 |
(ML_Antiquotation.value_embedded \<^binding>\<open>keyword\<close> |
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:
68482
diff
changeset
|
499 |
(Args.context -- |
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:
68482
diff
changeset
|
500 |
Scan.lift (Parse.embedded_position || Parse.position (Parse.keyword_with (K true))) |
64594 | 501 |
>> (fn (ctxt, (name, pos)) => |
502 |
if Keyword.is_keyword (Thy_Header.get_keywords' ctxt) name then |
|
503 |
(Context_Position.report ctxt pos (Token.keyword_markup (true, Markup.keyword2) name); |
|
504 |
"Parse.$$$ " ^ ML_Syntax.print_string name) |
|
505 |
else error ("Bad outer syntax keyword " ^ quote name ^ Position.here pos))) #> |
|
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
506 |
ML_Antiquotation.value_embedded \<^binding>\<open>command_keyword\<close> |
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:
68482
diff
changeset
|
507 |
(Args.context -- Scan.lift Parse.embedded_position >> (fn (ctxt, (name, pos)) => |
59934
b65c4370f831
more position information and PIDE markup for command keywords;
wenzelm
parents:
59878
diff
changeset
|
508 |
(case Keyword.command_markup (Thy_Header.get_keywords' ctxt) name of |
b65c4370f831
more position information and PIDE markup for command keywords;
wenzelm
parents:
59878
diff
changeset
|
509 |
SOME markup => |
b65c4370f831
more position information and PIDE markup for command keywords;
wenzelm
parents:
59878
diff
changeset
|
510 |
(Context_Position.reports ctxt [(pos, markup), (pos, Markup.keyword1)]; |
b65c4370f831
more position information and PIDE markup for command keywords;
wenzelm
parents:
59878
diff
changeset
|
511 |
ML_Syntax.print_pair ML_Syntax.print_string ML_Syntax.print_position (name, pos)) |
b65c4370f831
more position information and PIDE markup for command keywords;
wenzelm
parents:
59878
diff
changeset
|
512 |
| NONE => error ("Bad outer syntax command " ^ quote name ^ Position.here pos))))); |
56205 | 513 |
|
514 |
end; |