# HG changeset patch # User berghofe # Date 1170866703 -3600 # Node ID aac2ac7c32fdffbd514f45d0f5e1592acc868e5f # Parent 51a80e238b297337d50691b81af4ef9447ab49fd - Adapted to new inductive definition package - More elegant proof of eta postponement theorem inspired by Andreas Abel diff -r 51a80e238b29 -r aac2ac7c32fd src/HOL/Lambda/Eta.thy --- a/src/HOL/Lambda/Eta.thy Wed Feb 07 17:44:07 2007 +0100 +++ b/src/HOL/Lambda/Eta.thy Wed Feb 07 17:45:03 2007 +0100 @@ -18,32 +18,29 @@ "free (s \ t) i = (free s i \ free t i)" "free (Abs s) i = free s (i + 1)" -consts - eta :: "(dB \ dB) set" - -abbreviation - eta_red :: "[dB, dB] => bool" (infixl "-e>" 50) where - "s -e> t == (s, t) \ eta" +inductive2 eta :: "[dB, dB] => bool" (infixl "\\<^sub>\" 50) + where + eta [simp, intro]: "\ free s 0 ==> Abs (s \ Var 0) \\<^sub>\ s[dummy/0]" + | appL [simp, intro]: "s \\<^sub>\ t ==> s \ u \\<^sub>\ t \ u" + | appR [simp, intro]: "s \\<^sub>\ t ==> u \ s \\<^sub>\ u \ t" + | abs [simp, intro]: "s \\<^sub>\ t ==> Abs s \\<^sub>\ Abs t" abbreviation eta_reds :: "[dB, dB] => bool" (infixl "-e>>" 50) where - "s -e>> t == (s, t) \ eta^*" + "s -e>> t == eta^** s t" abbreviation eta_red0 :: "[dB, dB] => bool" (infixl "-e>=" 50) where - "s -e>= t == (s, t) \ eta^=" + "s -e>= t == eta^== s t" -inductive eta - intros - eta [simp, intro]: "\ free s 0 ==> Abs (s \ Var 0) -e> s[dummy/0]" - appL [simp, intro]: "s -e> t ==> s \ u -e> t \ u" - appR [simp, intro]: "s -e> t ==> u \ s -e> u \ t" - abs [simp, intro]: "s -e> t ==> Abs s -e> Abs t" +notation (xsymbols) + eta_reds (infixl "\\<^sub>\\<^sup>*" 50) and + eta_red0 (infixl "\\<^sub>\\<^sup>=" 50) -inductive_cases eta_cases [elim!]: - "Abs s -e> z" - "s \ t -e> u" - "Var i -e> t" +inductive_cases2 eta_cases [elim!]: + "Abs s \\<^sub>\ z" + "s \ t \\<^sub>\ u" + "Var i \\<^sub>\ t" subsection "Properties of eta, subst and free" @@ -69,15 +66,15 @@ apply (simp add: diff_Suc subst_Var split: nat.split) done -lemma free_eta: "s -e> t ==> free t i = free s i" +lemma free_eta: "s \\<^sub>\ t ==> free t i = free s i" by (induct arbitrary: i set: eta) (simp_all cong: conj_cong) lemma not_free_eta: - "[| s -e> t; \ free s i |] ==> \ free t i" + "[| s \\<^sub>\ t; \ free s i |] ==> \ free t i" by (simp add: free_eta) lemma eta_subst [simp]: - "s -e> t ==> s[u/i] -e> t[u/i]" + "s \\<^sub>\ t ==> s[u/i] \\<^sub>\ t[u/i]" by (induct arbitrary: u i set: eta) (simp_all add: subst_subst [symmetric]) theorem lift_subst_dummy: "\ free s i \ lift (s[dummy/i]) i = s" @@ -86,7 +83,7 @@ subsection "Confluence of eta" -lemma square_eta: "square eta eta (eta^=) (eta^=)" +lemma square_eta: "square eta eta (eta^==) (eta^==)" apply (unfold square_def id_def) apply (rule impI [THEN allI [THEN allI]]) apply simp @@ -105,38 +102,38 @@ subsection "Congruence rules for eta*" -lemma rtrancl_eta_Abs: "s -e>> s' ==> Abs s -e>> Abs s'" +lemma rtrancl_eta_Abs: "s \\<^sub>\\<^sup>* s' ==> Abs s \\<^sub>\\<^sup>* Abs s'" by (induct set: rtrancl) - (blast intro: rtrancl_refl rtrancl_into_rtrancl)+ + (blast intro: rtrancl.rtrancl_into_rtrancl)+ -lemma rtrancl_eta_AppL: "s -e>> s' ==> s \ t -e>> s' \ t" +lemma rtrancl_eta_AppL: "s \\<^sub>\\<^sup>* s' ==> s \ t \\<^sub>\\<^sup>* s' \ t" by (induct set: rtrancl) - (blast intro: rtrancl_refl rtrancl_into_rtrancl)+ + (blast intro: rtrancl.rtrancl_into_rtrancl)+ -lemma rtrancl_eta_AppR: "t -e>> t' ==> s \ t -e>> s \ t'" - by (induct set: rtrancl) (blast intro: rtrancl_refl rtrancl_into_rtrancl)+ +lemma rtrancl_eta_AppR: "t \\<^sub>\\<^sup>* t' ==> s \ t \\<^sub>\\<^sup>* s \ t'" + by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ lemma rtrancl_eta_App: - "[| s -e>> s'; t -e>> t' |] ==> s \ t -e>> s' \ t'" - by (blast intro!: rtrancl_eta_AppL rtrancl_eta_AppR intro: rtrancl_trans) + "[| s \\<^sub>\\<^sup>* s'; t \\<^sub>\\<^sup>* t' |] ==> s \ t \\<^sub>\\<^sup>* s' \ t'" + by (blast intro!: rtrancl_eta_AppL rtrancl_eta_AppR intro: rtrancl_trans') subsection "Commutation of beta and eta" lemma free_beta: - "s -> t ==> free t i \ free s i" + "s \\<^sub>\ t ==> free t i \ free s i" by (induct arbitrary: i set: beta) auto -lemma beta_subst [intro]: "s -> t ==> s[u/i] -> t[u/i]" +lemma beta_subst [intro]: "s \\<^sub>\ t ==> s[u/i] \\<^sub>\ t[u/i]" by (induct arbitrary: u i set: beta) (simp_all add: subst_subst [symmetric]) lemma subst_Var_Suc [simp]: "t[Var i/i] = t[Var(i)/i + 1]" by (induct t arbitrary: i) (auto elim!: linorder_neqE simp: subst_Var) -lemma eta_lift [simp]: "s -e> t ==> lift s i -e> lift t i" +lemma eta_lift [simp]: "s \\<^sub>\ t ==> lift s i \\<^sub>\ lift t i" by (induct arbitrary: i set: eta) simp_all -lemma rtrancl_eta_subst: "s -e> t \ u[s/i] -e>> u[t/i]" +lemma rtrancl_eta_subst: "s \\<^sub>\ t \ u[s/i] \\<^sub>\\<^sup>* u[t/i]" apply (induct u arbitrary: s t i) apply (simp_all add: subst_Var) apply blast @@ -144,7 +141,17 @@ apply (blast intro!: rtrancl_eta_Abs eta_lift) done -lemma square_beta_eta: "square beta eta (eta^*) (beta^=)" +lemma rtrancl_eta_subst': + assumes eta: "s \\<^sub>\\<^sup>* t" + shows "s[u/i] \\<^sub>\\<^sup>* t[u/i]" using eta + by induct (iprover intro: eta_subst)+ + +lemma rtrancl_eta_subst'': + assumes eta: "s \\<^sub>\\<^sup>* t" + shows "u[s/i] \\<^sub>\\<^sup>* u[t/i]" using eta + by induct (iprover intro: rtrancl_eta_subst rtrancl_trans')+ + +lemma square_beta_eta: "square beta eta (eta^**) (beta^==)" apply (unfold square_def) apply (rule impI [THEN allI [THEN allI]]) apply (erule beta.induct) @@ -156,7 +163,7 @@ iff del: dB.distinct simp: dB.distinct) (*23 seconds?*) done -lemma confluent_beta_eta: "confluent (beta \ eta)" +lemma confluent_beta_eta: "confluent (join beta eta)" apply (assumption | rule square_rtrancl_reflcl_commute confluent_Un beta_confluent eta_confluent square_beta_eta)+ @@ -165,7 +172,7 @@ subsection "Implicit definition of eta" -text {* @{term "Abs (lift s 0 \ Var 0) -e> s"} *} +text {* @{term "Abs (lift s 0 \ Var 0) \\<^sub>\ s"} *} lemma not_free_iff_lifted: "(\ free s i) = (\t. s = lift t i)" @@ -222,262 +229,163 @@ by (auto simp add: not_free_iff_lifted) -subsection {* Parallel eta-reduction *} - -consts - par_eta :: "(dB \ dB) set" - -abbreviation - par_eta_red :: "[dB, dB] => bool" (infixl "=e>" 50) where - "s =e> t == (s, t) \ par_eta" - -notation (xsymbols) - par_eta_red (infixl "\\<^sub>\" 50) - -inductive par_eta -intros - var [simp, intro]: "Var x \\<^sub>\ Var x" - eta [simp, intro]: "\ free s 0 \ s \\<^sub>\ s'\ Abs (s \ Var 0) \\<^sub>\ s'[dummy/0]" - app [simp, intro]: "s \\<^sub>\ s' \ t \\<^sub>\ t' \ s \ t \\<^sub>\ s' \ t'" - abs [simp, intro]: "s \\<^sub>\ t \ Abs s \\<^sub>\ Abs t" - -lemma free_par_eta [simp]: - assumes eta: "s \\<^sub>\ t" - shows "free t i = free s i" using eta - by (induct arbitrary: i) (simp_all cong: conj_cong) - -lemma par_eta_refl [simp]: "t \\<^sub>\ t" - by (induct t) simp_all - -lemma par_eta_lift [simp]: - assumes eta: "s \\<^sub>\ t" - shows "lift s i \\<^sub>\ lift t i" using eta - by (induct arbitrary: i) simp_all - -lemma par_eta_subst [simp]: - assumes eta: "s \\<^sub>\ t" - shows "u \\<^sub>\ u' \ s[u/i] \\<^sub>\ t[u'/i]" using eta - by (induct arbitrary: u u' i) (simp_all add: subst_subst [symmetric] subst_Var) - -theorem eta_subset_par_eta: "eta \ par_eta" - apply (rule subsetI) - apply clarify - apply (erule eta.induct) - apply (blast intro!: par_eta_refl)+ - done - -theorem par_eta_subset_eta: "par_eta \ eta\<^sup>*" - apply (rule subsetI) - apply clarify - apply (erule par_eta.induct) - apply blast - apply (rule rtrancl_into_rtrancl) - apply (rule rtrancl_eta_Abs) - apply (rule rtrancl_eta_AppL) - apply assumption - apply (rule eta.eta) - apply simp - apply (rule rtrancl_eta_App) - apply assumption+ - apply (rule rtrancl_eta_Abs) - apply assumption - done - - -subsection {* n-ary eta-expansion *} - -consts eta_expand :: "nat \ dB \ dB" -primrec - eta_expand_0: "eta_expand 0 t = t" - eta_expand_Suc: "eta_expand (Suc n) t = Abs (lift (eta_expand n t) 0 \ Var 0)" - -lemma eta_expand_Suc': - "eta_expand (Suc n) t = eta_expand n (Abs (lift t 0 \ Var 0))" - by (induct n arbitrary: t) simp_all - -theorem lift_eta_expand: "lift (eta_expand k t) i = eta_expand k (lift t i)" - by (induct k) (simp_all add: lift_lift) - -theorem eta_expand_beta: - assumes u: "u => u'" - shows "t => t' \ eta_expand k (Abs t) \ u => t'[u'/0]" -proof (induct k arbitrary: t t') - case 0 - with u show ?case by simp -next - case (Suc k) - hence "Abs (lift t (Suc 0)) \ Var 0 => lift t' (Suc 0)[Var 0/0]" - by (blast intro: par_beta_lift) - with Suc show ?case by (simp del: eta_expand_Suc add: eta_expand_Suc') -qed - -theorem eta_expand_red: - assumes t: "t => t'" - shows "eta_expand k t => eta_expand k t'" - by (induct k) (simp_all add: t) - -theorem eta_expand_eta: "t \\<^sub>\ t' \ eta_expand k t \\<^sub>\ t'" -proof (induct k arbitrary: t t') - case 0 - show ?case by simp -next - case (Suc k) - have "Abs (lift (eta_expand k t) 0 \ Var 0) \\<^sub>\ lift t' 0[arbitrary/0]" - by (fastsimp intro: par_eta_lift Suc) - thus ?case by simp -qed - - -subsection {* Elimination rules for parallel eta reduction *} - -theorem par_eta_elim_app: assumes eta: "t \\<^sub>\ u" - shows "u = u1' \ u2' \ - \u1 u2 k. t = eta_expand k (u1 \ u2) \ u1 \\<^sub>\ u1' \ u2 \\<^sub>\ u2'" using eta -proof (induct arbitrary: u1' u2') - case (app s s' t t') - have "s \ t = eta_expand 0 (s \ t)" by simp - with app show ?case by blast -next - case (eta dummy s s') - then obtain u1'' u2'' where s': "s' = u1'' \ u2''" - by (cases s') (auto simp add: subst_Var free_par_eta [symmetric] split: split_if_asm) - then have "\u1 u2 k. s = eta_expand k (u1 \ u2) \ u1 \\<^sub>\ u1'' \ u2 \\<^sub>\ u2''" by (rule eta) - then obtain u1 u2 k where s: "s = eta_expand k (u1 \ u2)" - and u1: "u1 \\<^sub>\ u1''" and u2: "u2 \\<^sub>\ u2''" by iprover - from u1 u2 eta s' have "\ free u1 0" and "\ free u2 0" - by (simp_all del: free_par_eta add: free_par_eta [symmetric]) - with s have "Abs (s \ Var 0) = eta_expand (Suc k) (u1[dummy/0] \ u2[dummy/0])" - by (simp del: lift_subst add: lift_subst_dummy lift_eta_expand) - moreover from u1 par_eta_refl have "u1[dummy/0] \\<^sub>\ u1''[dummy/0]" - by (rule par_eta_subst) - moreover from u2 par_eta_refl have "u2[dummy/0] \\<^sub>\ u2''[dummy/0]" - by (rule par_eta_subst) - ultimately show ?case using eta s' - by (simp only: subst.simps dB.simps) blast -next - case var thus ?case by simp -next - case abs thus ?case by simp -qed - -theorem par_eta_elim_abs: assumes eta: "t \\<^sub>\ t'" - shows "t' = Abs u' \ - \u k. t = eta_expand k (Abs u) \ u \\<^sub>\ u'" using eta -proof (induct arbitrary: u') - case (abs s t) - have "Abs s = eta_expand 0 (Abs s)" by simp - with abs show ?case by blast -next - case (eta dummy s s') - then obtain u'' where s': "s' = Abs u''" - by (cases s') (auto simp add: subst_Var free_par_eta [symmetric] split: split_if_asm) - then have "\u k. s = eta_expand k (Abs u) \ u \\<^sub>\ u''" by (rule eta) - then obtain u k where s: "s = eta_expand k (Abs u)" and u: "u \\<^sub>\ u''" by iprover - from eta u s' have "\ free u (Suc 0)" - by (simp del: free_par_eta add: free_par_eta [symmetric]) - with s have "Abs (s \ Var 0) = eta_expand (Suc k) (Abs (u[lift dummy 0/Suc 0]))" - by (simp del: lift_subst add: lift_eta_expand lift_subst_dummy) - moreover from u par_eta_refl have "u[lift dummy 0/Suc 0] \\<^sub>\ u''[lift dummy 0/Suc 0]" - by (rule par_eta_subst) - ultimately show ?case using eta s' by fastsimp -next - case var thus ?case by simp -next - case app thus ?case by simp -qed - - subsection {* Eta-postponement theorem *} text {* - Based on a proof by Masako Takahashi \cite{Takahashi-IandC}. + Based on a paper proof due to Andreas Abel. + Unlike the proof by Masako Takahashi \cite{Takahashi-IandC}, it does not + use parallel eta reduction, which only seems to complicate matters unnecessarily. *} -theorem par_eta_beta: "s \\<^sub>\ t \ t => u \ \t'. s => t' \ t' \\<^sub>\ u" -proof (induct t arbitrary: s u taking: "size :: dB \ nat" rule: measure_induct_rule) - case (less t) - from `t => u` - show ?case - proof cases - case (var n) - with less show ?thesis - by (auto intro: par_beta_refl) - next - case (abs r' r'') - with less have "s \\<^sub>\ Abs r'" by simp - then obtain r k where s: "s = eta_expand k (Abs r)" and rr: "r \\<^sub>\ r'" - by (blast dest: par_eta_elim_abs) - from abs have "size r' < size t" by simp - with abs(2) rr obtain t' where rt: "r => t'" and t': "t' \\<^sub>\ r''" - by (blast dest: less(1)) - with s abs show ?thesis - by (auto intro: eta_expand_red eta_expand_eta) - next - case (app q' q'' r' r'') - with less have "s \\<^sub>\ q' \ r'" by simp - then obtain q r k where s: "s = eta_expand k (q \ r)" - and qq: "q \\<^sub>\ q'" and rr: "r \\<^sub>\ r'" - by (blast dest: par_eta_elim_app [OF _ refl]) - from app have "size q' < size t" and "size r' < size t" by simp_all - with app(2,3) qq rr obtain t' t'' where "q => t'" and - "t' \\<^sub>\ q''" and "r => t''" and "t'' \\<^sub>\ r''" - by (blast dest: less(1)) - with s app show ?thesis - by (fastsimp intro: eta_expand_red eta_expand_eta) - next - case (beta q' q'' r' r'') - with less have "s \\<^sub>\ Abs q' \ r'" by simp - then obtain q r k k' where s: "s = eta_expand k (eta_expand k' (Abs q) \ r)" - and qq: "q \\<^sub>\ q'" and rr: "r \\<^sub>\ r'" - by (blast dest: par_eta_elim_app par_eta_elim_abs) - from beta have "size q' < size t" and "size r' < size t" by simp_all - with beta(2-3) qq rr obtain t' t'' where "q => t'" and - "t' \\<^sub>\ q''" and "r => t''" and "t'' \\<^sub>\ r''" - by (blast dest: less(1)) - with s beta show ?thesis - by (auto intro: eta_expand_red eta_expand_beta eta_expand_eta par_eta_subst) - qed +theorem eta_case: + assumes free: "\ free s 0" + and s: "s[dummy/0] => u" + shows "\t'. Abs (s \ Var 0) => t' \ t' \\<^sub>\\<^sup>* u" +proof - + from s have "lift (s[dummy/0]) 0 => lift u 0" by (simp del: lift_subst) + with free have "s => lift u 0" by (simp add: lift_subst_dummy del: lift_subst) + hence "Abs (s \ Var 0) => Abs (lift u 0 \ Var 0)" by simp + moreover have "\ free (lift u 0) 0" by simp + hence "Abs (lift u 0 \ Var 0) \\<^sub>\ lift u 0[dummy/0]" + by (rule eta.eta) + hence "Abs (lift u 0 \ Var 0) \\<^sub>\\<^sup>* u" by simp + ultimately show ?thesis by iprover qed -theorem eta_postponement': assumes eta: "s -e>> t" - shows "t => u \ \t'. s => t' \ t' -e>> u" - using eta [simplified rtrancl_subset - [OF eta_subset_par_eta par_eta_subset_eta, symmetric]] +theorem eta_par_beta: + assumes st: "s \\<^sub>\ t" + and tu: "t => u" + shows "\t'. s => t' \ t' \\<^sub>\\<^sup>* u" using tu st +proof (induct arbitrary: s) + case (var n) + thus ?case by (iprover intro: par_beta_refl) +next + case (abs s' t) + note abs' = this + from `s \\<^sub>\ Abs s'` show ?case + proof cases + case (eta s'' dummy) + from abs have "Abs s' => Abs t" by simp + with eta have "s''[dummy/0] => Abs t" by simp + with `\ free s'' 0` have "\t'. Abs (s'' \ Var 0) => t' \ t' \\<^sub>\\<^sup>* Abs t" + by (rule eta_case) + with eta show ?thesis by simp + next + case (abs r u) + hence "r \\<^sub>\ s'" by simp + then obtain t' where r: "r => t'" and t': "t' \\<^sub>\\<^sup>* t" by (iprover dest: abs') + from r have "Abs r => Abs t'" .. + moreover from t' have "Abs t' \\<^sub>\\<^sup>* Abs t" by (rule rtrancl_eta_Abs) + ultimately show ?thesis using abs by simp iprover + qed simp_all +next + case (app u u' t t') + from `s \\<^sub>\ u \ t` show ?case + proof cases + case (eta s' dummy) + from app have "u \ t => u' \ t'" by simp + with eta have "s'[dummy/0] => u' \ t'" by simp + with `\ free s' 0` have "\r. Abs (s' \ Var 0) => r \ r \\<^sub>\\<^sup>* u' \ t'" + by (rule eta_case) + with eta show ?thesis by simp + next + case (appL s' t'' u'') + hence "s' \\<^sub>\ u" by simp + then obtain r where s': "s' => r" and r: "r \\<^sub>\\<^sup>* u'" by (iprover dest: app) + from s' and app have "s' \ t => r \ t'" by simp + moreover from r have "r \ t' \\<^sub>\\<^sup>* u' \ t'" by (simp add: rtrancl_eta_AppL) + ultimately show ?thesis using appL by simp iprover + next + case (appR s' t'' u'') + hence "s' \\<^sub>\ t" by simp + then obtain r where s': "s' => r" and r: "r \\<^sub>\\<^sup>* t'" by (iprover dest: app) + from s' and app have "u \ s' => u' \ r" by simp + moreover from r have "u' \ r \\<^sub>\\<^sup>* u' \ t'" by (simp add: rtrancl_eta_AppR) + ultimately show ?thesis using appR by simp iprover + qed simp +next + case (beta u u' t t') + from `s \\<^sub>\ Abs u \ t` show ?case + proof cases + case (eta s' dummy) + from beta have "Abs u \ t => u'[t'/0]" by simp + with eta have "s'[dummy/0] => u'[t'/0]" by simp + with `\ free s' 0` have "\r. Abs (s' \ Var 0) => r \ r \\<^sub>\\<^sup>* u'[t'/0]" + by (rule eta_case) + with eta show ?thesis by simp + next + case (appL s' t'' u'') + hence "s' \\<^sub>\ Abs u" by simp + thus ?thesis + proof cases + case (eta s'' dummy) + have "Abs (lift u 1) = lift (Abs u) 0" by simp + also from eta have "\ = s''" by (simp add: lift_subst_dummy del: lift_subst) + finally have s: "s = Abs (Abs (lift u 1) \ Var 0) \ t" using appL and eta by simp + from beta have "lift u 1 => lift u' 1" by simp + hence "Abs (lift u 1) \ Var 0 => lift u' 1[Var 0/0]" + using par_beta.var .. + hence "Abs (Abs (lift u 1) \ Var 0) \ t => lift u' 1[Var 0/0][t'/0]" + using `t => t'` .. + with s have "s => u'[t'/0]" by simp + thus ?thesis by iprover + next + case (abs r r') + hence "r \\<^sub>\ u" by simp + then obtain r'' where r: "r => r''" and r'': "r'' \\<^sub>\\<^sup>* u'" by (iprover dest: beta) + from r and beta have "Abs r \ t => r''[t'/0]" by simp + moreover from r'' have "r''[t'/0] \\<^sub>\\<^sup>* u'[t'/0]" + by (rule rtrancl_eta_subst') + ultimately show ?thesis using abs and appL by simp iprover + qed simp_all + next + case (appR s' t'' u'') + hence "s' \\<^sub>\ t" by simp + then obtain r where s': "s' => r" and r: "r \\<^sub>\\<^sup>* t'" by (iprover dest: beta) + from s' and beta have "Abs u \ s' => u'[r/0]" by simp + moreover from r have "u'[r/0] \\<^sub>\\<^sup>* u'[t'/0]" + by (rule rtrancl_eta_subst'') + ultimately show ?thesis using appR by simp iprover + qed simp +qed + +theorem eta_postponement': + assumes eta: "s \\<^sub>\\<^sup>* t" and beta: "t => u" + shows "\t'. s => t' \ t' \\<^sub>\\<^sup>* u" using eta beta proof (induct arbitrary: u) case 1 thus ?case by blast next case (2 s' s'' s''') - from 2 obtain t' where s': "s' => t'" and t': "t' \\<^sub>\ s'''" - by (auto dest: par_eta_beta) - from s' obtain t'' where s: "s => t''" and t'': "t'' -e>> t'" using 2 + from 2 obtain t' where s': "s' => t'" and t': "t' \\<^sub>\\<^sup>* s'''" + by (auto dest: eta_par_beta) + from s' obtain t'' where s: "s => t''" and t'': "t'' \\<^sub>\\<^sup>* t'" using 2 by blast - from par_eta_subset_eta t' have "t' -e>> s'''" .. - with t'' have "t'' -e>> s'''" by (rule rtrancl_trans) + from t'' and t' have "t'' \\<^sub>\\<^sup>* s'''" by (rule rtrancl_trans') with s show ?case by iprover qed theorem eta_postponement: - assumes st: "(s, t) \ (beta \ eta)\<^sup>*" - shows "(s, t) \ eta\<^sup>* O beta\<^sup>*" using st + assumes st: "(join beta eta)\<^sup>*\<^sup>* s t" + shows "(eta\<^sup>*\<^sup>* OO beta\<^sup>*\<^sup>*) s t" using st proof induct case 1 show ?case by blast next case (2 s' s'') - from 2(3) obtain t' where s: "s \\<^sub>\\<^sup>* t'" and t': "t' -e>> s'" by blast + from 2(3) obtain t' where s: "s \\<^sub>\\<^sup>* t'" and t': "t' \\<^sub>\\<^sup>* s'" by blast from 2(2) show ?case proof - assume "s' -> s''" + assume "s' \\<^sub>\ s''" with beta_subset_par_beta have "s' => s''" .. - with t' obtain t'' where st: "t' => t''" and tu: "t'' -e>> s''" + with t' obtain t'' where st: "t' => t''" and tu: "t'' \\<^sub>\\<^sup>* s''" by (auto dest: eta_postponement') from par_beta_subset_beta st have "t' \\<^sub>\\<^sup>* t''" .. - with s have "s \\<^sub>\\<^sup>* t''" by (rule rtrancl_trans) + with s have "s \\<^sub>\\<^sup>* t''" by (rule rtrancl_trans') thus ?thesis using tu .. next - assume "s' -e> s''" - with t' have "t' -e>> s''" .. + assume "s' \\<^sub>\ s''" + with t' have "t' \\<^sub>\\<^sup>* s''" .. with s show ?thesis .. qed qed