src/HOL/Library/Mapping.thy
author haftmann
Fri, 15 Feb 2013 11:47:34 +0100
changeset 51161 6ed12ae3b3e1
parent 49975 faf4afed009f
child 51375 d9e62d9c98de
permissions -rw-r--r--
attempt to re-establish conventions which theories are loaded into the grand unified library theory; four different code generation tests for different code setup constellations; augment code generation setup where necessary
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
     1
(*  Title:      HOL/Library/Mapping.thy
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
     2
    Author:     Florian Haftmann and Ondrej Kuncar
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
     3
*)
29708
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
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49973
diff changeset
     8
imports Main Quotient_Option
29708
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
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 45231
diff changeset
    13
typedef ('a, 'b) mapping = "UNIV :: ('a \<rightharpoonup> 'b) set"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    14
  morphisms rep Mapping ..
37700
bd90378b8171 refrain from using datatype declaration -- opens chance for quickcheck later on
haftmann
parents: 37299
diff changeset
    15
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    16
setup_lifting(no_code) type_definition_mapping
37700
bd90378b8171 refrain from using datatype declaration -- opens chance for quickcheck later on
haftmann
parents: 37299
diff changeset
    17
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    18
lift_definition empty :: "('a, 'b) mapping" is "(\<lambda>_. None)" .
37700
bd90378b8171 refrain from using datatype declaration -- opens chance for quickcheck later on
haftmann
parents: 37299
diff changeset
    19
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    20
lift_definition lookup :: "('a, 'b) mapping \<Rightarrow> 'a \<Rightarrow> 'b option" is "\<lambda>m k. m k" .
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    21
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    22
lift_definition update :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" is "\<lambda>k v m. m(k \<mapsto> v)" .
37700
bd90378b8171 refrain from using datatype declaration -- opens chance for quickcheck later on
haftmann
parents: 37299
diff changeset
    23
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    24
lift_definition delete :: "'a \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" is "\<lambda>k m. m(k := None)" .
39380
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 39302
diff changeset
    25
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    26
lift_definition keys :: "('a, 'b) mapping \<Rightarrow> 'a set" is dom .
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    27
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    28
lift_definition tabulate :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping" is
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    29
  "\<lambda>ks f. (map_of (List.map (\<lambda>k. (k, f k)) ks))" .
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    30
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    31
lift_definition bulkload :: "'a list \<Rightarrow> (nat, 'a) mapping" is
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    32
  "\<lambda>xs k. if k < length xs then Some (xs ! k) else None" .
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    33
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    34
lift_definition map :: "('c \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('c, 'd) mapping" is
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    35
  "\<lambda>f g m. (Option.map g \<circ> m \<circ> f)" .
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    36
51161
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    37
40605
0bc28f978bcf mapper for mapping type
haftmann
parents: 39380
diff changeset
    38
subsection {* Functorial structure *}
0bc28f978bcf mapper for mapping type
haftmann
parents: 39380
diff changeset
    39
41505
6d19301074cf "enriched_type" replaces less specific "type_lifting"
haftmann
parents: 41372
diff changeset
    40
enriched_type map: map
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    41
  by (transfer, auto simp add: fun_eq_iff Option.map.compositionality Option.map.id)+
40605
0bc28f978bcf mapper for mapping type
haftmann
parents: 39380
diff changeset
    42
51161
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    43
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    44
subsection {* Derived operations *}
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    45
35194
a6c573d13385 added ordered_keys
haftmann
parents: 35157
diff changeset
    46
definition ordered_keys :: "('a\<Colon>linorder, 'b) mapping \<Rightarrow> 'a list" where
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
    47
  "ordered_keys m = (if finite (keys m) then sorted_list_of_set (keys m) else [])"
35194
a6c573d13385 added ordered_keys
haftmann
parents: 35157
diff changeset
    48
35157
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
    49
definition is_empty :: "('a, 'b) mapping \<Rightarrow> bool" where
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
    50
  "is_empty m \<longleftrightarrow> keys m = {}"
