src/HOL/Imperative_HOL/Ref.thy
author haftmann
Wed, 13 Feb 2013 13:38:52 +0100
changeset 52227 bee2678a3b61
parent 49088 1b609a7837ef
child 53572 6646bb548c6b
permissions -rw-r--r--
tuned spelling
haftmann@37787
     1
(*  Title:      HOL/Imperative_HOL/Ref.thy
haftmann@26170
     2
    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
haftmann@26170
     3
*)
haftmann@26170
     4
haftmann@26170
     5
header {* Monadic references *}
haftmann@26170
     6
haftmann@26170
     7
theory Ref
haftmann@37716
     8
imports Array
haftmann@26170
     9
begin
haftmann@26170
    10
haftmann@26170
    11
text {*
haftmann@26170
    12
  Imperative reference operations; modeled after their ML counterparts.
haftmann@26170
    13
  See http://caml.inria.fr/pub/docs/manual-caml-light/node14.15.html
haftmann@26170
    14
  and http://www.smlnj.org/doc/Conversion/top-level-comparison.html
haftmann@26170
    15
*}
haftmann@26170
    16
haftmann@37752
    17
subsection {* Primitives *}
haftmann@37716
    18
haftmann@37722
    19
definition present :: "heap \<Rightarrow> 'a\<Colon>heap ref \<Rightarrow> bool" where
haftmann@37722
    20
  "present h r \<longleftrightarrow> addr_of_ref r < lim h"
haftmann@37716
    21
haftmann@37722
    22
definition get :: "heap \<Rightarrow> 'a\<Colon>heap ref \<Rightarrow> 'a" where
haftmann@37722
    23
  "get h = from_nat \<circ> refs h TYPEREP('a) \<circ> addr_of_ref"
haftmann@37716
    24
haftmann@37722
    25
definition set :: "'a\<Colon>heap ref \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
haftmann@37722
    26
  "set r x = refs_update
haftmann@37722
    27
    (\<lambda>h. h(TYPEREP('a) := ((h (TYPEREP('a))) (addr_of_ref r := to_nat x))))"
haftmann@37716
    28
haftmann@37722
    29
definition alloc :: "'a \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap ref \<times> heap" where
haftmann@37722
    30
  "alloc x h = (let
haftmann@37716
    31
     l = lim h;
haftmann@37722
    32
     r = Ref l
haftmann@37722
    33
   in (r, set r x (h\<lparr>lim := l + 1\<rparr>)))"
haftmann@37716
    34
haftmann@37722
    35
definition noteq :: "'a\<Colon>heap ref \<Rightarrow> 'b\<Colon>heap ref \<Rightarrow> bool" (infix "=!=" 70) where
haftmann@37716
    36
  "r =!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_ref r \<noteq> addr_of_ref s"
haftmann@37716
    37
haftmann@37752
    38
haftmann@37752
    39
subsection {* Monad operations *}
haftmann@37752
    40
haftmann@37752
    41
definition ref :: "'a\<Colon>heap \<Rightarrow> 'a ref Heap" where
haftmann@37752
    42
  [code del]: "ref v = Heap_Monad.heap (alloc v)"
haftmann@37752
    43
haftmann@37752
    44
definition lookup :: "'a\<Colon>heap ref \<Rightarrow> 'a Heap" ("!_" 61) where
haftmann@37757
    45
  [code del]: "lookup r = Heap_Monad.tap (\<lambda>h. get h r)"
haftmann@37752
    46
haftmann@37752
    47
definition update :: "'a ref \<Rightarrow> 'a\<Colon>heap \<Rightarrow> unit Heap" ("_ := _" 62) where
haftmann@37752
    48
  [code del]: "update r v = Heap_Monad.heap (\<lambda>h. ((), set r v h))"
haftmann@37752
    49
haftmann@37752
    50
definition change :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a ref \<Rightarrow> 'a Heap" where
krauss@37792
    51
  "change f r = do {
haftmann@37752
    52
     x \<leftarrow> ! r;
haftmann@37752
    53
     let y = f x;
haftmann@37752
    54
     r := y;
haftmann@37752
    55
     return y
krauss@37792
    56
   }"
haftmann@37752
    57
haftmann@37752
    58
haftmann@37752
    59
subsection {* Properties *}
haftmann@37752
    60
haftmann@37757
    61
text {* Primitives *}
haftmann@37757
    62
haftmann@37722
    63
