src/HOL/Library/Ref.thy
author haftmann
Thu, 28 Feb 2008 16:50:52 +0100
changeset 26182 8262ec0e8782
parent 26170 66e6b967ccf1
child 26752 6b276119139b
permissions -rw-r--r--
added code generator setup
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/Ref.thy
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     2
    ID:         $Id$
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     3
    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     4
*)
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     5
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     6
header {* Monadic references *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     7
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     8
theory Ref
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     9
imports Heap_Monad
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    10
begin
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    11
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    12
text {*
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    13
  Imperative reference operations; modeled after their ML counterparts.
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    14
  See http://caml.inria.fr/pub/docs/manual-caml-light/node14.15.html
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    15
  and http://www.smlnj.org/doc/Conversion/top-level-comparison.html
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    16
*}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    17
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    18
subsection {* Primitives *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    19
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    20
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    21
  new :: "'a\<Colon>heap \<Rightarrow> 'a ref Heap" where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    22
  [code del]: "new v = Heap_Monad.heap (Heap.ref v)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    23
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    24
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    25
  lookup :: "'a\<Colon>heap ref \<Rightarrow> 'a Heap" ("!_" 61) where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    26
  [code del]: "lookup r = Heap_Monad.heap (\<lambda>h. (get_ref r h, h))"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    27
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    28
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    29
  update :: "'a ref \<Rightarrow> ('a\<Colon>heap) \<Rightarrow> unit Heap" ("_ := _" 62) where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    30
  [code del]: "update r e = Heap_Monad.heap (\<lambda>h. ((), set_ref r e h))"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    31
26182
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    32
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    33
subsection {* Derivates *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    34
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    35
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    36
  change :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a ref \<Rightarrow> 'a Heap"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    37
where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    38
  "change f r = (do x \<leftarrow> ! r;
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    39
                    let y = f x;
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    40
                    r := y;
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    41
                    return y
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    42
                 done)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    43
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    44
hide (open) const new lookup update change
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    45
26182
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    46
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    47
subsection {* Properties *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    48
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    49
lemma lookup_chain:
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    50
  "(!r \<guillemotright> f) = f"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    51
  by (cases f)
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    52
    (auto simp add: Let_def bindM_def lookup_def expand_fun_eq)
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    53
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    54
lemma update_change [code func]:
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    55
  "r := e = Ref.change (\<lambda>_. e) r \<guillemotright> return ()"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    56
  by (auto simp add: monad_simp change_def lookup_chain)
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    57
26182
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    58
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    59
subsection {* Code generator setup *}
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    60
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    61
subsubsection {* SML *}
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    62
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    63
code_type ref (SML "_/ ref")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    64
code_const Ref (SML "raise/ (Fail/ \"bare Ref\")")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    65
code_const Ref.new (SML "ref")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    66
code_const Ref.lookup (SML "'!")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    67
code_const Ref.update (SML infixl 3 ":=")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    68
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    69
code_reserved SML ref
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    70
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    71
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    72
subsubsection {* OCaml *}
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    73
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    74
code_type ref (OCaml "_/ ref")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    75
code_const Ref (OCaml "failwith/ \"bare Ref\"")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    76
code_const Ref.new (OCaml "ref")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    77
code_const Ref.lookup (OCaml "'!")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    78
code_const Ref.update (OCaml infixr 1 ":=")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    79
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    80
code_reserved OCaml ref
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    81
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    82
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    83
subsubsection {* Haskell *}
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    84
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    85
code_type ref (Haskell "STRef '_s _")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    86
code_const Ref (Haskell "error/ \"bare Ref\"")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    87
code_const Ref.new (Haskell "newSTRef")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    88
code_const Ref.lookup (Haskell "readSTRef")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    89
code_const Ref.update (Haskell "writeSTRef")
8262ec0e8782 added code generator setup
haftmann
parents: 26170
diff changeset
    90
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    91
end