| author | blanchet | 
| Thu, 05 May 2011 02:27:02 +0200 | |
| changeset 42690 | 4d29b4785f43 | 
| parent 42284 | 326f57825e1a | 
| child 44603 | a6f9a70d655d | 
| permissions | -rw-r--r-- | 
| 5368 | 1 | (* Title: HOL/ex/Antiquote.thy | 
| 2 | Author: Markus Wenzel, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 10357 | 5 | header {* Antiquotations *}
 | 
| 6 | ||
| 35113 | 7 | theory Antiquote | 
| 8 | imports Main | |
| 9 | begin | |
| 5368 | 10 | |
| 11586 | 11 | text {*
 | 
| 12 | A simple example on quote / antiquote in higher-order abstract | |
| 13 | syntax. | |
| 14 | *} | |
| 15 | ||
| 5368 | 16 | syntax | 
| 35113 | 17 |   "_Expr" :: "'a => 'a"    ("EXPR _" [1000] 999)
 | 
| 5368 | 18 | |
| 35113 | 19 | definition | 
| 20 |   var :: "'a => ('a => nat) => nat"    ("VAR _" [1000] 999)
 | |
| 21 | where "var x env = env x" | |
| 5368 | 22 | |
| 35113 | 23 | definition | 
| 8559 | 24 |   Expr :: "(('a => nat) => nat) => ('a => nat) => nat"
 | 
| 35113 | 25 | where "Expr exp env = exp env" | 
| 5368 | 26 | |
| 35113 | 27 | parse_translation {*
 | 
| 42284 | 28 | [Syntax_Trans.quote_antiquote_tr | 
| 29 |     @{syntax_const "_Expr"} @{const_syntax var} @{const_syntax Expr}]
 | |
| 35113 | 30 | *} | 
| 31 | ||
| 32 | print_translation {*
 | |
| 42284 | 33 | [Syntax_Trans.quote_antiquote_tr' | 
| 34 |     @{syntax_const "_Expr"} @{const_syntax var} @{const_syntax Expr}]
 | |
| 35113 | 35 | *} | 
| 8559 | 36 | |
| 9297 | 37 | term "EXPR (a + b + c)" | 
| 38 | term "EXPR (a + b + c + VAR x + VAR y + 1)" | |
| 39 | term "EXPR (VAR (f w) + VAR x)" | |
| 5368 | 40 | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 41 | term "Expr (\<lambda>env. env x)"    -- {* improper *}
 | 
| 10357 | 42 | term "Expr (\<lambda>env. f env)" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 43 | term "Expr (\<lambda>env. f env + env x)"    -- {* improper *}
 | 
| 10357 | 44 | term "Expr (\<lambda>env. f env y z)" | 
| 45 | term "Expr (\<lambda>env. f env + g y env)" | |
| 46 | term "Expr (\<lambda>env. f env + g env y + h a env z)" | |
| 5368 | 47 | |
| 9297 | 48 | end |