# HG changeset patch # User nipkow # Date 1227217152 -3600 # Node ID 3d9873c4c40903c293376baf6856de000d73038c # Parent 30cd9d89a0fb3e6326ff0a717f58a48004ddeee6 added optimizer diff -r 30cd9d89a0fb -r 3d9873c4c409 src/HOL/IMP/Live.thy --- a/src/HOL/IMP/Live.thy Thu Nov 20 19:43:34 2008 +0100 +++ b/src/HOL/IMP/Live.thy Thu Nov 20 22:39:12 2008 +0100 @@ -92,17 +92,80 @@ 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 - note IH1 = IH(1)[OF _ `\c,s\ \\<^sub>c s''` `\c,t\ \\<^sub>c t''`] - have L1: "\x\A. s'' x = t'' x" using IH1 WhileTrue(6,8) - by(simp add: ball_Un) (metis) - have L2: "\x\Dep b. s'' x = t'' x" - using IH1 WhileTrue(6,8) by (auto simp:L_gen_kill) - have L3: "\x\L c A. s'' x = t'' x" - using IH1 L_idemp[of c A] WhileTrue(6,8) by auto - have "\x\L (While b c) A. s'' x = t'' x" using L1 L2 L3 by auto - then show ?case using WhileTrue(5,6) `\While b c,t''\ \\<^sub>c t'` by metis + 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 qed auto } from this[OF IH(3) _ IH(4,2)] show ?case by metis qed + +primrec bury :: "com \ loc set \ com" where +"bury SKIP _ = SKIP" | +"bury (x :== e) A = (if x:A then x:== e else SKIP)" | +"bury (c1; c2) A = (bury c1 (L c2 A); bury c2 A)" | +"bury (IF b THEN c1 ELSE c2) A = (IF b THEN bury c1 A ELSE bury c2 A)" | +"bury (WHILE b DO c) A = (WHILE b DO bury c (Dep b \ A \ L c A))" + +theorem bury_sound: + "\ x \ L c A. s x = t x \ \c,s\ \\<^sub>c s' \ \bury c A,t\ \\<^sub>c t' \ + \x\A. s' x = t' x" +proof (induct c arbitrary: A s t s' t') + case SKIP then show ?case by auto +next + case (Assign x e) then show ?case + by (auto simp:update_def ball_Un split:split_if_asm dest!: dep_on) +next + case (Semi c1 c2) + from Semi(4) obtain s'' where s1: "\c1,s\ \\<^sub>c s''" and s2: "\c2,s''\ \\<^sub>c s'" + by auto + from Semi(5) obtain t'' where t1: "\bury c1 (L c2 A),t\ \\<^sub>c t''" and t2: "\bury c2 A,t''\ \\<^sub>c t'" + by auto + show ?case using Semi(1)[OF _ s1 t1] Semi(2)[OF _ s2 t2] Semi(3) by fastsimp +next + case (Cond b c1 c2) + show ?case + proof cases + assume "b s" + hence s: "\c1,s\ \\<^sub>c s'" using Cond(4) by simp + have "b t" using `b s` Cond(3) by (simp add: ball_Un)(blast dest: dep_on) + hence t: "\bury c1 A,t\ \\<^sub>c t'" using Cond(5) by auto + show ?thesis using Cond(1)[OF _ s t] Cond(3) by fastsimp + next + assume "\ b s" + hence s: "\c2,s\ \\<^sub>c s'" using Cond(4) by auto + have "\ b t" using `\ b s` Cond(3) by (simp add: ball_Un)(blast dest: dep_on) + hence t: "\bury c2 A,t\ \\<^sub>c t'" using Cond(5) by auto + show ?thesis using Cond(2)[OF _ s t] Cond(3) by fastsimp + qed +next + case (While b c) note IH = this + { fix cw + have "\cw,s\ \\<^sub>c s' \ cw = (While b c) \ \bury cw A,t\ \\<^sub>c t' \ + \ x \ L cw A. s x = t x \ \x\A. s' x = t' x" + proof (induct 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 auto + next + case (WhileTrue _ s _ s'' s') + have "\c,s\ \\<^sub>c s''" using WhileTrue(2,6) by simp + 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(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) `\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 +qed + + end \ No newline at end of file