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