diff -r 48c2eb84d5ce -r c0eac04ae3b4 Nominal/Ex/ExLF.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/Ex/ExLF.thy Sat Apr 03 22:31:11 2010 +0200 @@ -0,0 +1,27 @@ +theory ExLF +imports "../Parser" +begin + +atom_decl name +atom_decl ident + +nominal_datatype kind = + Type + | KPi "ty" n::"name" k::"kind" bind n in k +and ty = + TConst "ident" + | TApp "ty" "trm" + | TPi "ty" n::"name" t::"ty" bind n in t +and trm = + Const "ident" + | Var "name" + | App "trm" "trm" + | Lam "ty" n::"name" t::"trm" bind n in t + +thm kind_ty_trm.supp + +end + + + +