13020
|
1 |
|
|
2 |
header {* \section{Concrete Syntax} *}
|
|
3 |
|
16417
|
4 |
theory Quote_Antiquote imports Main begin
|
13020
|
5 |
|
|
6 |
syntax
|
|
7 |
"_quote" :: "'b \<Rightarrow> ('a \<Rightarrow> 'b)" ("(\<guillemotleft>_\<guillemotright>)" [0] 1000)
|
|
8 |
"_antiquote" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b" ("\<acute>_" [1000] 1000)
|
|
9 |
"_Assert" :: "'a \<Rightarrow> 'a set" ("(.{_}.)" [0] 1000)
|
|
10 |
|
|
11 |
syntax (xsymbols)
|
|
12 |
"_Assert" :: "'a \<Rightarrow> 'a set" ("(\<lbrace>_\<rbrace>)" [0] 1000)
|
|
13 |
|
|
14 |
translations
|
35107
|
15 |
".{b}." \<rightharpoonup> "CONST Collect \<guillemotleft>b\<guillemotright>"
|
13020
|
16 |
|
|
17 |
parse_translation {*
|
|
18 |
let
|
35113
|
19 |
fun quote_tr [t] = Syntax.quote_tr @{syntax_const "_antiquote"} t
|
13020
|
20 |
| quote_tr ts = raise TERM ("quote_tr", ts);
|
35113
|
21 |
in [(@{syntax_const "_quote"}, quote_tr)] end
|
13020
|
22 |
*}
|
|
23 |
|
|
24 |
end |