Nominal/Ex/TypeVarsTest.thy
author Christian Urban <urbanc@in.tum.de>
Fri, 27 Aug 2010 02:03:52 +0800
changeset 2438 abafea9b39bb
parent 2436 3885dc2669f9
child 2440 0a36825b16c1
permissions -rw-r--r--
corrected bug with fv-function generation (that was the problem with recursive binders)

theory TypeVarsTest
imports "../NewParser" 
begin

atom_decl name
declare [[STEPS = 21]]

class s1
class s2

nominal_datatype ('a, 'b) lam =
  Var "name"
| App "('a::s1, 'b::s2) lam" "('a, 'b) lam"
| Lam x::"name" l::"('a, 'b) lam"  bind x in l

thm lam.distinct
thm lam.induct
thm lam.exhaust
thm lam.fv_defs
thm lam.bn_defs
thm lam.perm_simps
thm lam.eq_iff
thm lam.fv_bn_eqvt
thm lam.size_eqvt


end