| 5368 |      1 | (*  Title:      HOL/ex/Antiquote.thy
 | 
|  |      2 |     Author:     Markus Wenzel, TU Muenchen
 | 
|  |      3 | *)
 | 
|  |      4 | 
 | 
| 59031 |      5 | section \<open>Antiquotations\<close>
 | 
| 10357 |      6 | 
 | 
| 35113 |      7 | theory Antiquote
 | 
|  |      8 | imports Main
 | 
|  |      9 | begin
 | 
| 5368 |     10 | 
 | 
| 59031 |     11 | text \<open>A simple example on quote / antiquote in higher-order abstract syntax.\<close>
 | 
| 11586 |     12 | 
 | 
| 56233 |     13 | definition var :: "'a \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> nat"  ("VAR _" [1000] 999)
 | 
| 35113 |     14 |   where "var x env = env x"
 | 
| 5368 |     15 | 
 | 
| 56233 |     16 | definition Expr :: "(('a \<Rightarrow> nat) \<Rightarrow> nat) \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> nat"
 | 
| 35113 |     17 |   where "Expr exp env = exp env"
 | 
| 5368 |     18 | 
 | 
| 44603 |     19 | syntax
 | 
| 56233 |     20 |   "_Expr" :: "'a \<Rightarrow> 'a"  ("EXPR _" [1000] 999)
 | 
| 44603 |     21 | 
 | 
| 59031 |     22 | parse_translation
 | 
|  |     23 |   \<open>[Syntax_Trans.quote_antiquote_tr
 | 
|  |     24 |     @{syntax_const "_Expr"} @{const_syntax var} @{const_syntax Expr}]\<close>
 | 
| 35113 |     25 | 
 | 
| 59031 |     26 | print_translation
 | 
|  |     27 |   \<open>[Syntax_Trans.quote_antiquote_tr'
 | 
|  |     28 |     @{syntax_const "_Expr"} @{const_syntax var} @{const_syntax Expr}]\<close>
 | 
| 8559 |     29 | 
 | 
| 9297 |     30 | term "EXPR (a + b + c)"
 | 
|  |     31 | term "EXPR (a + b + c + VAR x + VAR y + 1)"
 | 
|  |     32 | term "EXPR (VAR (f w) + VAR x)"
 | 
| 5368 |     33 | 
 | 
| 61933 |     34 | term "Expr (\<lambda>env. env x)"    \<comment> \<open>improper\<close>
 | 
| 10357 |     35 | term "Expr (\<lambda>env. f env)"
 | 
| 61933 |     36 | term "Expr (\<lambda>env. f env + env x)"    \<comment> \<open>improper\<close>
 | 
| 10357 |     37 | term "Expr (\<lambda>env. f env y z)"
 | 
|  |     38 | term "Expr (\<lambda>env. f env + g y env)"
 | 
|  |     39 | term "Expr (\<lambda>env. f env + g env y + h a env z)"
 | 
| 5368 |     40 | 
 | 
| 9297 |     41 | end
 |