lemma noteq_sym: "r =!= s \<Longrightarrow> s =!= r"
haftmann@37722
    64
  and unequal [simp]: "r \<noteq> r' \<longleftrightarrow> r =!= r'" -- "same types!"
haftmann@37722
    65
  by (auto simp add: noteq_def)
haftmann@37716
    66
haftmann@37722
    67
lemma noteq_irrefl: "r =!= r \<Longrightarrow> False"
haftmann@37722
    68
  by (auto simp add: noteq_def)
haftmann@37716
    69
haftmann@37722
    70
lemma present_alloc_neq: "present h r \<Longrightarrow> r =!= fst (alloc v h)"
haftmann@37722
    71
  by (simp add: present_def alloc_def noteq_def Let_def)
haftmann@37716
    72
haftmann@37722
    73
lemma next_fresh [simp]:
haftmann@37722
    74
  assumes "(r, h') = alloc x h"
haftmann@37722
    75
  shows "\<not> present h r"
haftmann@37722
    76
  using assms by (cases h) (auto simp add: alloc_def present_def Let_def)
haftmann@37716
    77
haftmann@37722
    78
lemma next_present [simp]:
haftmann@37722
    79
  assumes "(r, h') = alloc x h"
haftmann@37722
    80
  shows "present h' r"
haftmann@37722
    81
  using assms by (cases h) (auto simp add: alloc_def set_def present_def Let_def)
haftmann@37716
    82
haftmann@37722
    83
lemma get_set_eq [simp]:
haftmann@37722
    84
  "get (set r x h) r = x"
haftmann@37722
    85
  by (simp add: get_def set_def)
haftmann@37716
    86
haftmann@37722
    87
lemma get_set_neq [simp]:
haftmann@37722
    88
  "r =!= s \<Longrightarrow> get (set s x h) r = get h r"
haftmann@37722
    89
  by (simp add: noteq_def get_def set_def)
haftmann@37716
    90
haftmann@37722
    91
lemma set_same [simp]:
haftmann@37722
    92
  "set r x (set r y h) = set r x h"
haftmann@37722
    93
  by (simp add: set_def)
haftmann@37716
    94
haftmann@37771
    95
lemma not_present_alloc [simp]:
haftmann@37771
    96
  "\<not> present h (fst (alloc v h))"
haftmann@37771
    97
  by (simp add: present_def alloc_def Let_def)
haftmann@37771
    98
haftmann@37722
    99
lemma set_set_swap:
haftmann@37722
   100
  "r =!= r' \<Longrightarrow> set r x (set r' x' h) = set r' x' (set r x h)"
nipkow@39535
   101
  by (simp add: noteq_def set_def fun_eq_iff)
haftmann@37716
   102
haftmann@37722
   103
lemma alloc_set:
haftmann@37722
   104
  "fst (alloc x (set r x' h)) = fst (alloc x h)"
haftmann@37722
   105
  by (simp add: alloc_def set_def Let_def)
haftmann@37716
   106
haftmann@37722
   107
lemma get_alloc [simp]:
haftmann@37722
   108
  "get (snd (alloc x h)) (fst (alloc x' h)) = x"
haftmann@37722
   109
  by (simp add: alloc_def Let_def)
haftmann@37716
   110
haftmann@37722
   111
lemma set_alloc [simp]:
haftmann@37722
   112
  "set (fst (alloc v h)) v' (snd (alloc v h)) = snd (alloc v' h)"
haftmann@37722
   113
  by (simp add: alloc_def Let_def)
haftmann@37716
   114
haftmann@37722
   115
lemma get_alloc_neq: "r =!= fst (alloc v h) \<Longrightarrow> 
haftmann@37722
   116
  get (snd (alloc v h)) r  = get h r"
haftmann@37722
   117
  by (simp add: get_def set_def alloc_def Let_def noteq_def)
haftmann@37716
   118
haftmann@37722
   119
lemma lim_set [simp]:
haftmann@37722
   120
  "lim (set r v h) = lim h"
haftmann@37722
   121
  by (simp add: set_def)
haftmann@37716
   122
haftmann@37722
   123
lemma present_alloc [simp]: 
haftmann@37722
   124
  "present h r \<Longrightarrow> present (snd (alloc v h)) r"
haftmann@37722
   125
  by (simp add: present_def alloc_def Let_def)
haftmann@37716
   126
haftmann@37722
   127
