| author | wenzelm | 
| Fri, 03 Dec 2010 20:38:58 +0100 | |
| changeset 40945 | b8703f63bfb2 | 
| parent 25867 | c24395ea4e71 | 
| child 41589 | bbd861837ebc | 
| permissions | -rw-r--r-- | 
| 18269 | 1 | (* $Id$ *) | 
| 18106 | 2 | |
| 19501 | 3 | theory Lambda_mu | 
| 25867 | 4 | imports "../Nominal" | 
| 18106 | 5 | begin | 
| 6 | ||
| 18269 | 7 | section {* Lambda-Mu according to a paper by Gavin Bierman *}
 | 
| 18106 | 8 | |
| 9 | atom_decl var mvar | |
| 10 | ||
| 22829 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
19501diff
changeset | 11 | nominal_datatype trm = | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
19501diff
changeset | 12 | Var "var" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
19501diff
changeset | 13 |   | Lam  "\<guillemotleft>var\<guillemotright>trm"   ("Lam [_]._" [100,100] 100)
 | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
19501diff
changeset | 14 | | App "trm" "trm" | 
| 25867 | 15 | | Pss "mvar" "trm" (* passivate *) | 
| 16 |   | Act  "\<guillemotleft>mvar\<guillemotright>trm"  ("Act [_]._" [100,100] 100)       (* activate  *)
 | |
| 17 | ||
| 18106 | 18 | |
| 19501 | 19 | end |