src/HOL/Library/Mapping.thy
author huffman
Wed, 18 Feb 2009 19:32:26 -0800
changeset 29985 57975b45ab70
parent 29831 5dc920623bb1
child 30663 0b6aff7451b2
permissions -rw-r--r--
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/Mapping.thy
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
     3
*)
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
     4
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
     5
header {* An abstract view on maps for code generation. *}
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
     6
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
     7
theory Mapping
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
     8
imports Map
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
     9
begin
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    10
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    11
subsection {* Type definition and primitive operations *}
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    12
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    13
datatype ('a, 'b) map = Map "'a \<rightharpoonup> 'b"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    14
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    15
definition empty :: "('a, 'b) map" where
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    16
  "empty = Map (\<lambda>_. None)"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    17
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    18
primrec lookup :: "('a, 'b) map \<Rightarrow> 'a \<rightharpoonup> 'b" where
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    19
  "lookup (Map f) = f"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    20
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    21
primrec update :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) map \<Rightarrow> ('a, 'b) map" where
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    22
  "update k v (Map f) = Map (f (k \<mapsto> v))"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    23
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    24
primrec delete :: "'a \<Rightarrow> ('a, 'b) map \<Rightarrow> ('a, 'b) map" where
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    25
  "delete k (Map f) = Map (f (k := None))"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    26
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    27
primrec keys :: "('a, 'b) map \<Rightarrow> 'a set" where
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    28
  "keys (Map f) = dom f"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    29
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    30
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    31
subsection {* Derived operations *}
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    32
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    33
definition size :: "('a, 'b) map \<Rightarrow> nat" where
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    34
  "size m = (if finite (keys m) then card (keys m) else 0)"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    35
29814
15344c0899e1 added replace operation
haftmann
parents: 29708
diff changeset
    36
definition replace :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) map \<Rightarrow> ('a, 'b) map" where
15344c0899e1 added replace operation
haftmann
parents: 29708
diff changeset
    37
  "replace k v m = (if lookup m k = None then m else update k v m)"
15344c0899e1 added replace operation
haftmann
parents: 29708
diff changeset
    38
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    39
definition tabulate :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) map" where
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    40
  "tabulate ks f = Map (map_of (map (\<lambda>k. (k, f k)) ks))"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    41
29826
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
    42
definition bulkload :: "'a list \<Rightarrow> (nat, 'a) map" where
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
    43
  "bulkload xs = Map (\<lambda>k. if k < length xs then Some (xs ! k) else None)"
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
    44
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    45
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    46
subsection {* Properties *}
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    47
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    48
lemma lookup_inject:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    49
  "lookup m = lookup n \<longleftrightarrow> m = n"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    50
  by (cases m, cases n) simp
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    51
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    52
lemma lookup_empty [simp]:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    53
  "lookup empty = Map.empty"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    54
  by (simp add: empty_def)
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    55
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    56
lemma lookup_update [simp]:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    57
  "lookup (update k v m) = (lookup m) (k \<mapsto> v)"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    58
  by (cases m) simp
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    59
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    60
lemma lookup_delete:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    61
  "lookup (delete k m) k = None"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    62
  "k \<noteq> l \<Longrightarrow> lookup (delete k m) l = lookup m l"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    63
  by (cases m, simp)+
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    64
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    65
lemma lookup_tabulate:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    66
  "lookup (tabulate ks f) = (Some o f) |` set ks"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    67
  by (induct ks) (auto simp add: tabulate_def restrict_map_def expand_fun_eq)
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    68
29826
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
    69
lemma lookup_bulkload:
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
    70
  "lookup (bulkload xs) = (\<lambda>k. if k < length xs then Some (xs ! k) else None)"
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
    71
  unfolding bulkload_def by simp
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
    72
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    73
lemma update_update:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    74
  "update k v (update k w m) = update k v m"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    75
  "k \<noteq> l \<Longrightarrow> update k v (update l w m) = update l w (update k v m)"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    76
  by (cases m, simp add: expand_fun_eq)+
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    77
29814
15344c0899e1 added replace operation
haftmann
parents: 29708
diff changeset
    78
lemma replace_update:
15344c0899e1 added replace operation
haftmann
parents: 29708
diff changeset
    79
  "lookup m k = None \<Longrightarrow> replace k v m = m"
15344c0899e1 added replace operation
haftmann
parents: 29708
diff changeset
    80
  "lookup m k \<noteq> None \<Longrightarrow> replace k v m = update k v m"
15344c0899e1 added replace operation
haftmann
parents: 29708
diff changeset
    81
  by (auto simp add: replace_def)
15344c0899e1 added replace operation
haftmann
parents: 29708
diff changeset
    82
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    83
lemma delete_empty [simp]:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    84
  "delete k empty = empty"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    85
  by (simp add: empty_def)
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    86
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    87
lemma delete_update:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    88
  "delete k (update k v m) = delete k m"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    89
  "k \<noteq> l \<Longrightarrow> delete k (update l v m) = update l v (delete k m)"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    90
  by (cases m, simp add: expand_fun_eq)+
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    91
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    92
lemma update_delete [simp]:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    93
  "update k v (delete k m) = update k v m"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    94
  by (cases m) simp
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    95
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    96
lemma keys_empty [simp]:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    97
  "keys empty = {}"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    98
  unfolding empty_def by simp
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    99
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   100
lemma keys_update [simp]:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   101
  "keys (update k v m) = insert k (keys m)"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   102
  by (cases m) simp
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   103
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   104
lemma keys_delete [simp]:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   105
  "keys (delete k m) = keys m - {k}"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   106
  by (cases m) simp
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   107
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   108
lemma keys_tabulate [simp]:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   109
  "keys (tabulate ks f) = set ks"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   110
  by (auto simp add: tabulate_def dest: map_of_SomeD intro!: weak_map_of_SomeI)
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   111
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   112
lemma size_empty [simp]:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   113
  "size empty = 0"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   114
  by (simp add: size_def keys_empty)
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   115
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   116
lemma size_update:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   117
  "finite (keys m) \<Longrightarrow> size (update k v m) =
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   118
    (if k \<in> keys m then size m else Suc (size m))"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   119
  by (simp add: size_def keys_update)
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   120
    (auto simp only: card_insert card_Suc_Diff1)
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   121
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   122
lemma size_delete:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   123
  "size (delete k m) = (if k \<in> keys m then size m - 1 else size m)"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   124
  by (simp add: size_def keys_delete)
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   125
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   126
lemma size_tabulate:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   127
  "size (tabulate ks f) = length (remdups ks)"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   128
  by (simp add: size_def keys_tabulate distinct_card [of "remdups ks", symmetric])
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   129
29831
5dc920623bb1 Isar proof
haftmann
parents: 29828
diff changeset
   130
lemma bulkload_tabulate:
29826
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
   131
  "bulkload xs = tabulate [0..<length xs] (nth xs)"
29831
5dc920623bb1 Isar proof
haftmann
parents: 29828
diff changeset
   132
  by (rule sym)
5dc920623bb1 Isar proof
haftmann
parents: 29828
diff changeset
   133
    (auto simp add: bulkload_def tabulate_def expand_fun_eq map_of_eq_None_iff map_compose [symmetric] comp_def)
29826
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
   134
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   135
end