author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Thu, 26 Aug 2010 14:55:15 +0900 | |
changeset 2437 | 319f469b8b67 |
parent 2436 | 3885dc2669f9 |
child 2454 | 9ffee4eb1ae1 |
permissions | -rw-r--r-- |
theory ExPS7 imports "../NewParser" begin (* example 7 from Peter Sewell's bestiary *) atom_decl name nominal_datatype exp = Var name | Unit | Pair exp exp | LetRec l::"lrbs" e::"exp" bind (set) "bs l" in l e and lrb = Assign name exp and lrbs = Single lrb | More lrb lrbs binder b :: "lrb \<Rightarrow> atom set" and bs :: "lrbs \<Rightarrow> atom set" where "b (Assign x e) = {atom x}" | "bs (Single a) = b a" | "bs (More a as) = (b a) \<union> (bs as)" end