# HG changeset patch # User urbanc # Date 1161684173 -7200 # Node ID 286d68ce3f5bb7b5a6b090dfb9dce7b0206424e2 # Parent cda93bbf35dbe88b486af116b95da5765e763df4 adapted to Stefan's new inductive package diff -r cda93bbf35db -r 286d68ce3f5b src/HOL/Nominal/Examples/CR.thy --- a/src/HOL/Nominal/Examples/CR.thy Mon Oct 23 17:46:11 2006 +0200 +++ b/src/HOL/Nominal/Examples/CR.thy Tue Oct 24 12:02:53 2006 +0200 @@ -27,16 +27,17 @@ have vc: "z\x" "z\P" by fact have ih: "x\M \ M[x::=P] = M" by fact have asm: "x\Lam [z].M" by fact - with vc have "x\M" by (simp add: fresh_atm abs_fresh) - hence "M[x::=P] = M" using ih by simp - thus "(Lam [z].M)[x::=P] = Lam [z].M" using vc by simp + then have "x\M" using vc by (simp add: fresh_atm abs_fresh) + then have "M[x::=P] = M" using ih by simp + then show "(Lam [z].M)[x::=P] = Lam [z].M" using vc by simp qed lemma forget_automatic: assumes asm: "x\L" shows "L[x::=P] = L" - using asm by (nominal_induct L avoiding: x P rule: lam.induct) - (auto simp add: abs_fresh fresh_atm) + using asm +by (nominal_induct L avoiding: x P rule: lam.induct) + (auto simp add: abs_fresh fresh_atm) lemma fresh_fact: fixes z::"name" @@ -71,8 +72,9 @@ fixes z::"name" assumes asms: "z\N" "z\L" shows "z\(N[y::=L])" -using asms by (nominal_induct N avoiding: z y L rule: lam.induct) - (auto simp add: abs_fresh fresh_atm) + using asms +by (nominal_induct N avoiding: z y L rule: lam.induct) + (auto simp add: abs_fresh fresh_atm) lemma substitution_lemma: assumes a: "x\y" @@ -125,74 +127,79 @@ qed next case (App M1 M2) (* case 3: applications *) - thus ?case by simp + thus "(App M1 M2)[x::=N][y::=L] = (App M1 M2)[y::=L][x::=N[y::=L]]" by simp qed lemma substitution_lemma_automatic: assumes asm: "x\y" "x\L" shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" -using asm by (nominal_induct M avoiding: x y N L rule: lam.induct) - (auto simp add: fresh_fact forget) + using asm +by (nominal_induct M avoiding: x y N L rule: lam.induct) + (auto simp add: fresh_fact forget) lemma subst_rename: - assumes a: "c\t1" - shows "t1[a::=t2] = ([(c,a)]\t1)[c::=t2]" + assumes a: "y\N" + shows "N[x::=L] = ([(y,x)]\N)[y::=L]" using a -proof (nominal_induct t1 avoiding: a c t2 rule: lam.induct) +proof (nominal_induct N avoiding: x y L rule: lam.induct) case (Var b) - thus "(Var b)[a::=t2] = ([(c,a)]\(Var b))[c::=t2]" by (simp add: calc_atm fresh_atm) + thus "(Var b)[x::=L] = ([(y,x)]\(Var b))[y::=L]" by (simp add: calc_atm fresh_atm) next case App thus ?case by force next - case (Lam b s) - have i: "\a c t2. c\s \ (s[a::=t2] = ([(c,a)]\s)[c::=t2])" by fact - have f: "b\a" "b\c" "b\t2" by fact - from f have a:"b\c" and b: "b\a" and c: "b\t2" by (simp add: fresh_atm)+ - have "c\Lam [b].s" by fact - hence "c\s" using a by (simp add: abs_fresh) - hence d: "s[a::=t2] = ([(c,a)]\s)[c::=t2]" using i by simp - show "(Lam [b].s)[a::=t2] = ([(c,a)]\(Lam [b].s))[c::=t2]" (is "?LHS = ?RHS") + case (Lam b N1) + have ih: "y\N1 \ (N1[x::=L] = ([(y,x)]\N1)[y::=L])" by fact + have f: "b\y" "b\x" "b\L" by fact + from f have a:"b\y" and b: "b\x" and c: "b\L" by (simp add: fresh_atm)+ + have "y\Lam [b].N1" by fact + hence "y\N1" using a by (simp add: abs_fresh) + hence d: "N1[x::=L] = ([(y,x)]\N1)[y::=L]" using ih by simp + show "(Lam [b].N1)[x::=L] = ([(y,x)]\(Lam [b].N1))[y::=L]" (is "?LHS = ?RHS") proof - - have "?LHS = Lam [b].(s[a::=t2])" using b c by simp - also have "\ = Lam [b].(([(c,a)]\s)[c::=t2])" using d by simp - also have "\ = (Lam [b].([(c,a)]\s))[c::=t2]" using a c by simp + have "?LHS = Lam [b].(N1[x::=L])" using b c by simp + also have "\ = Lam [b].(([(y,x)]\N1)[y::=L])" using d by simp + also have "\ = (Lam [b].([(y,x)]\N1))[y::=L]" using a c by simp also have "\ = ?RHS" using a b by (simp add: calc_atm) finally show "?LHS = ?RHS" by simp qed qed lemma subst_rename_automatic: - assumes a: "c\t1" - shows "t1[a::=t2] = ([(c,a)]\t1)[c::=t2]" -using a -apply(nominal_induct t1 avoiding: a c t2 rule: lam.induct) -apply(auto simp add: calc_atm fresh_atm abs_fresh) -done + assumes a: "y\N" + shows "N[x::=L] = ([(y,x)]\N)[y::=L]" + using a +by (nominal_induct N avoiding: y x L rule: lam.induct) + (auto simp add: calc_atm fresh_atm abs_fresh) section {* Beta Reduction *} -consts - Beta :: "(lam\lam) set" -syntax - "_Beta" :: "lam\lam\bool" (" _ \\<^isub>\ _" [80,80] 80) - "_Beta_star" :: "lam\lam\bool" (" _ \\<^isub>\\<^sup>* _" [80,80] 80) -translations - "t1 \\<^isub>\ t2" \ "(t1,t2) \ Beta" - "t1 \\<^isub>\\<^sup>* t2" \ "(t1,t2) \ Beta\<^sup>*" -inductive Beta - intros - b1[intro!]: "s1\\<^isub>\s2 \ (App s1 t)\\<^isub>\(App s2 t)" - b2[intro!]: "s1\\<^isub>\s2 \ (App t s1)\\<^isub>\(App t s2)" - b3[intro!]: "s1\\<^isub>\s2 \ (Lam [a].s1)\\<^isub>\ (Lam [a].s2)" - b4[intro!]: "(App (Lam [a].s1) s2)\\<^isub>\(s1[a::=s2])" +inductive2 + "Beta" :: "lam\lam\bool" (" _ \\<^isub>\ _" [80,80] 80) +intros + b1[intro]: "s1\\<^isub>\s2 \ (App s1 t)\\<^isub>\(App s2 t)" + b2[intro]: "s1\\<^isub>\s2 \ (App t s1)\\<^isub>\(App t s2)" + b3[intro]: "s1\\<^isub>\s2 \ (Lam [a].s1)\\<^isub>\ (Lam [a].s2)" + b4[intro]: "(App (Lam [a].s1) s2)\\<^isub>\(s1[a::=s2])" -lemma eqvt_beta: +inductive2 + "Beta_star" :: "lam\lam\bool" (" _ \\<^isub>\\<^sup>* _" [80,80] 80) +intros + bs1[intro, simp]: "M \\<^isub>\\<^sup>* M" + bs2[intro]: "\M1\\<^isub>\\<^sup>* M2; M2 \\<^isub>\ M3\ \ M1 \\<^isub>\\<^sup>* M3" + +lemma beta_star_trans: + assumes a1: "M1\\<^isub>\\<^sup>* M2" + and a2: "M2\\<^isub>\\<^sup>* M3" + shows "M1 \\<^isub>\\<^sup>* M3" +using a2 a1 +by (induct) (auto) + +lemma eqvt_beta: fixes pi :: "name prm" - and t :: "lam" - and s :: "lam" assumes a: "t\\<^isub>\s" shows "(pi\t)\\<^isub>\(pi\s)" - using a by (induct, auto) + using a +by (induct) (auto) lemma beta_induct[consumes 1, case_names b1 b2 b3 b4]: fixes P :: "'a::fs_name\lam \ lam \bool" @@ -200,9 +207,9 @@ and s :: "lam" and x :: "'a::fs_name" assumes a: "t\\<^isub>\s" - and a1: "\t s1 s2 x. s1\\<^isub>\s2 \ (\z. P z s1 s2) \ P x (App s1 t) (App s2 t)" - and a2: "\t s1 s2 x. s1\\<^isub>\s2 \ (\z. P z s1 s2) \ P x (App t s1) (App t s2)" - and a3: "\a s1 s2 x. a\x \ s1\\<^isub>\s2 \ (\z. P z s1 s2) \ P x (Lam [a].s1) (Lam [a].s2)" + and a1: "\t s1 s2 x. \s1\\<^isub>\s2; (\z. P z s1 s2)\ \ P x (App s1 t) (App s2 t)" + and a2: "\t s1 s2 x. \s1\\<^isub>\s2; (\z. P z s1 s2)\ \ P x (App t s1) (App t s2)" + and a3: "\a s1 s2 x. \a\x; s1\\<^isub>\s2; (\z. P z s1 s2)\ \ P x (Lam [a].s1) (Lam [a].s2)" and a4: "\a t1 s1 x. a\x \ P x (App (Lam [a].t1) s1) (t1[a::=s1])" shows "P x t s" proof - @@ -212,7 +219,7 @@ next case b2 thus ?case using a2 by (simp, blast intro: eqvt_beta) next - case (b3 a s1 s2) + case (b3 s1 s2 a) have j1: "s1 \\<^isub>\ s2" by fact have j2: "\x (pi::name prm). P x (pi\s1) (pi\s2)" by fact show ?case @@ -256,21 +263,20 @@ section {* One-Reduction *} -consts - One :: "(lam\lam) set" -syntax - "_One" :: "lam\lam\bool" (" _ \\<^isub>1 _" [80,80] 80) - "_One_star" :: "lam\lam\bool" (" _ \\<^isub>1\<^sup>* _" [80,80] 80) -translations - "t1 \\<^isub>1 t2" \ "(t1,t2) \ One" - "t1 \\<^isub>1\<^sup>* t2" \ "(t1,t2) \ One\<^sup>*" -inductive One +inductive2 + One :: "lam\lam\bool" (" _ \\<^isub>1 _" [80,80] 80) intros o1[intro!]: "M\\<^isub>1M" o2[simp,intro!]: "\t1\\<^isub>1t2;s1\\<^isub>1s2\ \ (App t1 s1)\\<^isub>1(App t2 s2)" o3[simp,intro!]: "s1\\<^isub>1s2 \ (Lam [a].s1)\\<^isub>1(Lam [a].s2)" o4[simp,intro!]: "\s1\\<^isub>1s2;t1\\<^isub>1t2\ \ (App (Lam [a].t1) s1)\\<^isub>1(t2[a::=s2])" +inductive2 + "One_star" :: "lam\lam\bool" (" _ \\<^isub>1\<^sup>* _" [80,80] 80) +intros + os1[intro, simp]: "M \\<^isub>1\<^sup>* M" + os2[intro]: "\M1\\<^isub>1\<^sup>* M2; M2 \\<^isub>1 M3\ \ M1 \\<^isub>1\<^sup>* M3" + lemma eqvt_one: fixes pi :: "name prm" and t :: "lam" @@ -279,6 +285,13 @@ shows "(pi\t)\\<^isub>1(pi\s)" using a by (induct, auto) +lemma one_star_trans: + assumes a1: "M1\\<^isub>1\<^sup>* M2" + and a2: "M2\\<^isub>1\<^sup>* M3" + shows "M1\\<^isub>1\<^sup>* M3" +using a2 a1 +by (induct) (auto) + lemma one_induct[consumes 1, case_names o1 o2 o3 o4]: fixes P :: "'a::fs_name\lam \ lam \bool" and t :: "lam" @@ -286,11 +299,11 @@ and x :: "'a::fs_name" assumes a: "t\\<^isub>1s" and a1: "\t x. P x t t" - and a2: "\t1 t2 s1 s2 x. t1\\<^isub>1t2 \ (\z. P z t1 t2) \ s1\\<^isub>1s2 \ (\z. P z s1 s2) \ + and a2: "\t1 t2 s1 s2 x. \t1\\<^isub>1t2; (\z. P z t1 t2); s1\\<^isub>1s2; (\z. P z s1 s2)\ \ P x (App t1 s1) (App t2 s2)" - and a3: "\a s1 s2 x. a\x \ s1\\<^isub>1s2 \ (\z. P z s1 s2) \ P x (Lam [a].s1) (Lam [a].s2)" + and a3: "\a s1 s2 x. \a\x; s1\\<^isub>1s2; (\z. P z s1 s2)\ \ P x (Lam [a].s1) (Lam [a].s2)" and a4: "\a t1 t2 s1 s2 x. - a\x \ t1\\<^isub>1t2 \ (\z. P z t1 t2) \ s1\\<^isub>1s2 \ (\z. P z s1 s2) + \a\x; t1\\<^isub>1t2; (\z. P z t1 t2); s1\\<^isub>1s2; (\z. P z s1 s2)\ \ P x (App (Lam [a].t1) s1) (t2[a::=s2])" shows "P x t s" proof - @@ -301,7 +314,7 @@ case (o2 s1 s2 t1 t2) thus ?case using a2 by (simp, blast intro: eqvt_one) next - case (o3 a t1 t2) + case (o3 t1 t2 a) have j1: "t1 \\<^isub>1 t2" by fact have j2: "\(pi::name prm) x. P x (pi\t1) (pi\t2)" by fact show ?case @@ -321,7 +334,7 @@ using x alpha1 alpha2 by (simp only: pt_name2) qed next - case (o4 a s1 s2 t1 t2) + case (o4 s1 s2 t1 t2 a) have j0: "t1 \\<^isub>1 t2" by fact have j1: "s1 \\<^isub>1 s2" by fact have j2: "\(pi::name prm) x. P x (pi\t1) (pi\t2)" by fact @@ -375,7 +388,7 @@ next case o2 thus ?case by simp next - case (o3 c s1 s2) + case (o3 s1 s2 c) have ih: "a\s1 \ a\s2" by fact have c: "a\Lam [c].s1" by fact show ?case @@ -388,7 +401,7 @@ thus "a\Lam [c].s2" using d by (simp add: abs_fresh) qed next - case (o4 c t1 t2 s1 s2) + case (o4 t1 t2 s1 s2 c) have i1: "a\t1 \ a\t2" by fact have i2: "a\s1 \ a\s2" by fact have as: "a\App (Lam [c].s1) t1" by fact @@ -410,8 +423,11 @@ fixes t :: "lam" and t':: "lam" and a :: "name" - shows "(Lam [a].t)\\<^isub>1t'\\t''. t'=Lam [a].t'' \ t\\<^isub>1t''" - apply(ind_cases "(Lam [a].t)\\<^isub>1t'") + assumes a: "(Lam [a].t)\\<^isub>1t'" + shows "\t''. t'=Lam [a].t'' \ t\\<^isub>1t''" + using a + apply - + apply(ind_cases2 "(Lam [a].t)\\<^isub>1t'") apply(auto simp add: lam.distinct lam.inject alpha) apply(rule_tac x="[(a,aa)]\s2" in exI) apply(rule conjI) @@ -428,18 +444,22 @@ done lemma one_app: - "App t1 t2 \\<^isub>1 t' \ - (\s1 s2. t' = App s1 s2 \ t1 \\<^isub>1 s1 \ t2 \\<^isub>1 s2) \ - (\a s s1 s2. t1 = Lam [a].s \ t' = s1[a::=s2] \ s \\<^isub>1 s1 \ t2 \\<^isub>1 s2)" - apply(ind_cases "App t1 s1 \\<^isub>1 t'") + assumes a: "App t1 t2 \\<^isub>1 t'" + shows "(\s1 s2. t' = App s1 s2 \ t1 \\<^isub>1 s1 \ t2 \\<^isub>1 s2) \ + (\a s s1 s2. t1 = Lam [a].s \ t' = s1[a::=s2] \ s \\<^isub>1 s1 \ t2 \\<^isub>1 s2)" + using a + apply - + apply(ind_cases2 "App t1 s1 \\<^isub>1 t'") apply(auto simp add: lam.distinct lam.inject) done lemma one_red: - "App (Lam [a].t1) t2 \\<^isub>1 M \ - (\s1 s2. M = App (Lam [a].s1) s2 \ t1 \\<^isub>1 s1 \ t2 \\<^isub>1 s2) \ - (\s1 s2. M = s1[a::=s2] \ t1 \\<^isub>1 s1 \ t2 \\<^isub>1 s2)" - apply(ind_cases "App (Lam [a].t1) s1 \\<^isub>1 M") + assumes a: "App (Lam [a].t1) t2 \\<^isub>1 M" + shows "(\s1 s2. M = App (Lam [a].s1) s2 \ t1 \\<^isub>1 s1 \ t2 \\<^isub>1 s2) \ + (\s1 s2. M = s1[a::=s2] \ t1 \\<^isub>1 s1 \ t2 \\<^isub>1 s2)" + using a + apply - + apply(ind_cases2 "App (Lam [a].t1) s1 \\<^isub>1 M") apply(simp_all add: lam.inject) apply(force) apply(erule conjE) @@ -539,7 +559,7 @@ case (o1 M) (* case 1 --- M1 = M *) thus "\M3. M\\<^isub>1M3 \ M2\\<^isub>1M3" by blast next - case (o4 x Q Q' P P') (* case 2 --- a beta-reduction occurs*) + case (o4 Q Q' P P' x) (* case 2 --- a beta-reduction occurs*) have i1: "\M2. Q \\<^isub>1M2 \ (\M3. Q'\\<^isub>1M3 \ M2\\<^isub>1M3)" by fact have i2: "\M2. P \\<^isub>1M2 \ (\M3. P'\\<^isub>1M3 \ M2\\<^isub>1M3)" by fact have "App (Lam [x].P) Q \\<^isub>1 M2" by fact @@ -577,7 +597,7 @@ } ultimately show "\M3. P'[x::=Q']\\<^isub>1M3 \ M2\\<^isub>1M3" by blast next - case (o2 Q Q' P P') (* case 3 *) + case (o2 P P' Q Q') (* case 3 *) have i0: "P\\<^isub>1P'" by fact have i1: "\M2. Q \\<^isub>1M2 \ (\M3. Q'\\<^isub>1M3 \ M2\\<^isub>1M3)" by fact have i2: "\M2. P \\<^isub>1M2 \ (\M3. P'\\<^isub>1M3 \ M2\\<^isub>1M3)" by fact @@ -623,7 +643,7 @@ } ultimately show "\M3. App P' Q'\\<^isub>1M3 \ M2\\<^isub>1M3" by blast next - case (o3 x P P') (* case 4 *) + case (o3 P P' x) (* case 4 *) have i1: "P\\<^isub>1P'" by fact have i2: "\M2. P \\<^isub>1M2 \ (\M3. P'\\<^isub>1M3 \ M2\\<^isub>1M3)" by fact have "(Lam [x].P)\\<^isub>1 M2" by fact @@ -646,10 +666,10 @@ shows "(Lam [a].t1)\\<^isub>\\<^sup>*(Lam [a].t2)" using a proof induct - case 1 thus ?case by simp + case bs1 thus ?case by simp next - case (2 y z) - thus ?case by (blast dest: b3 intro: rtrancl_trans) + case (bs2 y z) + thus ?case by (blast dest: b3) qed lemma one_app_congL: @@ -657,9 +677,9 @@ shows "App t1 s\\<^isub>\\<^sup>* App t2 s" using a proof induct - case 1 thus ?case by simp + case bs1 thus ?case by simp next - case 2 thus ?case by (blast dest: b1 intro: rtrancl_trans) + case bs2 thus ?case by (blast dest: b1) qed lemma one_app_congR: @@ -667,20 +687,20 @@ shows "App s t1 \\<^isub>\\<^sup>* App s t2" using a proof induct - case 1 thus ?case by simp + case bs1 thus ?case by simp next - case 2 thus ?case by (blast dest: b2 intro: rtrancl_trans) + case bs2 thus ?case by (blast dest: b2) qed lemma one_app_cong: assumes a1: "t1\\<^isub>\\<^sup>*t2" - and a2: "s1\\<^isub>\\<^sup>*s2" + and a2: "s1\\<^isub>\\<^sup>*s2" shows "App t1 s1\\<^isub>\\<^sup>* App t2 s2" proof - have "App t1 s1 \\<^isub>\\<^sup>* App t2 s1" using a1 by (rule one_app_congL) moreover have "App t2 s1 \\<^isub>\\<^sup>* App t2 s2" using a2 by (rule one_app_congR) - ultimately show ?thesis by (blast intro: rtrancl_trans) + ultimately show ?thesis by (rule beta_star_trans) qed lemma one_beta_star: @@ -694,12 +714,12 @@ next case o3 thus ?case by (blast intro!: one_lam_cong) next - case (o4 a s1 s2 t1 t2) + case (o4 s1 s2 t1 t2 a) have a1: "t1\\<^isub>\\<^sup>*t2" and a2: "s1\\<^isub>\\<^sup>*s2" by fact have c1: "(App (Lam [a].t2) s2) \\<^isub>\ (t2 [a::= s2])" by (rule b4) from a1 a2 have c2: "App (Lam [a].t1 ) s1 \\<^isub>\\<^sup>* App (Lam [a].t2 ) s2" by (blast intro!: one_app_cong one_lam_cong) - show ?case using c1 c2 by (blast intro: rtrancl_trans) + show ?case using c2 c1 by (blast intro: beta_star_trans) qed lemma one_star_lam_cong: @@ -707,9 +727,9 @@ shows "(Lam [a].t1)\\<^isub>1\<^sup>* (Lam [a].t2)" using a proof induct - case 1 thus ?case by simp + case os1 thus ?case by simp next - case 2 thus ?case by (blast intro: rtrancl_trans) + case os2 thus ?case by (blast intro: one_star_trans) qed lemma one_star_app_congL: @@ -717,9 +737,9 @@ shows "App t1 s\\<^isub>1\<^sup>* App t2 s" using a proof induct - case 1 thus ?case by simp + case os1 thus ?case by simp next - case 2 thus ?case by (blast intro: rtrancl_trans) + case os2 thus ?case by (blast intro: one_star_trans) qed lemma one_star_app_congR: @@ -727,9 +747,9 @@ shows "App s t1 \\<^isub>1\<^sup>* App s t2" using a proof induct - case 1 thus ?case by simp + case os1 thus ?case by simp next - case 2 thus ?case by (blast intro: rtrancl_trans) + case os2 thus ?case by (blast intro: one_star_trans) qed lemma beta_one_star: @@ -747,22 +767,30 @@ qed lemma trans_closure: - shows "(t1\\<^isub>1\<^sup>*t2) = (t1\\<^isub>\\<^sup>*t2)" + shows "(M1\\<^isub>1\<^sup>*M2) = (M1\\<^isub>\\<^sup>*M2)" proof - assume "t1 \\<^isub>1\<^sup>* t2" - then show "t1\\<^isub>\\<^sup>*t2" + assume "M1 \\<^isub>1\<^sup>* M2" + then show "M1\\<^isub>\\<^sup>*M2" proof induct - case 1 thus ?case by simp + case (os1 M1) thus "M1\\<^isub>\\<^sup>*M1" by simp next - case 2 thus ?case by (force intro: rtrancl_trans simp add: one_beta_star) + case (os2 M1 M2 M3) + have "M2\\<^isub>1M3" by fact + then have "M2\\<^isub>\\<^sup>*M3" by (rule one_beta_star) + moreover have "M1\\<^isub>\\<^sup>*M2" by fact + ultimately show "M1\\<^isub>\\<^sup>*M3" by (auto intro: beta_star_trans) qed next - assume "t1 \\<^isub>\\<^sup>* t2" - then show "t1\\<^isub>1\<^sup>*t2" + assume "M1 \\<^isub>\\<^sup>* M2" + then show "M1\\<^isub>1\<^sup>*M2" proof induct - case 1 thus ?case by simp + case (bs1 M1) thus "M1\\<^isub>1\<^sup>*M1" by simp next - case 2 thus ?case by (force intro: rtrancl_trans simp add: beta_one_star) + case (bs2 M1 M2 M3) + have "M2\\<^isub>\M3" by fact + then have "M2\\<^isub>1\<^sup>*M3" by (rule beta_one_star) + moreover have "M1\\<^isub>1\<^sup>*M2" by fact + ultimately show "M1\\<^isub>1\<^sup>*M3" by (auto intro: one_star_trans) qed qed @@ -772,9 +800,9 @@ shows "\t3. t1\\<^isub>1t3 \ t2\\<^isub>1\<^sup>*t3" using a b proof (induct arbitrary: t2) - case 1 thus ?case by force + case os1 thus ?case by force next - case (2 s1 s2) + case (os2 t s1 s2 t2) have b: "s1 \\<^isub>1 s2" by fact have h: "\t2. t \\<^isub>1 t2 \ (\t3. s1 \\<^isub>1 t3 \ t2 \\<^isub>1\<^sup>* t3)" by fact have c: "t \\<^isub>1 t2" by fact @@ -783,7 +811,7 @@ from c h have "\t3. s1 \\<^isub>1 t3 \ t2 \\<^isub>1\<^sup>* t3" by blast then obtain t3 where c1: "s1 \\<^isub>1 t3" and c2: "t2 \\<^isub>1\<^sup>* t3" by blast have "\t4. s2 \\<^isub>1 t4 \ t3 \\<^isub>1 t4" using b c1 by (blast intro: diamond) - thus ?thesis using c2 by (blast intro: rtrancl_trans) + thus ?thesis using c2 by (blast intro: one_star_trans) qed qed @@ -791,20 +819,21 @@ assumes a: "t\\<^isub>1\<^sup>*t2" and b: "t\\<^isub>1\<^sup>*t1" shows "\t3. t1\\<^isub>1\<^sup>*t3\t2\\<^isub>1\<^sup>*t3" -using a -proof induct - case 1 - show ?case using b by force +using a b +proof (induct arbitrary: t1) + case (os1 t) then show ?case by force next - case (2 s1 s2) + case (os2 t s1 s2 t1) + have c: "t \\<^isub>1\<^sup>* s1" by fact + have c': "t \\<^isub>1\<^sup>* t1" by fact have d: "s1 \\<^isub>1 s2" by fact - have "\t3. t1 \\<^isub>1\<^sup>* t3 \ s1 \\<^isub>1\<^sup>* t3" by fact + have "t \\<^isub>1\<^sup>* t1 \ (\t3. t1 \\<^isub>1\<^sup>* t3 \ s1 \\<^isub>1\<^sup>* t3)" by fact then obtain t3 where f1: "t1 \\<^isub>1\<^sup>* t3" - and f2: "s1 \\<^isub>1\<^sup>* t3" by blast + and f2: "s1 \\<^isub>1\<^sup>* t3" using c' by blast from cr_one d f2 have "\t4. t3\\<^isub>1t4 \ s2\\<^isub>1\<^sup>*t4" by blast then obtain t4 where g1: "t3\\<^isub>1t4" and g2: "s2\\<^isub>1\<^sup>*t4" by blast - have "t1\\<^isub>1\<^sup>*t4" using f1 g1 by (blast intro: rtrancl_trans) + have "t1\\<^isub>1\<^sup>*t4" using f1 g1 by (blast intro: one_star_trans) thus ?case using g2 by blast qed