# HG changeset patch # User nipkow # Date 1268670843 -3600 # Node ID 3c1601857a6b033ca3f5917a8417be185377848a # Parent 533dd944e29ca70d049a0dcae76c4869323b1d2c# Parent 362431732b5e9d73d42054fa9e3701f71ce8e94f merged diff -r 533dd944e29c -r 3c1601857a6b src/HOL/IMP/Live.thy --- a/src/HOL/IMP/Live.thy Mon Mar 15 15:13:22 2010 +0100 +++ b/src/HOL/IMP/Live.thy Mon Mar 15 17:34:03 2010 +0100 @@ -95,10 +95,30 @@ have "\x\Dep b \ A \ L c A. s'' x = t'' x" using IH(1)[OF _ `\c,s\ \\<^sub>c s''` `\c,t\ \\<^sub>c t''`] WhileTrue(6,8) by (auto simp:L_gen_kill) - moreover then have "\x\L (While b c) A. s'' x = t'' x" by auto - ultimately show ?case using WhileTrue(5,6) `\While b c,t''\ \\<^sub>c t'` by metis + then have "\x\L (While b c) A. s'' x = t'' x" by auto + then show ?case using WhileTrue(5,6) `\While b c,t''\ \\<^sub>c t'` by metis qed auto } - from this[OF IH(3) _ IH(4,2)] show ?case by metis +-- "a terser version" + { let ?w = "While b c" + have "\?w,s\ \\<^sub>c s' \ \?w,t\ \\<^sub>c t' \ + \ x \ L ?w A. s x = t x \ \x\A. s' x = t' x" + proof (induct ?w s s' arbitrary: t A pred:evalc) + case WhileFalse + have "\ b t" using WhileFalse by (simp add: ball_Un)(blast dest:dep_on) + then have "t' = t" using WhileFalse by auto + then show ?case using WhileFalse by simp + next + case (WhileTrue s s'' s') + have "b t" using WhileTrue by (simp add: ball_Un)(blast dest:dep_on) + then obtain t'' where "\c,t\ \\<^sub>c t''" and "\While b c,t''\ \\<^sub>c t'" + using WhileTrue(6,7) by auto + have "\x\Dep b \ A \ L c A. s'' x = t'' x" + using IH(1)[OF _ `\c,s\ \\<^sub>c s''` `\c,t\ \\<^sub>c t''`] WhileTrue(7) + by (auto simp:L_gen_kill) + then have "\x\L (While b c) A. s'' x = t'' x" by auto + then show ?case using WhileTrue(5) `\While b c,t''\ \\<^sub>c t'` by metis + qed } + from this[OF IH(3) IH(4,2)] show ?case by metis qed @@ -164,7 +184,28 @@ ultimately show ?case using WhileTrue(5,6) `\bury (While b c) A,t''\ \\<^sub>c t'` by metis qed auto } - from this[OF IH(3) _ IH(4,2)] show ?case by metis + { let ?w = "While b c" + have "\?w,s\ \\<^sub>c s' \ \bury ?w A,t\ \\<^sub>c t' \ + \ x \ L ?w A. s x = t x \ \x\A. s' x = t' x" + proof (induct ?w s s' arbitrary: t A pred:evalc) + case WhileFalse + have "\ b t" using WhileFalse by (simp add: ball_Un)(blast dest:dep_on) + then have "t' = t" using WhileFalse by auto + then show ?case using WhileFalse by simp + next + case (WhileTrue s s'' s') + have "b t" using WhileTrue by (simp add: ball_Un)(blast dest:dep_on) + then obtain t'' where tt'': "\bury c (Dep b \ A \ L c A),t\ \\<^sub>c t''" + and "\bury (While b c) A,t''\ \\<^sub>c t'" + using WhileTrue(6,7) by auto + have "\x\Dep b \ A \ L c A. s'' x = t'' x" + using IH(1)[OF _ `\c,s\ \\<^sub>c s''` tt''] WhileTrue(7) + by (auto simp:L_gen_kill) + then have "\x\L (While b c) A. s'' x = t'' x" by auto + then show ?case + using WhileTrue(5) `\bury (While b c) A,t''\ \\<^sub>c t'` by metis + qed } + from this[OF IH(3) IH(4,2)] show ?case by metis qed