author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Tue, 04 May 2010 11:12:09 +0200 | |
changeset 2040 | 94e24da9ae75 |
parent 2030 | 43d7612f1429 |
child 2082 | 0854af516f14 |
permissions | -rw-r--r-- |
theory Ex2 imports "../NewParser" begin text {* example 2 *} atom_decl name nominal_datatype trm = Var "name" | App "trm" "trm" | Lam x::"name" t::"trm" bind_set x in t | Let p::"pat" "trm" t::"trm" bind_set "f p" in t and pat = PN | PS "name" | PD "name" "name" binder f::"pat \<Rightarrow> atom set" where "f PN = {}" | "f (PD x y) = {atom x, atom y}" | "f (PS x) = {atom x}" thm trm_pat.bn thm trm_pat.perm thm trm_pat.induct thm trm_pat.distinct thm trm_pat.fv[simplified trm_pat.supp(1-2)] end