Nominal/Ex/Ex2.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 12 May 2010 16:32:44 +0200
changeset 2114 3201a8c3456b
parent 2105 e25b0fff0dd2
child 2120 2786ff1df475
permissions -rw-r--r--
Use equivariance instead of alpha_eqvt

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)]

equivariance alpha_trm_raw


end