# HG changeset patch # User nipkow # Date 1325497888 -3600 # Node ID b9d4ec0f79acacaca504a2d0d33ad89df1ab96a6 # Parent a03bf644cb27f764d0d8cd12b8499f97f4385087 tuned proofs diff -r a03bf644cb27 -r b9d4ec0f79ac src/HOL/IMP/ACom.thy --- a/src/HOL/IMP/ACom.thy Sun Jan 01 18:12:11 2012 +0100 +++ b/src/HOL/IMP/ACom.thy Mon Jan 02 10:51:28 2012 +0100 @@ -57,16 +57,33 @@ lemma strip_acom[simp]: "strip (map_acom f c) = strip c" by (induction c) auto +lemma map_acom_SKIP: + "map_acom f c = SKIP {S'} \ (\S. c = SKIP {S} \ S' = f S)" +by (cases c) auto + +lemma map_acom_Assign: + "map_acom f c = x ::= e {S'} \ (\S. c = x::=e {S} \ S' = f S)" +by (cases c) auto + +lemma map_acom_Semi: + "map_acom f c = c1';c2' \ + (\c1 c2. c = c1;c2 \ map_acom f c1 = c1' \ map_acom f c2 = c2')" +by (cases c) auto + +lemma map_acom_If: + "map_acom f c = IF b THEN c1' ELSE c2' {S'} \ + (\S c1 c2. c = IF b THEN c1 ELSE c2 {S} \ map_acom f c1 = c1' \ map_acom f c2 = c2' \ S' = f S)" +by (cases c) auto + +lemma map_acom_While: + "map_acom f w = {I'} WHILE b DO c' {P'} \ + (\I P c. w = {I} WHILE b DO c {P} \ map_acom f c = c' \ I' = f I \ P' = f P)" +by (cases w) auto + lemma strip_anno[simp]: "strip (anno a c) = c" by(induct c) simp_all -lemma strip_eq_SKIP: "strip c = com.SKIP \ (EX P. c = SKIP {P})" -by (cases c) simp_all - -lemma strip_eq_Assign: "strip c = x::=e \ (EX P. c = x::=e {P})" -by (cases c) simp_all - lemma strip_eq_Semi: "strip c = c1;c2 \ (EX d1 d2. c = d1;d2 & strip d1 = c1 & strip d2 = c2)" by (cases c) simp_all diff -r a03bf644cb27 -r b9d4ec0f79ac src/HOL/IMP/Abs_Int0.thy --- a/src/HOL/IMP/Abs_Int0.thy Sun Jan 01 18:12:11 2012 +0100 +++ b/src/HOL/IMP/Abs_Int0.thy Mon Jan 02 10:51:28 2012 +0100 @@ -57,49 +57,40 @@ text{* The soundness proofs are textually identical to the ones for the step function operating on states as functions. *} -lemma step_preserves_le2: - "\ S \ \\<^isub>o S'; cs \ \\<^isub>c ca; strip cs = c; strip ca = c \ - \ step S cs \ \\<^isub>c (step' S' ca)" -proof(induction c arbitrary: cs ca S S') - case SKIP thus ?case - by(auto simp:strip_eq_SKIP) +lemma step_preserves_le: + "\ S \ \\<^isub>o S'; cs \ \\<^isub>c ca \ \ step S cs \ \\<^isub>c (step' S' ca)" +proof(induction cs arbitrary: ca S S') + case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) next case Assign thus ?case - by (fastforce simp: strip_eq_Assign intro: aval'_sound in_gamma_update + by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update split: option.splits del:subsetD) next - case Semi thus ?case apply (auto simp: strip_eq_Semi) + case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi) by (metis le_post post_map_acom) next - case (If b c1 c2) - then obtain cs1 cs2 ca1 ca2 P Pa where - "cs = IF b THEN cs1 ELSE cs2 {P}" "ca = IF b THEN ca1 ELSE ca2 {Pa}" + case (If b cs1 cs2 P) + then obtain ca1 ca2 Pa where + "ca= IF b THEN ca1 ELSE ca2 {Pa}" "P \ \\<^isub>o Pa" "cs1 \ \\<^isub>c ca1" "cs2 \ \\<^isub>c ca2" - "strip cs1 = c1" "strip ca1 = c1" "strip cs2 = c2" "strip ca2 = c2" - by (fastforce simp: strip_eq_If) + by (fastforce simp: If_le map_acom_If) moreover have "post cs1 \ \\<^isub>o(post ca1 \ post ca2)" by (metis (no_types) `cs1 \ \\<^isub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom) moreover have "post cs2 \ \\<^isub>o(post ca1 \ post ca2)" by (metis (no_types) `cs2 \ \\<^isub>c ca2` join_ge2 le_post mono_gamma_o order_trans post_map_acom) - ultimately show ?case using If.prems(1) by (simp add: If.IH subset_iff) + ultimately show ?case using `S \ \\<^isub>o S'` by (simp add: If.IH subset_iff) next - case (While b c1) - then obtain cs1 ca1 I P Ia Pa where - "cs = {I} WHILE b DO cs1 {P}" "ca = {Ia} WHILE b DO ca1 {Pa}" + case (While I b cs1 P) + then obtain ca1 Ia Pa where + "ca = {Ia} WHILE b DO ca1 {Pa}" "I \ \\<^isub>o Ia" "P \ \\<^isub>o Pa" "cs1 \ \\<^isub>c ca1" - "strip cs1 = c1" "strip ca1 = c1" - by (fastforce simp: strip_eq_While) + by (fastforce simp: map_acom_While While_le) moreover have "S \ post cs1 \ \\<^isub>o (S' \ post ca1)" using `S \ \\<^isub>o S'` le_post[OF `cs1 \ \\<^isub>c ca1`, simplified] by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) ultimately show ?case by (simp add: While.IH subset_iff) qed -lemma step_preserves_le: - "\ S \ \\<^isub>o S'; cs \ \\<^isub>c ca; strip cs = c \ - \ step S cs \ \\<^isub>c(step' S' ca)" -by (metis le_strip step_preserves_le2 strip_acom) - lemma AI_sound: "AI c = Some c' \ CS UNIV c \ \\<^isub>c c'" proof(simp add: CS_def AI_def) assume 1: "lpfp\<^isub>c (step' \) c = Some c'" @@ -109,7 +100,7 @@ have "lfp (step UNIV) c \ \\<^isub>c (step' \ c')" proof(rule lfp_lowerbound[simplified,OF 3]) show "step UNIV (\\<^isub>c (step' \ c')) \ \\<^isub>c (step' \ c')" - proof(rule step_preserves_le[OF _ _ 3]) + proof(rule step_preserves_le[OF _ _]) show "UNIV \ \\<^isub>o \" by simp show "\\<^isub>c (step' \ c') \ \\<^isub>c c'" by(rule mono_gamma_c[OF 2]) qed diff -r a03bf644cb27 -r b9d4ec0f79ac src/HOL/IMP/Abs_Int0_fun.thy --- a/src/HOL/IMP/Abs_Int0_fun.thy Sun Jan 01 18:12:11 2012 +0100 +++ b/src/HOL/IMP/Abs_Int0_fun.thy Mon Jan 02 10:51:28 2012 +0100 @@ -308,49 +308,40 @@ "\ s : \\<^isub>f S; i : \ a \ \ s(x := i) : \\<^isub>f(S(x := a))" by(simp add: \_fun_def) -lemma step_preserves_le2: - "\ S \ \\<^isub>o S'; cs \ \\<^isub>c ca; strip cs = c; strip ca = c \ - \ step S cs \ \\<^isub>c (step' S' ca)" -proof(induction c arbitrary: cs ca S S') - case SKIP thus ?case - by(auto simp:strip_eq_SKIP) +lemma step_preserves_le: + "\ S \ \\<^isub>o S'; cs \ \\<^isub>c ca \ \ step S cs \ \\<^isub>c (step' S' ca)" +proof(induction cs arbitrary: ca S S') + case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) next case Assign thus ?case - by (fastforce simp: strip_eq_Assign intro: aval'_sound in_gamma_update + by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update split: option.splits del:subsetD) next - case Semi thus ?case apply (auto simp: strip_eq_Semi) + case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi) by (metis le_post post_map_acom) next - case (If b c1 c2) - then obtain cs1 cs2 ca1 ca2 P Pa where - "cs= IF b THEN cs1 ELSE cs2 {P}" "ca= IF b THEN ca1 ELSE ca2 {Pa}" + case (If b cs1 cs2 P) + then obtain ca1 ca2 Pa where + "ca= IF b THEN ca1 ELSE ca2 {Pa}" "P \ \\<^isub>o Pa" "cs1 \ \\<^isub>c ca1" "cs2 \ \\<^isub>c ca2" - "strip cs1 = c1" "strip ca1 = c1" "strip cs2 = c2" "strip ca2 = c2" - by (fastforce simp: strip_eq_If) + by (fastforce simp: If_le map_acom_If) moreover have "post cs1 \ \\<^isub>o(post ca1 \ post ca2)" by (metis (no_types) `cs1 \ \\<^isub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom) moreover have "post cs2 \ \\<^isub>o(post ca1 \ post ca2)" by (metis (no_types) `cs2 \ \\<^isub>c ca2` join_ge2 le_post mono_gamma_o order_trans post_map_acom) ultimately show ?case using `S \ \\<^isub>o S'` by (simp add: If.IH subset_iff) next - case (While b c1) - then obtain cs1 ca1 I P Ia Pa where - "cs = {I} WHILE b DO cs1 {P}" "ca = {Ia} WHILE b DO ca1 {Pa}" + case (While I b cs1 P) + then obtain ca1 Ia Pa where + "ca = {Ia} WHILE b DO ca1 {Pa}" "I \ \\<^isub>o Ia" "P \ \\<^isub>o Pa" "cs1 \ \\<^isub>c ca1" - "strip cs1 = c1" "strip ca1 = c1" - by (fastforce simp: strip_eq_While) + by (fastforce simp: map_acom_While While_le) moreover have "S \ post cs1 \ \\<^isub>o (S' \ post ca1)" using `S \ \\<^isub>o S'` le_post[OF `cs1 \ \\<^isub>c ca1`, simplified] by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) ultimately show ?case by (simp add: While.IH subset_iff) qed -lemma step_preserves_le: - "\ S \ \\<^isub>o S'; cs \ \\<^isub>c ca; strip cs = c \ - \ step S cs \ \\<^isub>c(step' S' ca)" -by (metis le_strip step_preserves_le2 strip_acom) - lemma AI_sound: "AI c = Some c' \ CS UNIV c \ \\<^isub>c c'" proof(simp add: CS_def AI_def) assume 1: "lpfp\<^isub>c (step' \) c = Some c'" @@ -360,12 +351,12 @@ have "lfp (step UNIV) c \ \\<^isub>c (step' \ c')" proof(rule lfp_lowerbound[simplified,OF 3]) show "step UNIV (\\<^isub>c (step' \ c')) \ \\<^isub>c (step' \ c')" - proof(rule step_preserves_le[OF _ _ 3]) + proof(rule step_preserves_le[OF _ _]) show "UNIV \ \\<^isub>o \" by simp show "\\<^isub>c (step' \ c') \ \\<^isub>c c'" by(rule mono_gamma_c[OF 2]) qed qed - from this 2 show "lfp (step UNIV) c \ \\<^isub>c c'" + with 2 show "lfp (step UNIV) c \ \\<^isub>c c'" by (blast intro: mono_gamma_c order_trans) qed diff -r a03bf644cb27 -r b9d4ec0f79ac src/HOL/IMP/Abs_Int1.thy --- a/src/HOL/IMP/Abs_Int1.thy Sun Jan 01 18:12:11 2012 +0100 +++ b/src/HOL/IMP/Abs_Int1.thy Mon Jan 02 10:51:28 2012 +0100 @@ -168,27 +168,23 @@ "\ s : \\<^isub>f S; i : \ a \ \ s(x := i) : \\<^isub>f(update S x a)" by(simp add: \_st_def lookup_update) - -lemma step_preserves_le2: - "\ S \ \\<^isub>o S'; cs \ \\<^isub>c ca; strip cs = c; strip ca = c \ - \ step S cs \ \\<^isub>c (step' S' ca)" -proof(induction c arbitrary: cs ca S S') - case SKIP thus ?case - by(auto simp:strip_eq_SKIP) +lemma step_preserves_le: + "\ S \ \\<^isub>o S'; cs \ \\<^isub>c ca \ \ step S cs \ \\<^isub>c (step' S' ca)" +proof(induction cs arbitrary: ca S S') + case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) next case Assign thus ?case - by (fastforce simp: strip_eq_Assign intro: aval'_sound in_gamma_update + by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update split: option.splits del:subsetD) next - case Semi thus ?case apply (auto simp: strip_eq_Semi) + case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi) by (metis le_post post_map_acom) next - case (If b c1 c2) - then obtain cs1 cs2 ca1 ca2 P Pa where - "cs= IF b THEN cs1 ELSE cs2 {P}" "ca= IF b THEN ca1 ELSE ca2 {Pa}" + case (If b cs1 cs2 P) + then obtain ca1 ca2 Pa where + "ca= IF b THEN ca1 ELSE ca2 {Pa}" "P \ \\<^isub>o Pa" "cs1 \ \\<^isub>c ca1" "cs2 \ \\<^isub>c ca2" - "strip cs1 = c1" "strip ca1 = c1" "strip cs2 = c2" "strip ca2 = c2" - by (fastforce simp: strip_eq_If) + by (fastforce simp: If_le map_acom_If) moreover have "post cs1 \ \\<^isub>o(post ca1 \ post ca2)" by (metis (no_types) `cs1 \ \\<^isub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom) moreover have "post cs2 \ \\<^isub>o(post ca1 \ post ca2)" @@ -196,23 +192,17 @@ ultimately show ?case using `S \ \\<^isub>o S'` by (simp add: If.IH subset_iff bfilter_sound) next - case (While b c1) - then obtain cs1 ca1 I P Ia Pa where - "cs = {I} WHILE b DO cs1 {P}" "ca = {Ia} WHILE b DO ca1 {Pa}" + case (While I b cs1 P) + then obtain ca1 Ia Pa where + "ca = {Ia} WHILE b DO ca1 {Pa}" "I \ \\<^isub>o Ia" "P \ \\<^isub>o Pa" "cs1 \ \\<^isub>c ca1" - "strip cs1 = c1" "strip ca1 = c1" - by (fastforce simp: strip_eq_While) + by (fastforce simp: map_acom_While While_le) moreover have "S \ post cs1 \ \\<^isub>o (S' \ post ca1)" using `S \ \\<^isub>o S'` le_post[OF `cs1 \ \\<^isub>c ca1`, simplified] by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) ultimately show ?case by (simp add: While.IH subset_iff bfilter_sound) qed -lemma step_preserves_le: - "\ S \ \\<^isub>o S'; cs \ \\<^isub>c ca; strip cs = c \ - \ step S cs \ \\<^isub>c(step' S' ca)" -by (metis le_strip step_preserves_le2 strip_acom) - lemma AI_sound: "AI c = Some c' \ CS UNIV c \ \\<^isub>c c'" proof(simp add: CS_def AI_def) assume 1: "lpfp\<^isub>c (step' \) c = Some c'" @@ -222,7 +212,7 @@ have "lfp (step UNIV) c \ \\<^isub>c (step' \ c')" proof(rule lfp_lowerbound[simplified,OF 3]) show "step UNIV (\\<^isub>c (step' \ c')) \ \\<^isub>c (step' \ c')" - proof(rule step_preserves_le[OF _ _ 3]) + proof(rule step_preserves_le[OF _ _]) show "UNIV \ \\<^isub>o \" by simp show "\\<^isub>c (step' \ c') \ \\<^isub>c c'" by(rule mono_gamma_c[OF 2]) qed diff -r a03bf644cb27 -r b9d4ec0f79ac src/HOL/IMP/Abs_Int2.thy --- a/src/HOL/IMP/Abs_Int2.thy Sun Jan 01 18:12:11 2012 +0100 +++ b/src/HOL/IMP/Abs_Int2.thy Mon Jan 02 10:51:28 2012 +0100 @@ -196,7 +196,7 @@ have "lfp (step UNIV) c \ \\<^isub>c (step' \ c')" proof(rule lfp_lowerbound[simplified,OF 3]) show "step UNIV (\\<^isub>c (step' \ c')) \ \\<^isub>c (step' \ c')" - proof(rule step_preserves_le[OF _ _ 3]) + proof(rule step_preserves_le[OF _ _]) show "UNIV \ \\<^isub>o \" by simp show "\\<^isub>c (step' \ c') \ \\<^isub>c c'" by(rule mono_gamma_c[OF 2]) qed