Nominal/Ex/TypeVarsTest.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 24 Nov 2010 16:59:26 +0900
changeset 2576 67828f23c4e9
parent 2571 f0252365936c
child 2617 e44551d067e6
permissions -rw-r--r--
Foo2 strong_exhaust for first variable.

theory TypeVarsTest
imports "../Nominal2" 
begin

atom_decl name

class s1
class s2

instance nat :: s1 ..
instance int :: 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