theory Let
imports "../Nominal2"
begin
atom_decl name
nominal_datatype trm =
Var "name"
| App "trm" "trm"
| Lam x::"name" t::"trm" binds x in t
| Let as::"assn" t::"trm" binds "bn as" in t
and assn =
ANil
| ACons "name" "trm" "assn"
binder
bn
where
"bn ANil = []"
| "bn (ACons x t as) = (atom x) # (bn as)"
print_theorems
thm alpha_trm_raw_alpha_assn_raw_alpha_bn_raw.intros
thm bn_raw.simps
thm permute_bn_raw.simps
thm trm_assn.perm_bn_alpha
thm trm_assn.permute_bn
thm trm_assn.fv_defs
thm trm_assn.eq_iff
thm trm_assn.bn_defs
thm trm_assn.bn_inducts
thm trm_assn.perm_simps
thm trm_assn.permute_bn
thm trm_assn.induct
thm trm_assn.inducts
thm trm_assn.distinct
thm trm_assn.supp
thm trm_assn.fresh
thm trm_assn.exhaust
thm trm_assn.strong_exhaust
thm trm_assn.perm_bn_simps
end