| 5368 |      1 | (*  Title:      HOL/ex/Antiquote.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Markus Wenzel, TU Muenchen
 | 
|  |      4 | 
 | 
|  |      5 | Simple quote / antiquote example.
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
|  |      8 | Antiquote = Arith +
 | 
|  |      9 | 
 | 
|  |     10 | syntax
 | 
|  |     11 |   "_Expr" :: "'a => 'a"				("EXPR _" [1000] 999)
 | 
|  |     12 | 
 | 
|  |     13 | constdefs
 | 
|  |     14 |   var :: "'a => ('a => nat) => nat"		("VAR _" [1000] 999)
 | 
|  |     15 |   "var x env == env x"
 | 
|  |     16 | 
 | 
|  |     17 |   Expr :: "'a => 'a"
 | 
|  |     18 | 	(*"(('a => nat) => nat) => (('a => nat) => nat)"*)
 | 
|  |     19 |   "Expr == (%x. x)"
 | 
|  |     20 | 
 | 
|  |     21 | end
 | 
|  |     22 | 
 | 
|  |     23 | 
 | 
|  |     24 | ML
 | 
|  |     25 | 
 | 
|  |     26 | val parse_translation = [Syntax.quote_antiquote_tr "_Expr" "var" "Expr"];
 | 
|  |     27 | val print_translation = [Syntax.quote_antiquote_tr' "_Expr" "var" "Expr"];
 |