| author | wenzelm | 
| Sat, 30 Nov 2013 17:51:25 +0100 | |
| changeset 54655 | 9e8189a841f7 | 
| parent 52046 | bc01725d7918 | 
| child 55599 | 6535c537b243 | 
| permissions | -rw-r--r-- | 
| 44656 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 45111 | 3 | theory Abs_Int_den1_ivl | 
| 45990 
b7b905b23b2a
incorporated More_Set and More_List into the Main body -- to be consolidated later
 haftmann parents: 
45978diff
changeset | 4 | imports Abs_Int_den1 | 
| 44656 | 5 | begin | 
| 6 | ||
| 7 | subsection "Interval Analysis" | |
| 8 | ||
| 9 | datatype ivl = I "int option" "int option" | |
| 10 | ||
| 44932 | 11 | text{* We assume an important invariant: arithmetic operations are never
 | 
| 12 | applied to empty intervals @{term"I (Some i) (Some j)"} with @{term"j <
 | |
| 13 | i"}. This avoids special cases. Why can we assume this? Because an empty | |
| 14 | interval of values for a variable means that the current program point is | |
| 15 | unreachable. But this should actually translate into the bottom state, not a | |
| 16 | state where some variables have empty intervals. *} | |
| 17 | ||
| 44656 | 18 | definition "rep_ivl i = | 
| 19 |  (case i of I (Some l) (Some h) \<Rightarrow> {l..h} | I (Some l) None \<Rightarrow> {l..}
 | |
| 20 |   | I None (Some h) \<Rightarrow> {..h} | I None None \<Rightarrow> UNIV)"
 | |
| 21 | ||
| 44932 | 22 | definition "num_ivl n = I (Some n) (Some n)" | 
| 23 | ||
| 45978 | 24 | definition | 
| 46028 
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
 haftmann parents: 
45990diff
changeset | 25 | [code_abbrev]: "contained_in i k \<longleftrightarrow> k \<in> rep_ivl i" | 
| 45978 | 26 | |
| 27 | lemma contained_in_simps [code]: | |
| 28 | "contained_in (I (Some l) (Some h)) k \<longleftrightarrow> l \<le> k \<and> k \<le> h" | |
| 29 | "contained_in (I (Some l) None) k \<longleftrightarrow> l \<le> k" | |
| 30 | "contained_in (I None (Some h)) k \<longleftrightarrow> k \<le> h" | |
| 31 | "contained_in (I None None) k \<longleftrightarrow> True" | |
| 32 | by (simp_all add: contained_in_def rep_ivl_def) | |
| 33 | ||
| 44656 | 34 | instantiation option :: (plus)plus | 
| 35 | begin | |
| 36 | ||
| 37 | fun plus_option where | |
| 38 | "Some x + Some y = Some(x+y)" | | |
| 39 | "_ + _ = None" | |
| 40 | ||
| 41 | instance proof qed | |
| 42 | ||
| 43 | end | |
| 44 | ||
| 45 | definition empty where "empty = I (Some 1) (Some 0)" | |
| 46 | ||
| 47 | fun is_empty where | |
| 48 | "is_empty(I (Some l) (Some h)) = (h<l)" | | |
| 49 | "is_empty _ = False" | |
| 50 | ||
| 51 | lemma [simp]: "is_empty(I l h) = | |
| 52 | (case l of Some l \<Rightarrow> (case h of Some h \<Rightarrow> h<l | None \<Rightarrow> False) | None \<Rightarrow> False)" | |
| 53 | by(auto split:option.split) | |
| 54 | ||
| 55 | lemma [simp]: "is_empty i \<Longrightarrow> rep_ivl i = {}"
 | |
| 56 | by(auto simp add: rep_ivl_def split: ivl.split option.split) | |
| 57 | ||
| 44932 | 58 | definition "plus_ivl i1 i2 = ((*if is_empty i1 | is_empty i2 then empty else*) | 
| 59 | case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow> I (l1+l2) (h1+h2))" | |
| 44656 | 60 | |
| 61 | instantiation ivl :: SL_top | |
| 62 | begin | |
| 63 | ||
| 64 | definition le_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> bool" where | |
| 65 | "le_option pos x y = | |
| 44932 | 66 | (case x of (Some i) \<Rightarrow> (case y of Some j \<Rightarrow> i\<le>j | None \<Rightarrow> pos) | 
| 67 | | None \<Rightarrow> (case y of Some j \<Rightarrow> \<not>pos | None \<Rightarrow> True))" | |
| 44656 | 68 | |
| 69 | fun le_aux where | |
| 70 | "le_aux (I l1 h1) (I l2 h2) = (le_option False l2 l1 & le_option True h1 h2)" | |
| 71 | ||
| 72 | definition le_ivl where | |
| 73 | "i1 \<sqsubseteq> i2 = | |
| 74 | (if is_empty i1 then True else | |
| 75 | if is_empty i2 then False else le_aux i1 i2)" | |
| 76 | ||
| 77 | definition min_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> int option" where | |
| 78 | "min_option pos o1 o2 = (if le_option pos o1 o2 then o1 else o2)" | |
| 79 | ||
| 80 | definition max_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> int option" where | |
| 81 | "max_option pos o1 o2 = (if le_option pos o1 o2 then o2 else o1)" | |
| 82 | ||
| 83 | definition "i1 \<squnion> i2 = | |
| 84 | (if is_empty i1 then i2 else if is_empty i2 then i1 | |
| 85 | else case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow> | |
| 86 | I (min_option False l1 l2) (max_option True h1 h2))" | |
| 87 | ||
| 88 | definition "Top = I None None" | |
| 89 | ||
| 90 | instance | |
| 91 | proof | |
| 92 | case goal1 thus ?case | |
| 93 | by(cases x, simp add: le_ivl_def le_option_def split: option.split) | |
| 94 | next | |
| 95 | case goal2 thus ?case | |
| 96 | by(cases x, cases y, cases z, auto simp: le_ivl_def le_option_def split: option.splits if_splits) | |
| 97 | next | |
| 98 | case goal3 thus ?case | |
| 99 | by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits) | |
| 100 | next | |
| 101 | case goal4 thus ?case | |
| 102 | by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits) | |
| 103 | next | |
| 104 | case goal5 thus ?case | |
| 105 | by(cases x, cases y, cases z, auto simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits if_splits) | |
| 106 | next | |
| 107 | case goal6 thus ?case | |
| 108 | by(cases x, simp add: Top_ivl_def le_ivl_def le_option_def split: option.split) | |
| 109 | qed | |
| 110 | ||
| 111 | end | |
| 112 | ||
| 113 | ||
| 114 | instantiation ivl :: L_top_bot | |
| 115 | begin | |
| 116 | ||
| 44932 | 117 | definition "i1 \<sqinter> i2 = (if is_empty i1 \<or> is_empty i2 then empty else | 
| 118 | case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow> | |
| 119 | I (max_option False l1 l2) (min_option True h1 h2))" | |
| 44656 | 120 | |
| 121 | definition "Bot = empty" | |
| 122 | ||
| 123 | instance | |
| 124 | proof | |
| 125 | case goal1 thus ?case | |
| 126 | by (simp add:meet_ivl_def empty_def meet_ivl_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits) | |
| 127 | next | |
| 128 | case goal2 thus ?case | |
| 129 | by (simp add:meet_ivl_def empty_def meet_ivl_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits) | |
| 130 | next | |
| 131 | case goal3 thus ?case | |
| 132 | by (cases x, cases y, cases z, auto simp add: le_ivl_def meet_ivl_def empty_def le_option_def max_option_def min_option_def split: option.splits if_splits) | |
| 133 | next | |
| 134 | case goal4 show ?case by(cases x, simp add: Bot_ivl_def empty_def le_ivl_def) | |
| 135 | qed | |
| 136 | ||
| 137 | end | |
| 138 | ||
| 139 | instantiation option :: (minus)minus | |
| 140 | begin | |
| 141 | ||
| 142 | fun minus_option where | |
| 143 | "Some x - Some y = Some(x-y)" | | |
| 144 | "_ - _ = None" | |
| 145 | ||
| 146 | instance proof qed | |
| 147 | ||
| 148 | end | |
| 149 | ||
| 44932 | 150 | definition "minus_ivl i1 i2 = ((*if is_empty i1 | is_empty i2 then empty else*) | 
| 151 | case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow> I (l1-h2) (h1-l2))" | |
| 44656 | 152 | |
| 153 | lemma rep_minus_ivl: | |
| 154 | "n1 : rep_ivl i1 \<Longrightarrow> n2 : rep_ivl i2 \<Longrightarrow> n1-n2 : rep_ivl(minus_ivl i1 i2)" | |
| 155 | by(auto simp add: minus_ivl_def rep_ivl_def split: ivl.splits option.splits) | |
| 156 | ||
| 157 | ||
| 45023 | 158 | definition "filter_plus_ivl i i1 i2 = ((*if is_empty i then empty else*) | 
| 44932 | 159 | i1 \<sqinter> minus_ivl i i2, i2 \<sqinter> minus_ivl i i1)" | 
| 44656 | 160 | |
| 45023 | 161 | fun filter_less_ivl :: "bool \<Rightarrow> ivl \<Rightarrow> ivl \<Rightarrow> ivl * ivl" where | 
| 162 | "filter_less_ivl res (I l1 h1) (I l2 h2) = | |
| 45019 | 163 | ((*if is_empty(I l1 h1) \<or> is_empty(I l2 h2) then (empty, empty) else*) | 
| 164 | if res | |
| 44932 | 165 | then (I l1 (min_option True h1 (h2 - Some 1)), | 
| 166 | I (max_option False (l1 + Some 1) l2) h2) | |
| 44656 | 167 | else (I (max_option False l1 l2) h1, I l2 (min_option True h1 h2)))" | 
| 168 | ||
| 169 | interpretation Rep rep_ivl | |
| 170 | proof | |
| 171 | case goal1 thus ?case | |
| 172 | by(auto simp: rep_ivl_def le_ivl_def le_option_def split: ivl.split option.split if_splits) | |
| 173 | qed | |
| 174 | ||
| 44932 | 175 | interpretation Val_abs rep_ivl num_ivl plus_ivl | 
| 44656 | 176 | proof | 
| 44932 | 177 | case goal1 thus ?case by(simp add: rep_ivl_def num_ivl_def) | 
| 44656 | 178 | next | 
| 179 | case goal2 thus ?case | |
| 180 | by(auto simp add: rep_ivl_def plus_ivl_def split: ivl.split option.splits) | |
| 181 | qed | |
| 182 | ||
| 183 | interpretation Rep1 rep_ivl | |
| 184 | proof | |
| 185 | case goal1 thus ?case | |
| 186 | by(auto simp add: rep_ivl_def meet_ivl_def empty_def min_option_def max_option_def split: ivl.split option.split) | |
| 187 | next | |
| 188 | case goal2 show ?case by(auto simp add: Bot_ivl_def rep_ivl_def empty_def) | |
| 189 | qed | |
| 190 | ||
| 44932 | 191 | interpretation | 
| 45023 | 192 | Val_abs1 rep_ivl num_ivl plus_ivl filter_plus_ivl filter_less_ivl | 
| 44656 | 193 | proof | 
| 194 | case goal1 thus ?case | |
| 45023 | 195 | by(auto simp add: filter_plus_ivl_def) | 
| 44656 | 196 | (metis rep_minus_ivl add_diff_cancel add_commute)+ | 
| 197 | next | |
| 198 | case goal2 thus ?case | |
| 199 | by(cases a1, cases a2, | |
| 200 | auto simp: rep_ivl_def min_option_def max_option_def le_option_def split: if_splits option.splits) | |
| 201 | qed | |
| 202 | ||
| 44932 | 203 | interpretation | 
| 45023 | 204 | Abs_Int1 rep_ivl num_ivl plus_ivl filter_plus_ivl filter_less_ivl "(iter' 3)" | 
| 44656 | 205 | defines afilter_ivl is afilter | 
| 206 | and bfilter_ivl is bfilter | |
| 207 | and AI_ivl is AI | |
| 208 | and aval_ivl is aval' | |
| 44944 | 209 | proof qed (auto simp: iter'_pfp_above) | 
| 44656 | 210 | |
| 211 | ||
| 212 | fun list_up where | |
| 213 | "list_up bot = bot" | | |
| 214 | "list_up (Up x) = Up(list x)" | |
| 215 | ||
| 216 | value [code] "list_up(afilter_ivl (N 5) (I (Some 4) (Some 5)) Top)" | |
| 217 | value [code] "list_up(afilter_ivl (N 5) (I (Some 4) (Some 4)) Top)" | |
| 218 | value [code] "list_up(afilter_ivl (V ''x'') (I (Some 4) (Some 4)) | |
| 219 |  (Up(FunDom(Top(''x'':=I (Some 5) (Some 6))) [''x''])))"
 | |
