| 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)
 | 
| 53241 |      9 |   "_Assert"    :: "'a \<Rightarrow> 'a set"                    ("(\<lbrace>_\<rbrace>)" [0] 1000)
 | 
| 13020 |     10 | 
 | 
|  |     11 | translations
 | 
| 53241 |     12 |   "\<lbrace>b\<rbrace>" \<rightharpoonup> "CONST Collect \<guillemotleft>b\<guillemotright>"
 | 
| 13020 |     13 | 
 | 
|  |     14 | parse_translation {*
 | 
|  |     15 |   let
 | 
| 42284 |     16 |     fun quote_tr [t] = Syntax_Trans.quote_tr @{syntax_const "_antiquote"} t
 | 
| 13020 |     17 |       | quote_tr ts = raise TERM ("quote_tr", ts);
 | 
| 52143 |     18 |   in [(@{syntax_const "_quote"}, K quote_tr)] end
 | 
| 13020 |     19 | *}
 | 
|  |     20 | 
 | 
|  |     21 | end |