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