35157
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
    51
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
    52
definition size :: "('a, 'b) mapping \<Rightarrow> nat" where
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
    53
  "size m = (if finite (keys m) then card (keys m) else 0)"
35157
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
    54
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
    55
definition replace :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
    56
  "replace k v m = (if k \<in> keys m then update k v m else m)"
29814
15344c0899e1 added replace operation
haftmann
parents: 29708
diff changeset
    57
37026
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
    58
definition default :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
    59
  "default k v m = (if k \<in> keys m then m else update k v m)"
37026
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
    60
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    61
lift_definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" is
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    62
  "\<lambda>k f m. (case m k of None \<Rightarrow> m
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    63
    | Some v \<Rightarrow> m (k \<mapsto> (f v)))" .
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    64
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    65
lemma map_entry_code [code]: "map_entry k f m = (case lookup m k of None \<Rightarrow> m
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49973
diff changeset
    66
    | Some v \<Rightarrow> update k (f v) m)"
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49973
diff changeset
    67
  by transfer rule
37026
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
    68
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
    69
definition map_default :: "'a \<Rightarrow> 'b \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
    70
  "map_default k v f m = map_entry k f (default k v m)" 
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
    71
51161
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    72
instantiation mapping :: (type, type) equal
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    73
begin
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    74
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    75
definition
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    76
  "HOL.equal m1 m2 \<longleftrightarrow> (\<forall>k. lookup m1 k = lookup m2 k)"
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    77
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    78
instance proof
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    79
qed (unfold equal_mapping_def, transfer, auto)
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    80
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    81
end
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    82
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    83
lemma [transfer_rule]:
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    84
  "fun_rel cr_mapping (fun_rel cr_mapping HOL.iff) HOL.eq HOL.equal"
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    85
  by (unfold equal) transfer_prover
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    86
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
    87
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    88
subsection {* Properties *}
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    89
49973
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
    90
lemma lookup_update: "lookup (update k v m) k = Some v" 
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
    91
  by transfer simp
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
    92
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
    93
lemma lookup_update_neq: "k \<noteq> k' \<Longrightarrow> lookup (update k v m) k' = lookup m k'" 
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
    94
  by transfer simp
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
    95
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
    96
lemma lookup_empty: "lookup empty k = None" 
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
    97
  by transfer simp
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
    98
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
    99
lemma keys_is_none_rep [code_unfold]:
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   100
  "k \<in> keys m \<longleftrightarrow> \<not> (Option.is_none (lookup m k))"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   101
  by transfer (auto simp add: is_none_def)
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   102
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   103
lemma tabulate_alt_def:
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   104
  "map_of (List.map (\<lambda>k. (k, f k)) ks) = (Some o f) |` set ks"
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   105
  by (induct ks) (auto simp add: tabulate_def restrict_map_def)
29826
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
   106
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   107
lemma update_update:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   108
  "update k v (update k w m) = update k v m"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   109
  "k \<noteq> l \<Longrightarrow> update k v (update l w m) = update l w (update k v m)"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   110
  by (transfer, simp add: fun_upd_twist)+
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   111
35157
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   112
lemma update_delete [simp]:
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   113
  "update k v (delete k m) = update k v m"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   114
  by transfer simp
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   115
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   116
lemma delete_update:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   117
  "delete k (update k v m) = delete k m"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   118
  "k \<noteq> l \<Longrightarrow> delete k (update l v m) = update l v (delete k m)"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   119
  by (transfer, simp add: fun_upd_twist)+
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   120
35157
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   121
lemma delete_empty [simp]:
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   122
  "delete k empty = empty"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   123
  by transfer simp
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   124
35157
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   125
lemma replace_update:
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   126
  "k \<notin> keys m \<Longrightarrow> replace k v m = m"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   127
  "k \<in> keys m \<Longrightarrow> replace k v m = update k v m"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   128
  by (transfer, auto simp add: replace_def fun_upd_twist)+
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   129
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   130
lemma size_empty [simp]:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   131
  "size empty = 0"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   132
  unfolding size_def by transfer simp
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   133
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   134
lemma size_update:
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   135
  "finite (keys m) \<Longrightarrow> size (update k v m) =
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   136
    (if k \<in> keys m then size m else Suc (size m))"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   137
  unfolding size_def by transfer (auto simp add: insert_dom)
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   138
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   139
lemma size_delete:
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   140
  "size (delete k m) = (if k \<in> keys m then size m - 1 else size m)"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   141
  unfolding size_def by transfer simp
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   142
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   143
lemma size_tabulate [simp]:
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   144
  "size (tabulate ks f) = length (remdups ks)"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   145
  unfolding size_def by transfer (auto simp add: tabulate_alt_def card_set comp_def)
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   146
29831
5dc920623bb1 Isar proof
haftmann
parents: 29828
diff changeset
   147