| 220 | value [code] "list_up(afilter_ivl (V ''x'') (I (Some 4) (Some 5)) | |
| 221 |  (Up(FunDom(Top(''x'':=I (Some 5) (Some 6))) [''x''])))"
 | |
| 222 | value [code] "list_up(afilter_ivl (Plus (V ''x'') (V ''x'')) (I (Some 0) (Some 10)) | |
| 223 |   (Up(FunDom(Top(''x'':= I (Some 0) (Some 100)))[''x''])))"
 | |
| 224 | value [code] "list_up(afilter_ivl (Plus (V ''x'') (N 7)) (I (Some 0) (Some 10)) | |
| 225 |   (Up(FunDom(Top(''x'':= I (Some 0) (Some 100)))[''x''])))"
 | |
| 226 | ||
| 227 | value [code] "list_up(bfilter_ivl (Less (V ''x'') (V ''x'')) True | |
| 228 |   (Up(FunDom(Top(''x'':= I (Some 0) (Some 0)))[''x''])))"
 | |
| 229 | value [code] "list_up(bfilter_ivl (Less (V ''x'') (V ''x'')) True | |
| 230 |   (Up(FunDom(Top(''x'':= I (Some 0) (Some 2)))[''x''])))"
 | |
| 231 | value [code] "list_up(bfilter_ivl (Less (V ''x'') (Plus (N 10) (V ''y''))) True | |
| 232 |   (Up(FunDom(Top(''x'':= I (Some 15) (Some 20),''y'':= I (Some 5) (Some 7)))[''x'', ''y''])))"
 | |
