Nominal/Ex/TypeSchemes.thy
changeset 2950 0911cb7bf696
parent 2886 d7066575cbb9
child 2973 d1038e67923a
--- a/Nominal/Ex/TypeSchemes.thy	Tue Jul 05 18:01:54 2011 +0200
+++ b/Nominal/Ex/TypeSchemes.thy	Tue Jul 05 18:42:34 2011 +0200
@@ -12,7 +12,7 @@
   Var "name"
 | Fun "ty" "ty"
 and tys =
-  All xs::"name fset" ty::"ty" bind (set+) xs in ty
+  All xs::"name fset" ty::"ty" binds (set+) xs in ty
 
 ML {*
 get_all_info @{context}
@@ -282,7 +282,7 @@
 | Fun2 "ty2" "ty2"
 
 nominal_datatype tys2 =
-  All2 xs::"name fset" ty::"ty2" bind (set+) xs in ty
+  All2 xs::"name fset" ty::"ty2" binds (set+) xs in ty
 
 thm tys2.distinct
 thm tys2.induct tys2.strong_induct