lemma bulkload_tabulate:
29826
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
   148
  "bulkload xs = tabulate [0..<length xs] (nth xs)"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   149
  by transfer (auto simp add: tabulate_alt_def)
29826
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
   150
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   151
lemma is_empty_empty [simp]:
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   152
  "is_empty empty"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   153
  unfolding is_empty_def by transfer simp 
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   154
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   155
lemma is_empty_update [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   156
  "\<not> is_empty (update k v m)"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   157
  unfolding is_empty_def by transfer simp
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   158
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   159
lemma is_empty_delete:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   160
  "is_empty (delete k m) \<longleftrightarrow> is_empty m \<or> keys m = {k}"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   161
  unfolding is_empty_def by transfer (auto simp del: dom_eq_empty_conv)
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   162
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   163
lemma is_empty_replace [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   164
  "is_empty (replace k v m) \<longleftrightarrow> is_empty m"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   165
  unfolding is_empty_def replace_def by transfer auto
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   166
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   167
lemma is_empty_default [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   168
  "\<not> is_empty (default k v m)"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   169
  unfolding is_empty_def default_def by transfer auto
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   170
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   171
lemma is_empty_map_entry [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   172
  "is_empty (map_entry k f m) \<longleftrightarrow> is_empty m"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   173
  unfolding is_empty_def 
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   174
  apply transfer by (case_tac "m k") auto
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   175
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   176
lemma is_empty_map_default [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   177
  "\<not> is_empty (map_default k v f m)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   178
  by (simp add: map_default_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   179
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   180
lemma keys_empty [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   181
  "keys empty = {}"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   182
  by transfer simp
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   183
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   184
lemma keys_update [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   185
  "keys (update k v m) = insert k (keys m)"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   186
  by transfer simp
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   187
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   188
lemma keys_delete [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   189
  "keys (delete k m) = keys m - {k}"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   190
  by transfer simp
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   191
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   192
lemma keys_replace [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   193
  "keys (replace k v m) = keys m"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   194
  unfolding replace_def by transfer (simp add: insert_absorb)
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   195
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   196
lemma keys_default [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   197
  "keys (default k v m) = insert k (keys m)"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   198
  unfolding default_def by transfer (simp add: insert_absorb)
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   199
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   200
lemma keys_map_entry [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   201
  "keys (map_entry k f m) = keys m"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   202
  apply transfer by (case_tac "m k") auto
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   203
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   204
lemma keys_map_default [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   205
  "keys (map_default k v f m) = insert k (keys m)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   206
  by (simp add: map_default_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   207
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   208
lemma keys_tabulate [simp]:
37026
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
   209
  "keys (tabulate ks f) = set ks"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   210
  by transfer (simp add: map_of_map_restrict o_def)
37026
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
   211
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   212
lemma keys_bulkload [simp]:
37026
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
   213
  "keys (bulkload xs) = {0..<length xs}"
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
   214
  by (simp add: keys_tabulate bulkload_tabulate)
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
   215
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   216
lemma distinct_ordered_keys [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   217
  "distinct (ordered_keys m)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   218
  by (simp add: ordered_keys_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   219
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   220
lemma ordered_keys_infinite [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   221
  "\<not> finite (keys m) \<Longrightarrow> ordered_keys m = []"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   222
  by (simp add: ordered_keys_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   223
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   224
lemma ordered_keys_empty [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   225
  "ordered_keys empty = []"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   226
  by (simp add: ordered_keys_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   227
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   228
lemma ordered_keys_update [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   229
  "k \<in> keys m \<Longrightarrow> ordered_keys (update k v m) = ordered_keys m"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   230
  "finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (update k v m) = insort k (ordered_keys m)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   231
  by (simp_all add: ordered_keys_def) (auto simp only: sorted_list_of_set_insert [symmetric] insert_absorb)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   232
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   233
lemma ordered_keys_delete [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   234
  "ordered_keys (delete k m) = remove1 k (ordered_keys m)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   235
proof (cases "finite (keys m)")
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   236
  case False then show ?thesis by simp
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   237
next
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   238
  case True note fin = True
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   239
  show ?thesis
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   240
  proof (cases "k \<in> keys m")
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   241
    case False with fin have "k \<notin> set (sorted_list_of_set (keys m))" by simp
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   242
    with False show ?thesis by (simp add: ordered_keys_def remove1_idem)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   243
  next
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   244
    case True with fin show ?thesis by (simp add: ordered_keys_def sorted_list_of_set_remove)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   245
  qed
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   246
qed
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   247
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   248
lemma ordered_keys_replace [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   249
  "ordered_keys (replace k v m) = ordered_keys m"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   250
  by (simp add: replace_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   251
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   252
lemma ordered_keys_default [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   253
  "k \<in> keys m \<Longrightarrow> ordered_keys (default k v m) = ordered_keys m"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   254
  "finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (default k v m) = insort k (ordered_keys m)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   255
  by (simp_all add: default_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   256
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   257
lemma ordered_keys_map_entry [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   258
  "ordered_keys (map_entry k f m) = ordered_keys m"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   259
  by (simp add: ordered_keys_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   260
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   261
lemma ordered_keys_map_default [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   262
  "k \<in> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = ordered_keys m"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   263
  "finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = insort k (ordered_keys m)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   264
  by (simp_all add: map_default_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   265
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   266
lemma ordered_keys_tabulate [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   267
  "ordered_keys (tabulate ks f) = sort (remdups ks)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   268
  by (simp add: ordered_keys_def sorted_list_of_set_sort_remdups)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   269
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   270
lemma ordered_keys_bulkload [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   271
  "ordered_keys (bulkload ks) = [0..<length ks]"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   272
  by (simp add: ordered_keys_def)
36110
4ab91a42666a lemma is_empty_empty
haftmann
parents: 35194
diff changeset
   273
31459
ae39b7b2a68a added trees implementing mappings
haftmann
parents: 30663
diff changeset
   274
37700
bd90378b8171 refrain from using datatype declaration -- opens chance for quickcheck later on
haftmann
parents: 37299
diff changeset
   275
subsection {* Code generator setup *}
31459
ae39b7b2a68a added trees implementing mappings
haftmann
parents: 30663
diff changeset
   276
37701
411717732710 explicit code_datatype declaration prevents multiple instantiations later on
haftmann
parents: 37700
diff changeset
   277
code_datatype empty update
411717732710 explicit code_datatype declaration prevents multiple instantiations later on
haftmann
parents: 37700
diff changeset
   278
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   279
hide_const (open) empty is_empty rep lookup update delete ordered_keys keys size
40605
0bc28f978bcf mapper for mapping type
haftmann
parents: 39380
diff changeset
   280
  replace default map_entry map_default tabulate bulkload map
35157
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   281
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49973
diff changeset
   282
end
51161
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   283