# HG changeset patch # User nipkow # Date 1326877514 -3600 # Node ID e69684c1c1426faff1c458ccfafe0937149ee13d # Parent 549755ebf4d26393c732ae1d316f0b3f625992bc introduced commands over a set of vars diff -r 549755ebf4d2 -r e69684c1c142 src/HOL/IMP/Abs_Int0.thy --- a/src/HOL/IMP/Abs_Int0.thy Tue Jan 17 18:25:36 2012 +0100 +++ b/src/HOL/IMP/Abs_Int0.thy Wed Jan 18 10:05:14 2012 +0100 @@ -341,10 +341,15 @@ apply (auto simp: strip_eq_SKIP strip_eq_Assign strip_eq_Semi strip_eq_If strip_eq_While) done +lemmas size_annos_same2 = eqTrueI[OF size_annos_same] + +lemma set_annos_anno: "set (annos (anno a c)) = {a}" +by(induction c)(auto) + lemma le_iff_le_annos: "c1 \ c2 \ (annos c1, annos c2) : listrel{(x,y). x \ y} \ strip c1 = strip c2" apply(induct c1 c2 rule: le_acom.induct) -apply (auto simp: listrel.Nil listrel_Cons_iff listrel_app size_annos_same) +apply (auto simp: listrel.Nil listrel_Cons_iff listrel_app size_annos_same2) apply (metis listrel_app_same_size size_annos_same)+ done diff -r 549755ebf4d2 -r e69684c1c142 src/HOL/IMP/Abs_Int1.thy --- a/src/HOL/IMP/Abs_Int1.thy Tue Jan 17 18:25:36 2012 +0100 +++ b/src/HOL/IMP/Abs_Int1.thy Wed Jan 18 10:05:14 2012 +0100 @@ -1,7 +1,7 @@ (* Author: Tobias Nipkow *) theory Abs_Int1 -imports Abs_Int0 +imports Abs_Int0 Vars begin instantiation prod :: (preord,preord) preord @@ -18,6 +18,29 @@ end +instantiation com :: vars +begin + +fun vars_com :: "com \ vname set" where +"vars com.SKIP = {}" | +"vars (x::=e) = {x} \ vars e" | +"vars (c1;c2) = vars c1 \ vars c2" | +"vars (IF b THEN c1 ELSE c2) = vars b \ vars c1 \ vars c2" | +"vars (WHILE b DO c) = vars b \ vars c" + +instance .. + +end + +lemma finite_avars: "finite(vars(a::aexp))" +by(induction a) simp_all + +lemma finite_bvars: "finite(vars(b::bexp))" +by(induction b) (simp_all add: finite_avars) + +lemma finite_cvars: "finite(vars(c::com))" +by(induction c) (simp_all add: finite_avars finite_bvars) + subsection "Backward Analysis of Expressions" @@ -221,6 +244,65 @@ by (blast intro: mono_gamma_c order_trans) qed + +subsubsection "Commands over a set of variables" + +text{* Key invariant: the domains of all abstract states are subsets of the +set of variables of the program. *} + +definition "domo S = (case S of None \ {} | Some S' \ set(dom S'))" + +definition Com :: "vname set \ 'a st option acom set" where +"Com X = {c. \S \ set(annos c). domo S \ X}" + +lemma domo_Top[simp]: "domo \ = {}" +by(simp add: domo_def Top_st_def Top_option_def) + +lemma bot_acom_Dom[simp]: "\\<^sub>c c \ Com X" +by(simp add: bot_acom_def Com_def domo_def set_annos_anno) + +lemma post_in_annos: "post c : set(annos c)" +by(induction c) simp_all + +lemma domo_join: "domo (S \ T) \ domo S \ domo T" +by(auto simp: domo_def join_st_def split: option.split) + +lemma domo_afilter: "vars a \ X \ domo S \ X \ domo(afilter a i S) \ X" +apply(induction a arbitrary: i S) +apply(simp add: domo_def) +apply(simp add: domo_def Let_def update_def lookup_def split: option.splits) +apply blast +apply(simp split: prod.split) +done + +lemma domo_bfilter: "vars b \ X \ domo S \ X \ domo(bfilter b bv S) \ X" +apply(induction b arbitrary: bv S) +apply(simp add: domo_def) +apply(simp) +apply(simp) +apply rule +apply (metis le_sup_iff subset_trans[OF domo_join]) +apply(simp split: prod.split) +by (metis domo_afilter) + +lemma step'_Com: + "domo S \ X \ vars(strip c) \ X \ c : Com X \ step' S c : Com X" +apply(induction c arbitrary: S) +apply(simp add: Com_def) +apply(simp add: Com_def domo_def update_def split: option.splits) +apply(simp (no_asm_use) add: Com_def ball_Un) +apply (metis post_in_annos) +apply(simp (no_asm_use) add: Com_def ball_Un) +apply rule +apply (metis Un_assoc domo_join order_trans post_in_annos subset_Un_eq) +apply (metis domo_bfilter) +apply(simp (no_asm_use) add: Com_def) +apply rule +apply (metis domo_join le_sup_iff post_in_annos subset_trans) +apply rule +apply (metis domo_bfilter) +by (metis domo_bfilter) + end