diff -r 90bccef65004 -r aca2c7f80e2f src/HOL/Nominal/Examples/CR.thy --- a/src/HOL/Nominal/Examples/CR.thy Wed Jul 11 11:35:17 2007 +0200 +++ b/src/HOL/Nominal/Examples/CR.thy Wed Jul 11 11:36:06 2007 +0200 @@ -147,7 +147,7 @@ section {* Beta Reduction *} -inductive2 +inductive "Beta" :: "lam\lam\bool" (" _ \\<^isub>\ _" [80,80] 80) where b1[intro]: "s1\\<^isub>\s2 \ (App s1 t)\\<^isub>\(App s2 t)" @@ -160,7 +160,7 @@ nominal_inductive Beta by (simp_all add: abs_fresh fresh_fact') -inductive2 +inductive "Beta_star" :: "lam\lam\bool" (" _ \\<^isub>\\<^sup>* _" [80,80] 80) where bs1[intro, simp]: "M \\<^isub>\\<^sup>* M" @@ -177,7 +177,7 @@ section {* One-Reduction *} -inductive2 +inductive One :: "lam\lam\bool" (" _ \\<^isub>1 _" [80,80] 80) where o1[intro!]: "M\\<^isub>1M" @@ -190,7 +190,7 @@ nominal_inductive One by (simp_all add: abs_fresh fresh_fact') -inductive2 +inductive "One_star" :: "lam\lam\bool" (" _ \\<^isub>1\<^sup>* _" [80,80] 80) where os1[intro, simp]: "M \\<^isub>1\<^sup>* M" @@ -272,7 +272,7 @@ shows "\t''. t'=Lam [a].t'' \ t\\<^isub>1t''" using a apply - - apply(ind_cases2 "(Lam [a].t)\\<^isub>1t'") + apply(ind_cases "(Lam [a].t)\\<^isub>1t'") apply(auto simp add: lam.inject alpha) apply(rule_tac x="[(a,aa)]\s2" in exI) apply(rule conjI) @@ -289,7 +289,7 @@ (\a s s1 s2. t1 = Lam [a].s \ a\(t2,s2) \ t' = s1[a::=s2] \ s \\<^isub>1 s1 \ t2 \\<^isub>1 s2)" using a apply - - apply(ind_cases2 "App t1 t2 \\<^isub>1 t'") + apply(ind_cases "App t1 t2 \\<^isub>1 t'") apply(auto simp add: lam.distinct lam.inject) done @@ -299,7 +299,7 @@ (\s1 s2. M = s1[a::=s2] \ t1 \\<^isub>1 s1 \ t2 \\<^isub>1 s2)" using a apply - - apply(ind_cases2 "App (Lam [a].t1) t2 \\<^isub>1 M") + apply(ind_cases "App (Lam [a].t1) t2 \\<^isub>1 M") apply(simp_all add: lam.inject) apply(force) apply(erule conjE)