author | Christian Urban <urbanc@in.tum.de> |
Fri, 21 May 2010 00:44:39 +0100 | |
changeset 2291 | 20ee31bc34d5 |
parent 2120 | 2786ff1df475 |
child 2311 | 4da5c5c29009 |
permissions | -rw-r--r-- |
theory LF imports "../NewParser" begin atom_decl name atom_decl ident nominal_datatype kind = Type | KPi "ty" n::"name" k::"kind" bind_set n in k and ty = TConst "ident" | TApp "ty" "trm" | TPi "ty" n::"name" t::"ty" bind_set n in t and trm = Const "ident" | Var "name" | App "trm" "trm" | Lam "ty" n::"name" t::"trm" bind_set n in t thm kind_ty_trm.supp end