lemma present_set [simp]:
haftmann@37722
   128
  "present (set r v h) = present h"
nipkow@39535
   129
  by (simp add: present_def fun_eq_iff)
haftmann@37716
   130
haftmann@37722
   131
lemma noteq_I:
haftmann@37722
   132
  "present h r \<Longrightarrow> \<not> present h r' \<Longrightarrow> r =!= r'"
haftmann@37722
   133
  by (auto simp add: noteq_def present_def)
haftmann@37716
   134
haftmann@37757
   135
haftmann@37757
   136
text {* Monad operations *}
haftmann@37757
   137
haftmann@37787
   138
lemma execute_ref [execute_simps]:
haftmann@37757
   139
  "execute (ref v) h = Some (alloc v h)"
haftmann@37787
   140
  by (simp add: ref_def execute_simps)
haftmann@37716
   141
haftmann@37787
   142
lemma success_refI [success_intros]:
haftmann@37757
   143
  "success (ref v) h"
haftmann@37787
   144
  by (auto intro: success_intros simp add: ref_def)
haftmann@37757
   145
haftmann@40919
   146
lemma effect_refI [effect_intros]:
haftmann@37771
   147
  assumes "(r, h') = alloc v h"
haftmann@40919
   148
  shows "effect (ref v) h h' r"
haftmann@40919
   149
  by (rule effectI) (insert assms, simp add: execute_simps)
haftmann@37771
   150
haftmann@40919
   151
lemma effect_refE [effect_elims]:
haftmann@40919
   152
  assumes "effect (ref v) h h' r"
haftmann@37796
   153
  obtains "get h' r = v" and "present h' r" and "\<not> present h r"
haftmann@40919
   154
  using assms by (rule effectE) (simp add: execute_simps)
haftmann@37771
   155
haftmann@37787
   156
lemma execute_lookup [execute_simps]:
haftmann@37752
   157
  "Heap_Monad.execute (lookup r) h = Some (get h r, h)"
haftmann@37787
   158
  by (simp add: lookup_def execute_simps)
haftmann@26170
   159
haftmann@37787
   160
lemma success_lookupI [success_intros]:
haftmann@37757
   161
  "success (lookup r) h"
haftmann@37787
   162
  by (auto intro: success_intros  simp add: lookup_def)
haftmann@37757
   163
haftmann@40919
   164
lemma effect_lookupI [effect_intros]:
haftmann@37796
   165
  assumes "h' = h" "x = get h r"
haftmann@40919
   166
  shows "effect (!r) h h' x"
haftmann@40919
   167
  by (rule effectI) (insert assms, simp add: execute_simps)
haftmann@37771
   168
haftmann@40919
   169
lemma effect_lookupE [effect_elims]:
haftmann@40919
   170
  assumes "effect (!r) h h' x"
haftmann@37796
   171
  obtains "h' = h" "x = get h r"
haftmann@40919
   172
  using assms by (rule effectE) (simp add: execute_simps)
haftmann@37771
   173
haftmann@37787
   174
lemma execute_update [execute_simps]:
haftmann@37752
   175
  "Heap_Monad.execute (update r v) h = Some ((), set r v h)"
haftmann@37787
   176
  by (simp add: update_def execute_simps)
haftmann@26170
   177
haftmann@37787
   178
lemma success_updateI [success_intros]:
haftmann@37757
   179
  "success (update r v) h"
haftmann@37787
   180
  by (auto intro: success_intros  simp add: update_def)
haftmann@37757
   181
haftmann@40919
   182
lemma effect_updateI [effect_intros]:
haftmann@37796
   183
  assumes "h' = set r v h"
haftmann@40919
   184
  shows "effect (r := v) h h' x"
haftmann@40919
   185
  by (rule effectI) (insert assms, simp add: execute_simps)
haftmann@37771
   186
haftmann@40919
   187
lemma effect_updateE [effect_elims]:
haftmann@40919
   188
  assumes "effect (r' := v) h h' r"
haftmann@37796
   189
  obtains "h' = set r' v h"
haftmann@40919
   190
  using assms by (rule effectE) (simp add: execute_simps)
haftmann@37771
   191
haftmann@37787
   192
lemma execute_change [execute_simps]:
haftmann@37752
   193
  "Heap_Monad.execute (change f r) h = Some (f (get h r), set r (f (get h r)) h)"
haftmann@37787
   194
  by (simp add: change_def bind_def Let_def execute_simps)
