Nominal/Ex/TypeSchemes.thy
changeset 2486 b4ea19604b0b
parent 2480 ac7dff1194e8
child 2493 2e174807c891
--- a/Nominal/Ex/TypeSchemes.thy	Sat Sep 25 02:53:39 2010 +0200
+++ b/Nominal/Ex/TypeSchemes.thy	Sat Sep 25 08:28:45 2010 -0400
@@ -6,7 +6,7 @@
 
 atom_decl name
 
-declare [[STEPS = 100]]
+(* defined as a single nominal datatype *)
 
 nominal_datatype ty =
   Var "name"
@@ -26,6 +26,8 @@
 thm ty_tys.supports
 thm ty_tys.fsupp
 
+(* defined as two separate nominal datatypes *)
+
 nominal_datatype ty2 =
   Var2 "name"
 | Fun2 "ty2" "ty2"
@@ -50,8 +52,6 @@
 text {* *}
 
 (*
-ML {* Sign.of_sort @{theory} (@{typ ty}, @{sort fs}) *}
-
 lemma strong_induct:
   assumes a1: "\<And>name b. P b (Var name)"
   and     a2: "\<And>t1 t2 b. \<lbrakk>\<And>c. P c t1; \<And>c. P c t2\<rbrakk> \<Longrightarrow> P b (Fun t1 t2)"
@@ -74,7 +74,7 @@
                    s="pa \<bullet> (p \<bullet> All fset ty)" in subst)
     apply (rule supp_perm_eq)
     apply assumption
-    apply (simp only: ty_tys.perm)
+    apply (simp only: ty_tys.perm_simps)
     apply (rule a3)
     apply(erule_tac x="(pa + p)" in allE)
     apply simp
@@ -89,7 +89,8 @@
     apply (simp add: fresh_star_def)
     apply clarify
     apply (simp add: fresh_def)
-    apply (simp add: ty_tys_supp)
+    apply(auto)
+    apply (simp add: ty_tys.supp)
     done
   then have "P a (0 \<bullet> t) \<and> P' d (0 \<bullet> ts)" by blast
   then show ?thesis by simp