| author | wenzelm | 
| Wed, 12 May 2021 16:47:52 +0200 | |
| changeset 73687 | 54fe8cc0e1c6 | 
| parent 69597 | ff784d5a5bfb | 
| child 82774 | 2865a6618cba | 
| permissions | -rw-r--r-- | 
| 44236 
b73b7832b384
moved theory Nested_Environment to HOL-Unix (a bit too specific for HOL-Library);
 wenzelm parents: 
38857diff
changeset | 1 | (* Title: HOL/Unix/Nested_Environment.thy | 
| 10943 | 2 | Author: Markus Wenzel, TU Muenchen | 
| 3 | *) | |
| 4 | ||
| 58889 | 5 | section \<open>Nested environments\<close> | 
| 10943 | 6 | |
| 15131 | 7 | theory Nested_Environment | 
| 30663 
0b6aff7451b2
Main is (Complex_Main) base entry point in library theories
 haftmann parents: 
28562diff
changeset | 8 | imports Main | 
| 15131 | 9 | begin | 
| 10943 | 10 | |
| 58613 | 11 | text \<open> | 
| 61893 | 12 |   Consider a partial function @{term [source] "e :: 'a \<Rightarrow> 'b option"}; this may
 | 
| 69597 | 13 | be understood as an \<^emph>\<open>environment\<close> mapping indexes \<^typ>\<open>'a\<close> to optional | 
| 14 | entry values \<^typ>\<open>'b\<close> (cf.\ the basic theory \<open>Map\<close> of Isabelle/HOL). This | |
| 61893 | 15 | basic idea is easily generalized to that of a \<^emph>\<open>nested environment\<close>, where | 
| 16 | entries may be either basic values or again proper environments. Then each | |
| 17 | entry is accessed by a \<^emph>\<open>path\<close>, i.e.\ a list of indexes leading to its | |
| 10948 | 18 | position within the structure. | 
| 58613 | 19 | \<close> | 
| 10943 | 20 | |
| 58310 | 21 | datatype (dead 'a, dead 'b, dead 'c) env = | 
| 10943 | 22 | Val 'a | 
| 44703 | 23 |   | Env 'b  "'c \<Rightarrow> ('a, 'b, 'c) env option"
 | 
| 10943 | 24 | |
| 58613 | 25 | text \<open> | 
| 61893 | 26 | \<^medskip> | 
| 69597 | 27 |   In the type \<^typ>\<open>('a, 'b, 'c) env\<close> the parameter \<^typ>\<open>'a\<close> refers to
 | 
| 28 | basic values (occurring in terminal positions), type \<^typ>\<open>'b\<close> to values | |
| 29 | associated with proper (inner) environments, and type \<^typ>\<open>'c\<close> with the | |
| 61893 | 30 | index type for branching. Note that there is no restriction on any of these | 
| 31 | types. In particular, arbitrary branching may yield rather large | |
| 32 | (transfinite) tree structures. | |
| 58613 | 33 | \<close> | 
| 10943 | 34 | |
| 35 | ||
| 58613 | 36 | subsection \<open>The lookup operation\<close> | 
| 10943 | 37 | |
| 58613 | 38 | text \<open> | 
| 61893 | 39 | Lookup in nested environments works by following a given path of index | 
| 40 | elements, leading to an optional result (a terminal value or nested | |
| 41 | environment). A \<^emph>\<open>defined position\<close> within a nested environment is one where | |
| 69597 | 42 | \<^term>\<open>lookup\<close> at its path does not yield \<^term>\<open>None\<close>. | 
| 58613 | 43 | \<close> | 
| 10943 | 44 | |
| 44703 | 45 | primrec lookup :: "('a, 'b, 'c) env \<Rightarrow> 'c list \<Rightarrow> ('a, 'b, 'c) env option"
 | 
| 46 |   and lookup_option :: "('a, 'b, 'c) env option \<Rightarrow> 'c list \<Rightarrow> ('a, 'b, 'c) env option"
 | |
| 44601 | 47 | where | 
| 34941 | 48 | "lookup (Val a) xs = (if xs = [] then Some (Val a) else None)" | 
| 49 | | "lookup (Env b es) xs = | |
| 50 | (case xs of | |
| 44703 | 51 | [] \<Rightarrow> Some (Env b es) | 
| 52 | | y # ys \<Rightarrow> lookup_option (es y) ys)" | |
| 34941 | 53 | | "lookup_option None xs = None" | 
| 54 | | "lookup_option (Some e) xs = lookup e xs" | |
| 10943 | 55 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
34941diff
changeset | 56 | hide_const lookup_option | 
| 10943 | 57 | |
| 58613 | 58 | text \<open> | 
| 61893 | 59 | \<^medskip> | 
| 69597 | 60 | The characteristic cases of \<^term>\<open>lookup\<close> are expressed by the following | 
| 61893 | 61 | equalities. | 
| 58613 | 62 | \<close> | 
| 10943 | 63 | |
| 64 | theorem lookup_nil: "lookup e [] = Some e" | |
| 65 | by (cases e) simp_all | |
| 66 | ||
| 67 | theorem lookup_val_cons: "lookup (Val a) (x # xs) = None" | |
| 68 | by simp | |
| 69 | ||
| 70 | theorem lookup_env_cons: | |
| 71 | "lookup (Env b es) (x # xs) = | |
| 72 | (case es x of | |
| 44703 | 73 | None \<Rightarrow> None | 
| 74 | | Some e \<Rightarrow> lookup e xs)" | |
| 10943 | 75 | by (cases "es x") simp_all | 
| 76 | ||
| 58261 | 77 | lemmas lookup.simps [simp del] lookup_option.simps [simp del] | 
| 10943 | 78 | and lookup_simps [simp] = lookup_nil lookup_val_cons lookup_env_cons | 
| 79 | ||
| 80 | theorem lookup_eq: | |
| 81 | "lookup env xs = | |
| 82 | (case xs of | |
| 44703 | 83 | [] \<Rightarrow> Some env | 
| 84 | | x # xs \<Rightarrow> | |
| 10943 | 85 | (case env of | 
| 44703 | 86 | Val a \<Rightarrow> None | 
| 87 | | Env b es \<Rightarrow> | |
| 10943 | 88 | (case es x of | 
| 44703 | 89 | None \<Rightarrow> None | 
| 90 | | Some e \<Rightarrow> lookup e xs)))" | |
| 10943 | 91 | by (simp split: list.split env.split) | 
| 92 | ||
| 58613 | 93 | text \<open> | 
| 61893 | 94 | \<^medskip> | 
| 69597 | 95 | Displaced \<^term>\<open>lookup\<close> operations, relative to a certain base path prefix, | 
| 61893 | 96 | may be reduced as follows. There are two cases, depending whether the | 
| 97 | environment actually extends far enough to follow the base path. | |
| 58613 | 98 | \<close> | 
| 10943 | 99 | |
| 100 | theorem lookup_append_none: | |
| 18153 | 101 | assumes "lookup env xs = None" | 
| 102 | shows "lookup env (xs @ ys) = None" | |
| 23394 | 103 | using assms | 
| 20503 | 104 | proof (induct xs arbitrary: env) | 
| 18153 | 105 | case Nil | 
| 106 | then have False by simp | |
| 107 | then show ?case .. | |
| 108 | next | |
| 109 | case (Cons x xs) | |
| 110 | show ?case | |
| 111 | proof (cases env) | |
| 112 | case Val | |
| 113 | then show ?thesis by simp | |
| 10943 | 114 | next | 
| 18153 | 115 | case (Env b es) | 
| 116 | show ?thesis | |
| 117 | proof (cases "es x") | |
| 118 | case None | |
| 119 | with Env show ?thesis by simp | |
| 10943 | 120 | next | 
| 18153 | 121 | case (Some e) | 
| 58613 | 122 | note es = \<open>es x = Some e\<close> | 
| 10943 | 123 | show ?thesis | 
| 18153 | 124 | proof (cases "lookup e xs") | 
| 125 | case None | |
| 126 | then have "lookup e (xs @ ys) = None" by (rule Cons.hyps) | |
| 127 | with Env Some show ?thesis by simp | |
| 10943 | 128 | next | 
| 18153 | 129 | case Some | 
| 130 | with Env es have False using Cons.prems by simp | |
| 131 | then show ?thesis .. | |
| 10943 | 132 | qed | 
| 133 | qed | |
| 18153 | 134 | qed | 
| 10943 | 135 | qed | 
| 136 | ||
| 137 | theorem lookup_append_some: | |
| 18153 | 138 | assumes "lookup env xs = Some e" | 
| 139 | shows "lookup env (xs @ ys) = lookup e ys" | |
| 23394 | 140 | using assms | 
| 20503 | 141 | proof (induct xs arbitrary: env e) | 
| 18153 | 142 | case Nil | 
| 143 | then have "env = e" by simp | |
| 144 | then show "lookup env ([] @ ys) = lookup e ys" by simp | |
| 145 | next | |
| 146 | case (Cons x xs) | |
| 58613 | 147 | note asm = \<open>lookup env (x # xs) = Some e\<close> | 
| 18153 | 148 | show "lookup env ((x # xs) @ ys) = lookup e ys" | 
| 149 | proof (cases env) | |
| 150 | case (Val a) | |
| 151 | with asm have False by simp | |
| 152 | then show ?thesis .. | |
| 10943 | 153 | next | 
| 18153 | 154 | case (Env b es) | 
| 155 | show ?thesis | |
| 156 | proof (cases "es x") | |
| 157 | case None | |
| 158 | with asm Env have False by simp | |
| 159 | then show ?thesis .. | |
| 10943 | 160 | next | 
| 18153 | 161 | case (Some e') | 
| 58613 | 162 | note es = \<open>es x = Some e'\<close> | 
| 10943 | 163 | show ?thesis | 
| 18153 | 164 | proof (cases "lookup e' xs") | 
| 165 | case None | |
| 166 | with asm Env es have False by simp | |
| 167 | then show ?thesis .. | |
| 10943 | 168 | next | 
| 18153 | 169 | case Some | 
| 170 | with asm Env es have "lookup e' xs = Some e" | |
| 171 | by simp | |
| 172 | then have "lookup e' (xs @ ys) = lookup e ys" by (rule Cons.hyps) | |
| 173 | with Env es show ?thesis by simp | |
| 10943 | 174 | qed | 
| 175 | qed | |
| 18153 | 176 | qed | 
| 10943 | 177 | qed | 
| 178 | ||
| 58613 | 179 | text \<open> | 
| 61893 | 180 | \<^medskip> | 
| 69597 | 181 | Successful \<^term>\<open>lookup\<close> deeper down an environment structure means we are | 
| 61893 | 182 | able to peek further up as well. Note that this is basically just the | 
| 183 |   contrapositive statement of @{thm [source] lookup_append_none} above.
 | |
| 58613 | 184 | \<close> | 
| 10943 | 185 | |
| 186 | theorem lookup_some_append: | |
| 18153 | 187 | assumes "lookup env (xs @ ys) = Some e" | 
| 188 | shows "\<exists>e. lookup env xs = Some e" | |
| 10943 | 189 | proof - | 
| 23394 | 190 | from assms have "lookup env (xs @ ys) \<noteq> None" by simp | 
| 18153 | 191 | then have "lookup env xs \<noteq> None" | 
| 10943 | 192 | by (rule contrapos_nn) (simp only: lookup_append_none) | 
| 18576 | 193 | then show ?thesis by (simp) | 
| 10943 | 194 | qed | 
| 195 | ||
| 58613 | 196 | text \<open> | 
| 69597 | 197 | The subsequent statement describes in more detail how a successful \<^term>\<open>lookup\<close> with a non-empty path results in a certain situation at any upper | 
| 61893 | 198 | position. | 
| 58613 | 199 | \<close> | 
| 10943 | 200 | |
| 18153 | 201 | theorem lookup_some_upper: | 
| 202 | assumes "lookup env (xs @ y # ys) = Some e" | |
| 203 | shows "\<exists>b' es' env'. | |
| 204 | lookup env xs = Some (Env b' es') \<and> | |
| 205 | es' y = Some env' \<and> | |
| 206 | lookup env' ys = Some e" | |
| 23394 | 207 | using assms | 
| 20503 | 208 | proof (induct xs arbitrary: env e) | 
| 18153 | 209 | case Nil | 
| 210 | from Nil.prems have "lookup env (y # ys) = Some e" | |
| 211 | by simp | |
| 212 | then obtain b' es' env' where | |
| 213 | env: "env = Env b' es'" and | |
| 214 | es': "es' y = Some env'" and | |
| 215 | look': "lookup env' ys = Some e" | |
| 216 | by (auto simp add: lookup_eq split: option.splits env.splits) | |
| 217 | from env have "lookup env [] = Some (Env b' es')" by simp | |
| 218 | with es' look' show ?case by blast | |
| 219 | next | |
| 220 | case (Cons x xs) | |
| 221 | from Cons.prems | |
| 222 | obtain b' es' env' where | |
| 223 | env: "env = Env b' es'" and | |
| 224 | es': "es' x = Some env'" and | |
| 225 | look': "lookup env' (xs @ y # ys) = Some e" | |
| 226 | by (auto simp add: lookup_eq split: option.splits env.splits) | |
| 227 | from Cons.hyps [OF look'] obtain b'' es'' env'' where | |
| 228 | upper': "lookup env' xs = Some (Env b'' es'')" and | |
| 229 | es'': "es'' y = Some env''" and | |
| 230 | look'': "lookup env'' ys = Some e" | |
| 231 | by blast | |
| 232 | from env es' upper' have "lookup env (x # xs) = Some (Env b'' es'')" | |
| 233 | by simp | |
| 234 | with es'' look'' show ?case by blast | |
| 10943 | 235 | qed | 
| 236 | ||
| 237 | ||
| 58613 | 238 | subsection \<open>The update operation\<close> | 
| 10943 | 239 | |
| 58613 | 240 | text \<open> | 
| 61893 | 241 | Update at a certain position in a nested environment may either delete an | 
| 242 | existing entry, or overwrite an existing one. Note that update at undefined | |
| 243 | positions is simple absorbed, i.e.\ the environment is left unchanged. | |
| 58613 | 244 | \<close> | 
| 10943 | 245 | |
| 44703 | 246 | primrec update :: "'c list \<Rightarrow> ('a, 'b, 'c) env option \<Rightarrow>
 | 
| 247 |     ('a, 'b, 'c) env \<Rightarrow> ('a, 'b, 'c) env"
 | |
| 248 |   and update_option :: "'c list \<Rightarrow> ('a, 'b, 'c) env option \<Rightarrow>
 | |
| 249 |     ('a, 'b, 'c) env option \<Rightarrow> ('a, 'b, 'c) env option"
 | |
| 44601 | 250 | where | 
| 251 | "update xs opt (Val a) = | |
| 44703 | 252 | (if xs = [] then (case opt of None \<Rightarrow> Val a | Some e \<Rightarrow> e) | 
| 44601 | 253 | else Val a)" | 
| 254 | | "update xs opt (Env b es) = | |
| 255 | (case xs of | |
| 44703 | 256 | [] \<Rightarrow> (case opt of None \<Rightarrow> Env b es | Some e \<Rightarrow> e) | 
| 257 | | y # ys \<Rightarrow> Env b (es (y := update_option ys opt (es y))))" | |
| 44601 | 258 | | "update_option xs opt None = | 
| 259 | (if xs = [] then opt else None)" | |
| 260 | | "update_option xs opt (Some e) = | |
| 261 | (if xs = [] then opt else Some (update xs opt e))" | |
| 10943 | 262 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
34941diff
changeset | 263 | hide_const update_option | 
| 10943 | 264 | |
| 58613 | 265 | text \<open> | 
| 61893 | 266 | \<^medskip> | 
| 69597 | 267 | The characteristic cases of \<^term>\<open>update\<close> are expressed by the following | 
| 61893 | 268 | equalities. | 
| 58613 | 269 | \<close> | 
| 10943 | 270 | |
| 271 | theorem update_nil_none: "update [] None env = env" | |
| 272 | by (cases env) simp_all | |
| 273 | ||
| 274 | theorem update_nil_some: "update [] (Some e) env = e" | |
| 275 | by (cases env) simp_all | |
| 276 | ||
| 277 | theorem update_cons_val: "update (x # xs) opt (Val a) = Val a" | |
| 278 | by simp | |
| 279 | ||
| 280 | theorem update_cons_nil_env: | |
| 281 | "update [x] opt (Env b es) = Env b (es (x := opt))" | |
| 282 | by (cases "es x") simp_all | |
| 283 | ||
| 284 | theorem update_cons_cons_env: | |
| 285 | "update (x # y # ys) opt (Env b es) = | |
| 286 | Env b (es (x := | |
| 287 | (case es x of | |
| 44703 | 288 | None \<Rightarrow> None | 
| 289 | | Some e \<Rightarrow> Some (update (y # ys) opt e))))" | |
| 10943 | 290 | by (cases "es x") simp_all | 
| 291 | ||
| 58261 | 292 | lemmas update.simps [simp del] update_option.simps [simp del] | 
| 10943 | 293 | and update_simps [simp] = update_nil_none update_nil_some | 
| 294 | update_cons_val update_cons_nil_env update_cons_cons_env | |
| 295 | ||
| 296 | lemma update_eq: | |
| 297 | "update xs opt env = | |
| 298 | (case xs of | |
| 44703 | 299 | [] \<Rightarrow> | 
| 10943 | 300 | (case opt of | 
| 44703 | 301 | None \<Rightarrow> env | 
| 302 | | Some e \<Rightarrow> e) | |
| 303 | | x # xs \<Rightarrow> | |
| 10943 | 304 | (case env of | 
| 44703 | 305 | Val a \<Rightarrow> Val a | 
| 306 | | Env b es \<Rightarrow> | |
| 10943 | 307 | (case xs of | 
| 44703 | 308 | [] \<Rightarrow> Env b (es (x := opt)) | 
| 309 | | y # ys \<Rightarrow> | |
| 10943 | 310 | Env b (es (x := | 
| 311 | (case es x of | |
| 44703 | 312 | None \<Rightarrow> None | 
| 313 | | Some e \<Rightarrow> Some (update (y # ys) opt e)))))))" | |
| 10943 | 314 | by (simp split: list.split env.split option.split) | 
| 315 | ||
| 58613 | 316 | text \<open> | 
| 61893 | 317 | \<^medskip> | 
| 69597 | 318 | The most basic correspondence of \<^term>\<open>lookup\<close> and \<^term>\<open>update\<close> states | 
| 319 | that after \<^term>\<open>update\<close> at a defined position, subsequent \<^term>\<open>lookup\<close> | |
| 61893 | 320 | operations would yield the new value. | 
| 58613 | 321 | \<close> | 
| 10943 | 322 | |
| 323 | theorem lookup_update_some: | |
| 18153 | 324 | assumes "lookup env xs = Some e" | 
| 325 | shows "lookup (update xs (Some env') env) xs = Some env'" | |
| 23394 | 326 | using assms | 
| 20503 | 327 | proof (induct xs arbitrary: env e) | 
| 18153 | 328 | case Nil | 
| 329 | then have "env = e" by simp | |
| 330 | then show ?case by simp | |
| 331 | next | |
| 332 | case (Cons x xs) | |
| 333 | note hyp = Cons.hyps | |
| 58613 | 334 | and asm = \<open>lookup env (x # xs) = Some e\<close> | 
| 18153 | 335 | show ?case | 
| 336 | proof (cases env) | |
| 337 | case (Val a) | |
| 338 | with asm have False by simp | |
| 339 | then show ?thesis .. | |
| 10943 | 340 | next | 
| 18153 | 341 | case (Env b es) | 
| 342 | show ?thesis | |
| 343 | proof (cases "es x") | |
| 344 | case None | |
| 345 | with asm Env have False by simp | |
| 346 | then show ?thesis .. | |
| 10943 | 347 | next | 
| 18153 | 348 | case (Some e') | 
| 58613 | 349 | note es = \<open>es x = Some e'\<close> | 
| 10943 | 350 | show ?thesis | 
| 18153 | 351 | proof (cases xs) | 
| 352 | case Nil | |
| 353 | with Env show ?thesis by simp | |
| 10943 | 354 | next | 
| 18153 | 355 | case (Cons x' xs') | 
| 356 | from asm Env es have "lookup e' xs = Some e" by simp | |
| 357 | then have "lookup (update xs (Some env') e') xs = Some env'" by (rule hyp) | |
| 358 | with Env es Cons show ?thesis by simp | |
| 10943 | 359 | qed | 
| 360 | qed | |
| 18153 | 361 | qed | 
| 10943 | 362 | qed | 
| 363 | ||
| 58613 | 364 | text \<open> | 
| 61893 | 365 | \<^medskip> | 
| 69597 | 366 | The properties of displaced \<^term>\<open>update\<close> operations are analogous to those | 
| 367 | of \<^term>\<open>lookup\<close> above. There are two cases: below an undefined position | |
| 368 | \<^term>\<open>update\<close> is absorbed altogether, and below a defined positions \<^term>\<open>update\<close> affects subsequent \<^term>\<open>lookup\<close> operations in the obvious way. | |
| 58613 | 369 | \<close> | 
| 10943 | 370 | |
| 371 | theorem update_append_none: | |
| 18153 | 372 | assumes "lookup env xs = None" | 
| 373 | shows "update (xs @ y # ys) opt env = env" | |
| 23394 | 374 | using assms | 
| 20503 | 375 | proof (induct xs arbitrary: env) | 
| 18153 | 376 | case Nil | 
| 377 | then have False by simp | |
| 378 | then show ?case .. | |
| 379 | next | |
| 380 | case (Cons x xs) | |
| 381 | note hyp = Cons.hyps | |
| 58613 | 382 | and asm = \<open>lookup env (x # xs) = None\<close> | 
| 18153 | 383 | show "update ((x # xs) @ y # ys) opt env = env" | 
| 384 | proof (cases env) | |
| 385 | case (Val a) | |
| 386 | then show ?thesis by simp | |
| 10943 | 387 | next | 
| 18153 | 388 | case (Env b es) | 
| 389 | show ?thesis | |
| 390 | proof (cases "es x") | |
| 391 | case None | |
| 58613 | 392 | note es = \<open>es x = None\<close> | 
| 10943 | 393 | show ?thesis | 
| 18153 | 394 | by (cases xs) (simp_all add: es Env fun_upd_idem_iff) | 
| 395 | next | |
| 396 | case (Some e) | |
| 58613 | 397 | note es = \<open>es x = Some e\<close> | 
| 18153 | 398 | show ?thesis | 
| 399 | proof (cases xs) | |
| 400 | case Nil | |
| 401 | with asm Env Some have False by simp | |
| 402 | then show ?thesis .. | |
| 10943 | 403 | next | 
| 18153 | 404 | case (Cons x' xs') | 
| 405 | from asm Env es have "lookup e xs = None" by simp | |
| 406 | then have "update (xs @ y # ys) opt e = e" by (rule hyp) | |
| 407 | with Env es Cons show "update ((x # xs) @ y # ys) opt env = env" | |
| 408 | by (simp add: fun_upd_idem_iff) | |
| 10943 | 409 | qed | 
| 410 | qed | |
| 18153 | 411 | qed | 
| 10943 | 412 | qed | 
| 413 | ||
| 414 | theorem update_append_some: | |
| 18153 | 415 | assumes "lookup env xs = Some e" | 
| 416 | shows "lookup (update (xs @ y # ys) opt env) xs = Some (update (y # ys) opt e)" | |
| 23394 | 417 | using assms | 
| 20503 | 418 | proof (induct xs arbitrary: env e) | 
| 18153 | 419 | case Nil | 
| 420 | then have "env = e" by simp | |
| 421 | then show ?case by simp | |
| 422 | next | |
| 423 | case (Cons x xs) | |
| 424 | note hyp = Cons.hyps | |
| 58613 | 425 | and asm = \<open>lookup env (x # xs) = Some e\<close> | 
| 18153 | 426 | show "lookup (update ((x # xs) @ y # ys) opt env) (x # xs) = | 
| 427 | Some (update (y # ys) opt e)" | |
| 428 | proof (cases env) | |
| 429 | case (Val a) | |
| 430 | with asm have False by simp | |
| 431 | then show ?thesis .. | |
| 10943 | 432 | next | 
| 18153 | 433 | case (Env b es) | 
| 434 | show ?thesis | |
| 435 | proof (cases "es x") | |
| 436 | case None | |
| 437 | with asm Env have False by simp | |
| 438 | then show ?thesis .. | |
| 10943 | 439 | next | 
| 18153 | 440 | case (Some e') | 
| 58613 | 441 | note es = \<open>es x = Some e'\<close> | 
| 10943 | 442 | show ?thesis | 
| 18153 | 443 | proof (cases xs) | 
| 444 | case Nil | |
| 445 | with asm Env es have "e = e'" by simp | |
| 446 | with Env es Nil show ?thesis by simp | |
| 10943 | 447 | next | 
| 18153 | 448 | case (Cons x' xs') | 
| 449 | from asm Env es have "lookup e' xs = Some e" by simp | |
| 450 | then have "lookup (update (xs @ y # ys) opt e') xs = | |
| 451 | Some (update (y # ys) opt e)" by (rule hyp) | |
| 452 | with Env es Cons show ?thesis by simp | |
| 10943 | 453 | qed | 
| 454 | qed | |
| 18153 | 455 | qed | 
| 10943 | 456 | qed | 
| 457 | ||
| 58613 | 458 | text \<open> | 
| 61893 | 459 | \<^medskip> | 
| 69597 | 460 | Apparently, \<^term>\<open>update\<close> does not affect the result of subsequent \<^term>\<open>lookup\<close> operations at independent positions, i.e.\ in case that the paths | 
| 461 | for \<^term>\<open>update\<close> and \<^term>\<open>lookup\<close> fork at a certain point. | |
| 58613 | 462 | \<close> | 
| 10943 | 463 | |
| 464 | theorem lookup_update_other: | |
| 18153 | 465 | assumes neq: "y \<noteq> (z::'c)" | 
| 466 | shows "lookup (update (xs @ z # zs) opt env) (xs @ y # ys) = | |
| 10943 | 467 | lookup env (xs @ y # ys)" | 
| 20503 | 468 | proof (induct xs arbitrary: env) | 
| 18153 | 469 | case Nil | 
| 470 | show ?case | |
| 471 | proof (cases env) | |
| 472 | case Val | |
| 473 | then show ?thesis by simp | |
| 474 | next | |
| 475 | case Env | |
| 476 | show ?thesis | |
| 477 | proof (cases zs) | |
| 478 | case Nil | |
| 479 | with neq Env show ?thesis by simp | |
| 10943 | 480 | next | 
| 18153 | 481 | case Cons | 
| 482 | with neq Env show ?thesis by simp | |
| 483 | qed | |
| 484 | qed | |
| 485 | next | |
| 486 | case (Cons x xs) | |
| 487 | note hyp = Cons.hyps | |
| 488 | show ?case | |
| 489 | proof (cases env) | |
| 490 | case Val | |
| 491 | then show ?thesis by simp | |
| 492 | next | |
| 493 | case (Env y es) | |
| 494 | show ?thesis | |
| 495 | proof (cases xs) | |
| 496 | case Nil | |
| 10943 | 497 | show ?thesis | 
| 18153 | 498 | proof (cases "es x") | 
| 499 | case None | |
| 500 | with Env Nil show ?thesis by simp | |
| 10943 | 501 | next | 
| 18153 | 502 | case Some | 
| 503 | with neq hyp and Env Nil show ?thesis by simp | |
| 504 | qed | |
| 505 | next | |
| 506 | case (Cons x' xs') | |
| 507 | show ?thesis | |
| 508 | proof (cases "es x") | |
| 509 | case None | |
| 510 | with Env Cons show ?thesis by simp | |
| 511 | next | |
| 512 | case Some | |
| 513 | with neq hyp and Env Cons show ?thesis by simp | |
| 10943 | 514 | qed | 
| 515 | qed | |
| 18153 | 516 | qed | 
| 10943 | 517 | qed | 
| 518 | ||
| 44267 | 519 | |
| 58613 | 520 | subsection \<open>Code generation\<close> | 
| 24433 
4a405457e9d6
added explicit equation for equality of nested environments
 haftmann parents: 
23394diff
changeset | 521 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
36176diff
changeset | 522 | lemma equal_env_code [code]: | 
| 49679 | 523 | fixes x y :: "'a::equal" | 
| 524 |     and f g :: "'c::{equal, finite} \<Rightarrow> ('b::equal, 'a, 'c) env option"
 | |
| 44267 | 525 | shows | 
| 526 | "HOL.equal (Env x f) (Env y g) \<longleftrightarrow> | |
| 527 | HOL.equal x y \<and> (\<forall>z \<in> UNIV. | |
| 528 | case f z of | |
| 529 | None \<Rightarrow> (case g z of None \<Rightarrow> True | Some _ \<Rightarrow> False) | |
| 530 | | Some a \<Rightarrow> (case g z of None \<Rightarrow> False | Some b \<Rightarrow> HOL.equal a b))" (is ?env) | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
36176diff
changeset | 531 | and "HOL.equal (Val a) (Val b) \<longleftrightarrow> HOL.equal a b" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
36176diff
changeset | 532 | and "HOL.equal (Val a) (Env y g) \<longleftrightarrow> False" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
36176diff
changeset | 533 | and "HOL.equal (Env x f) (Val b) \<longleftrightarrow> False" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
36176diff
changeset | 534 | proof (unfold equal) | 
| 44267 | 535 | have "f = g \<longleftrightarrow> | 
| 536 | (\<forall>z. case f z of | |
| 537 | None \<Rightarrow> (case g z of None \<Rightarrow> True | Some _ \<Rightarrow> False) | |
| 538 | | Some a \<Rightarrow> (case g z of None \<Rightarrow> False | Some b \<Rightarrow> a = b))" (is "?lhs = ?rhs") | |
| 24433 
4a405457e9d6
added explicit equation for equality of nested environments
 haftmann parents: 
23394diff
changeset | 539 | proof | 
| 
4a405457e9d6
added explicit equation for equality of nested environments
 haftmann parents: 
23394diff
changeset | 540 | assume ?lhs | 
| 
4a405457e9d6
added explicit equation for equality of nested environments
 haftmann parents: 
23394diff
changeset | 541 | then show ?rhs by (auto split: option.splits) | 
| 
4a405457e9d6
added explicit equation for equality of nested environments
 haftmann parents: 
23394diff
changeset | 542 | next | 
| 44267 | 543 | assume ?rhs (is "\<forall>z. ?prop z") | 
| 44236 
b73b7832b384
moved theory Nested_Environment to HOL-Unix (a bit too specific for HOL-Library);
 wenzelm parents: 
38857diff
changeset | 544 | show ?lhs | 
| 24433 
4a405457e9d6
added explicit equation for equality of nested environments
 haftmann parents: 
23394diff
changeset | 545 | proof | 
| 
4a405457e9d6
added explicit equation for equality of nested environments
 haftmann parents: 
23394diff
changeset | 546 | fix z | 
| 58613 | 547 | from \<open>?rhs\<close> have "?prop z" .. | 
| 24433 
4a405457e9d6
added explicit equation for equality of nested environments
 haftmann parents: 
23394diff
changeset | 548 | then show "f z = g z" by (auto split: option.splits) | 
| 
4a405457e9d6
added explicit equation for equality of nested environments
 haftmann parents: 
23394diff
changeset | 549 | qed | 
| 
4a405457e9d6
added explicit equation for equality of nested environments
 haftmann parents: 
23394diff
changeset | 550 | qed | 
| 26513 | 551 | then show "Env x f = Env y g \<longleftrightarrow> | 
| 49679 | 552 | x = y \<and> (\<forall>z \<in> UNIV. | 
| 44267 | 553 | case f z of | 
| 554 | None \<Rightarrow> (case g z of None \<Rightarrow> True | Some _ \<Rightarrow> False) | |
| 555 | | Some a \<Rightarrow> (case g z of None \<Rightarrow> False | Some b \<Rightarrow> a = b))" by simp | |
| 24433 
4a405457e9d6
added explicit equation for equality of nested environments
 haftmann parents: 
23394diff
changeset | 556 | qed simp_all | 
| 
4a405457e9d6
added explicit equation for equality of nested environments
 haftmann parents: 
23394diff
changeset | 557 | |
| 49679 | 558 | lemma [code nbe]: "HOL.equal (x :: (_, _, _) env) x \<longleftrightarrow> True" | 
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
36176diff
changeset | 559 | by (fact equal_refl) | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
36176diff
changeset | 560 | |
| 10943 | 561 | end |