haftmann@37757
   195
haftmann@37787
   196
lemma success_changeI [success_intros]:
haftmann@37757
   197
  "success (change f r) h"
haftmann@40919
   198
  by (auto intro!: success_intros effect_intros simp add: change_def)
haftmann@37771
   199
haftmann@40919
   200
lemma effect_changeI [effect_intros]: 
haftmann@37796
   201
  assumes "h' = set r (f (get h r)) h" "x = f (get h r)"
haftmann@40919
   202
  shows "effect (change f r) h h' x"
haftmann@40919
   203
  by (rule effectI) (insert assms, simp add: execute_simps)  
haftmann@37771
   204
haftmann@40919
   205
lemma effect_changeE [effect_elims]:
haftmann@40919
   206
  assumes "effect (change f r') h h' r"
haftmann@37796
   207
  obtains "h' = set r' (f (get h r')) h" "r = f (get h r')"
haftmann@40919
   208
  using assms by (rule effectE) (simp add: execute_simps)
haftmann@26170
   209
haftmann@26170
   210
lemma lookup_chain:
haftmann@26170
   211
  "(!r \<guillemotright> f) = f"
haftmann@37771
   212
  by (rule Heap_eqI) (auto simp add: lookup_def execute_simps intro: execute_bind)
haftmann@26170
   213
haftmann@28562
   214
lemma update_change [code]:
haftmann@37722
   215
  "r := e = change (\<lambda>_. e) r \<guillemotright> return ()"
haftmann@37722
   216
  by (rule Heap_eqI) (simp add: change_def lookup_chain)
haftmann@26170
   217
haftmann@26182
   218
haftmann@52227
   219
text {* Non-interaction between imperative arrays and imperative references *}
haftmann@37716
   220
haftmann@37806
   221
lemma array_get_set [simp]:
haftmann@37806
   222
  "Array.get (set r v h) = Array.get h"
nipkow@39535
   223
  by (simp add: Array.get_def set_def fun_eq_iff)
haftmann@37716
   224
haftmann@37796
   225
lemma get_update [simp]:
haftmann@37805
   226
  "get (Array.update a i v h) r = get h r"
haftmann@37804
   227
  by (simp add: get_def Array.update_def Array.set_def)
haftmann@37716
   228
haftmann@37796
   229
lemma alloc_update:
haftmann@37796
   230
  "fst (alloc v (Array.update a i v' h)) = fst (alloc v h)"
haftmann@37806
   231
  by (simp add: Array.update_def Array.get_def Array.set_def alloc_def Let_def)
haftmann@37716
   232
haftmann@37796
   233
lemma update_set_swap:
haftmann@37796
   234
  "Array.update a i v (set r v' h) = set r v' (Array.update a i v h)"
haftmann@37806
   235
  by (simp add: Array.update_def Array.get_def Array.set_def set_def)
haftmann@37716
   236
haftmann@37722
   237
lemma length_alloc [simp]: 
haftmann@37802
   238
  "Array.length (snd (alloc v h)) a = Array.length h a"
haftmann@37806
   239
  by (simp add: Array.length_def Array.get_def alloc_def set_def Let_def)
haftmann@37716
   240
haftmann@37806
   241
lemma array_get_alloc [simp]: 
haftmann@37806
   242
  "Array.get (snd (alloc v h)) = Array.get h"
nipkow@39535
   243
  by (simp add: Array.get_def alloc_def set_def Let_def fun_eq_iff)
haftmann@37716
   244
haftmann@37796
   245
lemma present_update [simp]: 
haftmann@37796
   246
  "present (Array.update a i v h) = present h"
nipkow@39535
   247
  by (simp add: Array.update_def Array.set_def fun_eq_iff present_def)
haftmann@37716
   248
haftmann@37722
   249
lemma array_present_set [simp]:
haftmann@37804
   250
  "Array.present (set r v h) = Array.present h"
nipkow@39535
   251
  by (simp add: Array.present_def set_def fun_eq_iff)
haftmann@37716
   252
haftmann@37722
   253
lemma array_present_alloc [simp]:
haftmann@37804
   254
  "Array.present h a \<Longrightarrow> Array.present (snd (alloc v h)) a"
haftmann@37804
   255
  by (simp add: Array.present_def alloc_def Let_def)
haftmann@37716
   256
haftmann@37722
   257
lemma set_array_set_swap:
haftmann@37804
   258
  "Array.set a xs (set r x' h) = set r x' (Array.set a xs h)"