| 233 | ||
| 234 | definition "test_ivl1 = | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
46028diff
changeset | 235 | ''y'' ::= N 7;; | 
| 44656 | 236 | IF Less (V ''x'') (V ''y'') | 
| 237 | THEN ''y'' ::= Plus (V ''y'') (V ''x'') | |
| 238 | ELSE ''x'' ::= Plus (V ''x'') (V ''y'')" | |
| 239 | value [code] "list_up(AI_ivl test_ivl1 Top)" | |
| 240 | ||
| 241 | value "list_up (AI_ivl test3_const Top)" | |
| 242 | ||
| 243 | value "list_up (AI_ivl test5_const Top)" | |
| 244 | ||
| 245 | value "list_up (AI_ivl test6_const Top)" | |
| 246 | ||
| 247 | definition "test2_ivl = | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
46028diff
changeset | 248 | ''y'' ::= N 7;; | 
| 44656 | 249 | WHILE Less (V ''x'') (N 100) DO ''y'' ::= Plus (V ''y'') (N 1)" | 
| 250 | value [code] "list_up(AI_ivl test2_ivl Top)" | |
| 251 | ||
| 44932 | 252 | definition "test3_ivl = | 
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
46028diff
changeset | 253 | ''x'' ::= N 0;; ''y'' ::= N 100;; ''z'' ::= Plus (V ''x'') (V ''y'');; | 
| 44932 | 254 | WHILE Less (V ''x'') (N 11) | 
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
46028diff
changeset | 255 |  DO (''x'' ::= Plus (V ''x'') (N 1);; ''y'' ::= Plus (V ''y'') (N -1))"
 | 
| 44932 | 256 | value [code] "list_up(AI_ivl test3_ivl Top)" | 
| 257 | ||
| 258 | definition "test4_ivl = | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
46028diff
changeset | 259 | ''x'' ::= N 0;; ''y'' ::= N 0;; | 
| 44932 | 260 | WHILE Less (V ''x'') (N 1001) | 
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
46028diff
changeset | 261 |  DO (''y'' ::= V ''x'';; ''x'' ::= Plus (V ''x'') (N 1))"
 | 
| 44932 | 262 | value [code] "list_up(AI_ivl test4_ivl Top)" | 
| 263 | ||
| 45020 | 264 | text{* Nontermination not detected: *}
 | 
| 265 | definition "test5_ivl = | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
46028diff
changeset | 266 | ''x'' ::= N 0;; | 
| 45020 | 267 | WHILE Less (V ''x'') (N 1) DO ''x'' ::= Plus (V ''x'') (N -1)" | 
| 268 | value [code] "list_up(AI_ivl test5_ivl Top)" | |
| 269 | ||
| 44656 | 270 | end |