src/HOL/Library/Array.thy
author haftmann
Wed, 27 Feb 2008 21:41:08 +0100
changeset 26170 66e6b967ccf1
child 26182 8262ec0e8782
permissions -rw-r--r--
added theories for imperative HOL
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/Array.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 arrays *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     7
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     8
theory Array
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
subsection {* Primitives *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    13
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    14
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    15
  new :: "nat \<Rightarrow> 'a\<Colon>heap \<Rightarrow> 'a array Heap" where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    16
  [code del]: "new n x = Heap_Monad.heap (Heap.array n x)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    17
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    18
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    19
  of_list :: "'a\<Colon>heap list \<Rightarrow> 'a array Heap" where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    20
  [code del]: "of_list xs = Heap_Monad.heap (Heap.array_of_list xs)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    21
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    22
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    23
  length :: "'a\<Colon>heap array \<Rightarrow> nat Heap" where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    24
  [code del]: "length arr = Heap_Monad.heap (\<lambda>h. (Heap.length arr h, h))"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    25
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    26
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    27
  nth :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a Heap"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    28
where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    29
  [code del]: "nth a i = (do len \<leftarrow> length a;
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    30
                 (if i < len
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    31
                     then Heap_Monad.heap (\<lambda>h. (get_array a h ! i, h))
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    32
                     else raise (''array lookup: index out of range''))
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    33
              done)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    34
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    35
-- {* FIXME adjustion for List theory *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    36
no_syntax
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    37
  nth  :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!" 100)
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    38
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    39
abbreviation
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    40
  nth_list :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!" 100)
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    41
where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    42
  "nth_list \<equiv> List.nth"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    43
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    44
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    45
  upd :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a\<Colon>heap array Heap"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    46
where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    47
  [code del]: "upd i x a = (do len \<leftarrow> length a;
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    48
                      (if i < len
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    49
                           then Heap_Monad.heap (\<lambda>h. ((), Heap.upd a i x h))
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    50
                           else raise (''array update: index out of range''));
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    51
                      return a
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    52
                   done)" 
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    53
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    54
lemma upd_return:
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    55
  "upd i x a \<guillemotright> return a = upd i x a"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    56
  unfolding upd_def by (simp add: monad_simp)
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    57
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    58
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    59
subsection {* Derivates *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    60
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    61
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    62
  map_entry :: "nat \<Rightarrow> ('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    63
where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    64
  "map_entry i f a = (do
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    65
     x \<leftarrow> nth a i;
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    66
     upd i (f x) a
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    67
   done)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    68
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    69
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    70
  swap :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a Heap"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    71
where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    72
  "swap i x a = (do
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    73
     y \<leftarrow> nth a i;
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    74
     upd i x a;
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    75
     return x
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    76
   done)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    77
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    78
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    79
  make :: "nat \<Rightarrow> (nat \<Rightarrow> 'a\<Colon>heap) \<Rightarrow> 'a array Heap"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    80
where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    81
  "make n f = of_list (map f [0 ..< n])"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    82
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    83
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    84
  freeze :: "'a\<Colon>heap array \<Rightarrow> 'a list Heap"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    85
where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    86
  "freeze a = (do
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    87
     n \<leftarrow> length a;
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    88
     mapM (nth a) [0..<n]
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    89
   done)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    90
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    91
definition
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    92
  map :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    93
where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    94
  "map f a = (do
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    95
     n \<leftarrow> length a;
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    96
     foldM (\<lambda>n. map_entry n f) [0..<n] a
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    97
   done)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    98
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    99
hide (open) const new map -- {* avoid clashed with some popular names *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   100
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   101
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   102
subsection {* Converting arrays to lists *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   103
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   104
primrec list_of_aux :: "nat \<Rightarrow> ('a\<Colon>heap) array \<Rightarrow> 'a list \<Rightarrow> 'a list Heap" where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   105
  "list_of_aux 0 a xs = return xs"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   106
  | "list_of_aux (Suc n) a xs = (do
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   107
        x \<leftarrow> Array.nth a n;
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   108
        list_of_aux n a (x#xs)
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   109
     done)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   110
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   111
definition list_of :: "('a\<Colon>heap) array \<Rightarrow> 'a list Heap" where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   112
  "list_of a = (do n \<leftarrow> Array.length a;
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   113
                   list_of_aux n a []
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   114
                done)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   115
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   116
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   117
subsection {* Properties *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   118
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   119
lemma array_make [code func]:
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   120
  "Array.new n x = make n (\<lambda>_. x)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   121
  by (induct n) (simp_all add: make_def new_def Heap_Monad.heap_def
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   122
    monad_simp array_of_list_replicate [symmetric]
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   123
    map_replicate_trivial replicate_append_same
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   124
    of_list_def)
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   125
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   126
lemma array_of_list_make [code func]:
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   127
  "of_list xs = make (List.length xs) (\<lambda>n. xs ! n)"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   128
  unfolding make_def map_nth ..
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   129
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
   130
end