src/HOL/Nominal/Examples/Lambda_mu.thy
author haftmann
Fri, 23 Nov 2007 21:09:32 +0100
changeset 25459 d1dce7d0731c
parent 22829 f1db55c7534d
child 25867 c24395ea4e71
permissions -rw-r--r--
deleted card definition as code lemma; authentic syntax for card

(* $Id$ *)

theory Lambda_mu 
imports "../Nominal" 
begin

section {* Lambda-Mu according to a paper by Gavin Bierman *}

atom_decl var mvar

nominal_datatype trm = 
    Var   "var" 
  | Lam  "\<guillemotleft>var\<guillemotright>trm"   ("Lam [_]._" [100,100] 100)
  | App  "trm" "trm" 
  | Pss  "mvar" "trm"
  | Act  "\<guillemotleft>mvar\<guillemotright>trm"  ("Act [_]._" [100,100] 100)

end