# HG changeset patch # User berghofe # Date 1264867426 -3600 # Node ID 81e8fdfeb8491ff112e7b1c69b3c3ee99d92ba2f # Parent b5c6e59e2cd721943c96f8ac90481f15236982a9 Adapted to changes in cases method. diff -r b5c6e59e2cd7 -r 81e8fdfeb849 src/HOL/Bali/DeclConcepts.thy --- a/src/HOL/Bali/DeclConcepts.thy Sat Jan 30 17:01:01 2010 +0100 +++ b/src/HOL/Bali/DeclConcepts.thy Sat Jan 30 17:03:46 2010 +0100 @@ -915,23 +915,15 @@ assume "G \ m member_of C" then show "n=m" proof (cases) - case (Immediate m' _) - with eqid - have "m=m'" - "memberid n = memberid m" - "G\ mbr m declared_in C" - "declclass m = C" - by auto - with member_n + case Immediate + with eqid member_n show ?thesis by (cases n, cases m) (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def split: memberdecl.splits) next - case (Inherited m' _ _) - then have "G\ memberid m undeclared_in C" - by simp + case Inherited with eqid member_n show ?thesis by (cases n) (auto dest: declared_not_undeclared) @@ -1656,10 +1648,7 @@ from member_of show "?Methd C" proof (cases) - case (Immediate membr Ca) - then have "Ca=C" "membr = method sig m" and - "G\Methd sig m declared_in C" "declclass m = C" - by (cases m,auto) + case Immediate with clsC have "table_of (map (\(s, m). (s, C, m)) (methods c)) sig = Some m" by (cases m) @@ -1669,13 +1658,12 @@ show ?thesis by (simp add: methd_rec) next - case (Inherited membr Ca S) + case (Inherited S) with clsC - have eq_Ca_C: "Ca=C" and - undecl: "G\mid sig undeclared_in C" and + have undecl: "G\mid sig undeclared_in C" and super: "G \Methd sig m member_of (super c)" by (auto dest: subcls1D) - from eq_Ca_C clsC undecl + from clsC undecl have "table_of (map (\(s, m). (s, C, m)) (methods c)) sig = None" by (auto simp add: undeclared_in_def cdeclaredmethd_def intro: table_of_mapconst_NoneI) diff -r b5c6e59e2cd7 -r 81e8fdfeb849 src/HOL/IMP/Transition.thy --- a/src/HOL/IMP/Transition.thy Sat Jan 30 17:01:01 2010 +0100 +++ b/src/HOL/IMP/Transition.thy Sat Jan 30 17:03:46 2010 +0100 @@ -205,20 +205,16 @@ (is "\i j s'. ?Q i j s'") proof (cases set: evalc1) case Semi1 - then obtain s' where - "co = Some c2" and "s''' = s'" and "\c1, s\ \\<^sub>1 \s'\" - by auto - with 1 n have "?Q 1 n s'" by simp + from `co = Some c2` and `\c1, s\ \\<^sub>1 \s'''\` and 1 n + have "?Q 1 n s'''" by simp thus ?thesis by blast next - case Semi2 - then obtain c1' s' where - "co = Some (c1'; c2)" "s''' = s'" and - c1: "\c1, s\ \\<^sub>1 \c1', s'\" - by auto - with n have "\c1'; c2, s'\ -n\\<^sub>1 \s''\" by simp + case (Semi2 c1') + note c1 = `\c1, s\ \\<^sub>1 \c1', s'''\` + with `co = Some (c1'; c2)` and n + have "\c1'; c2, s'''\ -n\\<^sub>1 \s''\" by simp with Suc.hyps obtain i j s0 where - c1': "\c1',s'\ -i\\<^sub>1 \s0\" and + c1': "\c1',s'''\ -i\\<^sub>1 \s0\" and c2: "\c2,s0\ -j\\<^sub>1 \s''\" and i: "n = i+j" by fast @@ -228,7 +224,7 @@ with c2 i have "?Q (i+1) j s0" by simp thus ?thesis by blast - qed auto -- "the remaining cases cannot occur" + qed qed diff -r b5c6e59e2cd7 -r 81e8fdfeb849 src/HOL/Lambda/Eta.thy --- a/src/HOL/Lambda/Eta.thy Sat Jan 30 17:01:01 2010 +0100 +++ b/src/HOL/Lambda/Eta.thy Sat Jan 30 17:03:46 2010 +0100 @@ -273,13 +273,13 @@ 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') + case (abs r) + from `r \\<^sub>\ s'` + 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 + qed next case (app u u' t t') from `s \\<^sub>\ u \ t` show ?case @@ -291,20 +291,20 @@ 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) + case (appL s') + from `s' \\<^sub>\ u` + 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) + case (appR s') + from `s' \\<^sub>\ t` + 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 + qed next case (beta u u' t t') from `s \\<^sub>\ Abs u \ t` show ?case @@ -316,9 +316,8 @@ 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 + case (appL s') + from `s' \\<^sub>\ Abs u` show ?thesis proof cases case (eta s'' dummy) have "Abs (lift u 1) = lift (Abs u) 0" by simp @@ -332,23 +331,23 @@ 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) + case (abs r) + from `r \\<^sub>\ u` + 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 + qed 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) + case (appR s') + from `s' \\<^sub>\ t` + 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 qed theorem eta_postponement': diff -r b5c6e59e2cd7 -r 81e8fdfeb849 src/HOL/Nominal/Examples/Pattern.thy --- a/src/HOL/Nominal/Examples/Pattern.thy Sat Jan 30 17:01:01 2010 +0100 +++ b/src/HOL/Nominal/Examples/Pattern.thy Sat Jan 30 17:03:46 2010 +0100 @@ -575,13 +575,13 @@ and R: "\U. S = T \ U \ (x, T) # \ \ t : U \ P" shows P using ty proof cases - case (Abs x' T' \' t' U) + case (Abs x' T' t' U) obtain y::name where y: "y \ (x, \, \x':T'. t')" by (rule exists_fresh) (auto intro: fin_supp) from `(\x:T. t) = (\x':T'. t')` [symmetric] have x: "x \ (\x':T'. t')" by (simp add: abs_fresh) have x': "x' \ (\x':T'. t')" by (simp add: abs_fresh) - from `(x', T') # \' \ t' : U` have x'': "x' \ \'" + from `(x', T') # \ \ t' : U` have x'': "x' \ \" by (auto dest: valid_typing) have "(\x:T. t) = (\x':T'. t')" by fact also from x x' y have "\ = [(x, y)] \ [(x', y)] \ (\x':T'. t')" @@ -592,10 +592,10 @@ then have T: "T = T'" and t: "[(x, y)] \ [(x', y)] \ t' = t" by (simp_all add: trm.inject alpha) from Abs T have "S = T \ U" by simp - moreover from `(x', T') # \' \ t' : U` - have "[(x, y)] \ [(x', y)] \ ((x', T') # \' \ t' : U)" + moreover from `(x', T') # \ \ t' : U` + have "[(x, y)] \ [(x', y)] \ ((x', T') # \ \ t' : U)" by (simp add: perm_bool) - with T t y `\ = \'` x'' fresh have "(x, T) # \ \ t : U" + with T t y x'' fresh have "(x, T) # \ \ t : U" by (simp add: eqvts swap_simps perm_fresh_fresh fresh_prod) ultimately show ?thesis by (rule R) qed simp_all @@ -764,7 +764,7 @@ and R: "\T \. \ \ t : T \ \ p : T \ \ \ \ @ \ \ u : U \ P" shows P using ty proof cases - case (Let p' t' \' T \ u' U') + case (Let p' t' T \ u') then have "(supp \::name set) \* \" by (auto intro: valid_typing valid_app_freshs) with Let have "(supp p'::name set) \* \" @@ -776,7 +776,7 @@ moreover from Let have "pat_type p = pat_type p'" by (simp add: trm.inject) moreover note distinct - moreover from `\ @ \' \ u' : U'` have "valid (\ @ \')" + moreover from `\ @ \ \ u' : U` have "valid (\ @ \)" by (rule valid_typing) then have "valid \" by (rule valid_appD) with `\ p' : T \ \` have "distinct (pat_vars p')"