haftmann@37804
   259
  by (simp add: Array.set_def set_def)
haftmann@37722
   260
haftmann@37796
   261
hide_const (open) present get set alloc noteq lookup update change
haftmann@37716
   262
haftmann@37716
   263
haftmann@26182
   264
subsection {* Code generator setup *}
haftmann@26182
   265
haftmann@52227
   266
text {* Intermediate operation avoids invariance problem in @{text Scala} (similar to value restriction) *}
haftmann@38314
   267
haftmann@38314
   268
definition ref' where
haftmann@38314
   269
  [code del]: "ref' = ref"
haftmann@38314
   270
haftmann@38314
   271
lemma [code]:
haftmann@38314
   272
  "ref x = ref' x"
haftmann@38314
   273
  by (simp add: ref'_def)
haftmann@38314
   274
haftmann@38314
   275
haftmann@39831
   276
text {* SML / Eval *}
haftmann@26182
   277
haftmann@39831
   278
code_type ref (SML "_/ ref")
haftmann@39831
   279
code_type ref (Eval "_/ Unsynchronized.ref")
haftmann@26182
   280
code_const Ref (SML "raise/ (Fail/ \"bare Ref\")")
haftmann@39831
   281
code_const ref' (SML "(fn/ ()/ =>/ ref/ _)")
haftmann@39831
   282
code_const ref' (Eval "(fn/ ()/ =>/ Unsynchronized.ref/ _)")
haftmann@26752
   283
code_const Ref.lookup (SML "(fn/ ()/ =>/ !/ _)")
haftmann@26752
   284
code_const Ref.update (SML "(fn/ ()/ =>/ _/ :=/ _)")
haftmann@39948
   285
code_const "HOL.equal :: 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool" (SML infixl 6 "=")
haftmann@26182
   286
haftmann@39831
   287
code_reserved Eval Unsynchronized
haftmann@26182
   288
haftmann@26182
   289
haftmann@37752
   290
text {* OCaml *}
haftmann@26182
   291
haftmann@26182
   292
code_type ref (OCaml "_/ ref")
haftmann@37830
   293
code_const Ref (OCaml "failwith/ \"bare Ref\"")
haftmann@38314
   294
code_const ref' (OCaml "(fun/ ()/ ->/ ref/ _)")
haftmann@37830
   295
code_const Ref.lookup (OCaml "(fun/ ()/ ->/ !/ _)")
haftmann@37830
   296
code_const Ref.update (OCaml "(fun/ ()/ ->/ _/ :=/ _)")
haftmann@39948
   297
code_const "HOL.equal :: 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool" (OCaml infixl 4 "=")
haftmann@26182
   298
haftmann@26182
   299
code_reserved OCaml ref
haftmann@26182
   300
haftmann@26182
   301
haftmann@37752
   302
text {* Haskell *}
haftmann@26182
   303
haftmann@29730
   304
code_type ref (Haskell "Heap.STRef/ Heap.RealWorld/ _")
haftmann@26182
   305
code_const Ref (Haskell "error/ \"bare Ref\"")
haftmann@38314
   306
code_const ref' (Haskell "Heap.newSTRef")
haftmann@29730
   307
code_const Ref.lookup (Haskell "Heap.readSTRef")
haftmann@29730
   308
code_const Ref.update (Haskell "Heap.writeSTRef")
haftmann@39948
   309
code_const "HOL.equal :: 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool" (Haskell infix 4 "==")
haftmann@39948
   310
code_instance ref :: HOL.equal (Haskell -)
haftmann@26182
   311
haftmann@37842
   312
haftmann@37842
   313
text {* Scala *}
haftmann@37842
   314
haftmann@39194
   315
code_type ref (Scala "!Ref[_]")
haftmann@49088
   316
code_const Ref (Scala "!sys.error(\"bare Ref\")")
haftmann@39194
   317
code_const ref' (Scala "('_: Unit)/ =>/ Ref((_))")
haftmann@39194
   318
code_const Ref.lookup (Scala "('_: Unit)/ =>/ Ref.lookup((_))")
haftmann@39194
   319
code_const Ref.update (Scala "('_: Unit)/ =>/ Ref.update((_), (_))")
haftmann@39948
   320
code_const "HOL.equal :: 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool" (Scala infixl 5 "==")
haftmann@37842
   321
haftmann@37757
   322
end