| author | Lars Hupel <lars.hupel@mytum.de> | 
| Thu, 20 Jul 2017 23:59:09 +0200 | |
| changeset 66293 | 2eae295c8fc3 | 
| parent 58310 | 91ea607a34d8 | 
| child 67406 | 23307fd33906 | 
| permissions | -rw-r--r-- | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 1 | (* Author: Tobias Nipkow *) | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 2 | |
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 3 | theory ACom | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 4 | imports Com | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 5 | begin | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 6 | |
| 46225 | 7 | subsection "Annotated Commands" | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 8 | |
| 58310 | 9 | datatype 'a acom = | 
| 47818 | 10 |   SKIP 'a                           ("SKIP {_}" 61) |
 | 
| 11 |   Assign vname aexp 'a              ("(_ ::= _/ {_})" [1000, 61, 0] 61) |
 | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
52019diff
changeset | 12 |   Seq "('a acom)" "('a acom)"       ("_;;//_"  [60, 61] 60) |
 | 
| 49095 | 13 |   If bexp 'a "('a acom)" 'a "('a acom)" 'a
 | 
| 14 |     ("(IF _/ THEN ({_}/ _)/ ELSE ({_}/ _)//{_})"  [0, 0, 0, 61, 0, 0] 61) |
 | |
| 15 |   While 'a bexp 'a "('a acom)" 'a
 | |
| 49138 | 16 |     ("({_}//WHILE _//DO ({_}//_)//{_})"  [0, 0, 0, 61, 0] 61)
 | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 17 | |
| 54941 | 18 | notation com.SKIP ("SKIP")
 | 
| 19 | ||
| 49603 | 20 | text_raw{*\snip{stripdef}{1}{1}{% *}
 | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 21 | fun strip :: "'a acom \<Rightarrow> com" where | 
| 54941 | 22 | "strip (SKIP {P}) = SKIP" |
 | 
| 49603 | 23 | "strip (x ::= e {P}) = x ::= e" |
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 24 | "strip (C\<^sub>1;;C\<^sub>2) = strip C\<^sub>1;; strip C\<^sub>2" | | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 25 | "strip (IF b THEN {P\<^sub>1} C\<^sub>1 ELSE {P\<^sub>2} C\<^sub>2 {P}) =
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 26 | IF b THEN strip C\<^sub>1 ELSE strip C\<^sub>2" | | 
| 49603 | 27 | "strip ({I} WHILE b DO {P} C {Q}) = WHILE b DO strip C"
 | 
| 28 | text_raw{*}%endsnip*}
 | |
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 29 | |
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 30 | text_raw{*\snip{asizedef}{1}{1}{% *}
 | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 31 | fun asize :: "com \<Rightarrow> nat" where | 
| 54941 | 32 | "asize SKIP = 1" | | 
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 33 | "asize (x ::= e) = 1" | | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 34 | "asize (C\<^sub>1;;C\<^sub>2) = asize C\<^sub>1 + asize C\<^sub>2" | | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 35 | "asize (IF b THEN C\<^sub>1 ELSE C\<^sub>2) = asize C\<^sub>1 + asize C\<^sub>2 + 3" | | 
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 36 | "asize (WHILE b DO C) = asize C + 3" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 37 | text_raw{*}%endsnip*}
 | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 38 | |
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 39 | text_raw{*\snip{annotatedef}{1}{1}{% *}
 | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 40 | definition shift :: "(nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a" where | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 41 | "shift f n = (\<lambda>p. f(p+n))" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 42 | |
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 43 | fun annotate :: "(nat \<Rightarrow> 'a) \<Rightarrow> com \<Rightarrow> 'a acom" where | 
| 54941 | 44 | "annotate f SKIP = SKIP {f 0}" |
 | 
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 45 | "annotate f (x ::= e) = x ::= e {f 0}" |
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 46 | "annotate f (c\<^sub>1;;c\<^sub>2) = annotate f c\<^sub>1;; annotate (shift f (asize c\<^sub>1)) c\<^sub>2" | | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 47 | "annotate f (IF b THEN c\<^sub>1 ELSE c\<^sub>2) = | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 48 |   IF b THEN {f 0} annotate (shift f 1) c\<^sub>1
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 49 |   ELSE {f(asize c\<^sub>1 + 1)} annotate (shift f (asize c\<^sub>1 + 2)) c\<^sub>2
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 50 |   {f(asize c\<^sub>1 + asize c\<^sub>2 + 2)}" |
 | 
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 51 | "annotate f (WHILE b DO c) = | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 52 |   {f 0} WHILE b DO {f 1} annotate (shift f 2) c {f(asize c + 2)}"
 | 
| 49603 | 53 | text_raw{*}%endsnip*}
 | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 54 | |
| 49603 | 55 | text_raw{*\snip{annosdef}{1}{1}{% *}
 | 
| 47613 | 56 | fun annos :: "'a acom \<Rightarrow> 'a list" where | 
| 49603 | 57 | "annos (SKIP {P}) = [P]" |
 | 
| 58 | "annos (x ::= e {P}) = [P]" |
 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 59 | "annos (C\<^sub>1;;C\<^sub>2) = annos C\<^sub>1 @ annos C\<^sub>2" | | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 60 | "annos (IF b THEN {P\<^sub>1} C\<^sub>1 ELSE {P\<^sub>2} C\<^sub>2 {Q}) =
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 61 | P\<^sub>1 # annos C\<^sub>1 @ P\<^sub>2 # annos C\<^sub>2 @ [Q]" | | 
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 62 | "annos ({I} WHILE b DO {P} C {Q}) = I # P # annos C @ [Q]"
 | 
| 49603 | 63 | text_raw{*}%endsnip*}
 | 
| 47613 | 64 | |
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 65 | definition anno :: "'a acom \<Rightarrow> nat \<Rightarrow> 'a" where | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 66 | "anno C p = annos C ! p" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 67 | |
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 68 | definition post :: "'a acom \<Rightarrow>'a" where | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 69 | "post C = last(annos C)" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 70 | |
| 49603 | 71 | text_raw{*\snip{mapacomdef}{1}{2}{% *}
 | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 72 | fun map_acom :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a acom \<Rightarrow> 'b acom" where
 | 
| 45746 | 73 | "map_acom f (SKIP {P}) = SKIP {f P}" |
 | 
| 49603 | 74 | "map_acom f (x ::= e {P}) = x ::= e {f P}" |
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 75 | "map_acom f (C\<^sub>1;;C\<^sub>2) = map_acom f C\<^sub>1;; map_acom f C\<^sub>2" | | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 76 | "map_acom f (IF b THEN {P\<^sub>1} C\<^sub>1 ELSE {P\<^sub>2} C\<^sub>2 {Q}) =
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 77 |   IF b THEN {f P\<^sub>1} map_acom f C\<^sub>1 ELSE {f P\<^sub>2} map_acom f C\<^sub>2
 | 
| 49603 | 78 |   {f Q}" |
 | 
| 79 | "map_acom f ({I} WHILE b DO {P} C {Q}) =
 | |
| 80 |   {f I} WHILE b DO {f P} map_acom f C {f Q}"
 | |
| 81 | text_raw{*}%endsnip*}
 | |
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 82 | |
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 83 | |
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 84 | lemma annos_ne: "annos C \<noteq> []" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 85 | by(induction C) auto | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 86 | |
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 87 | lemma strip_annotate[simp]: "strip(annotate f c) = c" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 88 | by(induction c arbitrary: f) auto | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 89 | |
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 90 | lemma length_annos_annotate[simp]: "length (annos (annotate f c)) = asize c" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 91 | by(induction c arbitrary: f) auto | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 92 | |
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 93 | lemma size_annos: "size(annos C) = asize(strip C)" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 94 | by(induction C)(auto) | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 95 | |
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 96 | lemma size_annos_same: "strip C1 = strip C2 \<Longrightarrow> size(annos C1) = size(annos C2)" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 97 | apply(induct C2 arbitrary: C1) | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 98 | apply(case_tac C1, simp_all)+ | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 99 | done | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 100 | |
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 101 | lemmas size_annos_same2 = eqTrueI[OF size_annos_same] | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 102 | |
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 103 | lemma anno_annotate[simp]: "p < asize c \<Longrightarrow> anno (annotate f c) p = f p" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 104 | apply(induction c arbitrary: f p) | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 105 | apply (auto simp: anno_def nth_append nth_Cons numeral_eq_Suc shift_def | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 106 | split: nat.split) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
54941diff
changeset | 107 | apply (metis add_Suc_right add_diff_inverse add.commute) | 
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 108 | apply(rule_tac f=f in arg_cong) | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 109 | apply arith | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 110 | apply (metis less_Suc_eq) | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 111 | done | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 112 | |
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 113 | lemma eq_acom_iff_strip_annos: | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 114 | "C1 = C2 \<longleftrightarrow> strip C1 = strip C2 \<and> annos C1 = annos C2" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 115 | apply(induction C1 arbitrary: C2) | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 116 | apply(case_tac C2, auto simp: size_annos_same2)+ | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 117 | done | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 118 | |
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 119 | lemma eq_acom_iff_strip_anno: | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 120 | "C1=C2 \<longleftrightarrow> strip C1 = strip C2 \<and> (\<forall>p<size(annos C1). anno C1 p = anno C2 p)" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 121 | by(auto simp add: eq_acom_iff_strip_annos anno_def | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 122 | list_eq_iff_nth_eq size_annos_same2) | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 123 | |
| 49477 | 124 | lemma post_map_acom[simp]: "post(map_acom f C) = f(post C)" | 
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 125 | by (induction C) (auto simp: post_def last_append annos_ne) | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 126 | |
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 127 | lemma strip_map_acom[simp]: "strip (map_acom f C) = strip C" | 
| 49477 | 128 | by (induction C) auto | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 129 | |
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 130 | lemma anno_map_acom: "p < size(annos C) \<Longrightarrow> anno (map_acom f C) p = f(anno C p)" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 131 | apply(induction C arbitrary: p) | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 132 | apply(auto simp: anno_def nth_append nth_Cons' size_annos) | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 133 | done | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 134 | |
| 46157 | 135 | lemma strip_eq_SKIP: | 
| 54941 | 136 |   "strip C = SKIP \<longleftrightarrow> (EX P. C = SKIP {P})"
 | 
| 49095 | 137 | by (cases C) simp_all | 
| 46157 | 138 | |
| 139 | lemma strip_eq_Assign: | |
| 49095 | 140 |   "strip C = x::=e \<longleftrightarrow> (EX P. C = x::=e {P})"
 | 
| 141 | by (cases C) simp_all | |
| 46157 | 142 | |
| 47818 | 143 | lemma strip_eq_Seq: | 
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
52019diff
changeset | 144 | "strip C = c1;;c2 \<longleftrightarrow> (EX C1 C2. C = C1;;C2 & strip C1 = c1 & strip C2 = c2)" | 
| 49095 | 145 | by (cases C) simp_all | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 146 | |
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 147 | lemma strip_eq_If: | 
| 49095 | 148 | "strip C = IF b THEN c1 ELSE c2 \<longleftrightarrow> | 
| 49477 | 149 |   (EX P1 P2 C1 C2 Q. C = IF b THEN {P1} C1 ELSE {P2} C2 {Q} & strip C1 = c1 & strip C2 = c2)"
 | 
| 49095 | 150 | by (cases C) simp_all | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 151 | |
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 152 | lemma strip_eq_While: | 
| 49095 | 153 | "strip C = WHILE b DO c1 \<longleftrightarrow> | 
| 49477 | 154 |   (EX I P C1 Q. C = {I} WHILE b DO {P} C1 {Q} & strip C1 = c1)"
 | 
| 49095 | 155 | by (cases C) simp_all | 
| 47613 | 156 | |
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 157 | lemma [simp]: "shift (\<lambda>p. a) n = (\<lambda>p. a)" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 158 | by(simp add:shift_def) | 
| 47613 | 159 | |
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 160 | lemma set_annos_anno[simp]: "set (annos (annotate (\<lambda>p. a) c)) = {a}"
 | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 161 | by(induction c) simp_all | 
| 47613 | 162 | |
| 51712 | 163 | lemma post_in_annos: "post C \<in> set(annos C)" | 
| 52019 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 164 | by(auto simp: post_def annos_ne) | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 165 | |
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 166 | lemma post_anno_asize: "post C = anno C (size(annos C) - 1)" | 
| 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 nipkow parents: 
51712diff
changeset | 167 | by(simp add: post_def last_conv_nth[OF annos_ne] anno_def) | 
| 51712 | 168 | |
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 169 | end |