src/HOL/Library/Mapping.thy
author wenzelm
Tue, 21 Oct 2014 20:19:14 +0200
changeset 58752 2077bc9558cf
parent 56545 8f1e7596deb7
child 58881 b9556a055632
permissions -rw-r--r--
support for begin/end matching;
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
53013
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 51379
diff changeset
     8
imports Main
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
     9
begin
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    10
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    11
subsection {* Parametricity transfer rules *}
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    12
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    13
lemma map_of_foldr: -- {* FIXME move *}
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    14
  "map_of xs = foldr (\<lambda>(k, v) m. m(k \<mapsto> v)) xs Map.empty"
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    15
  using map_add_map_of_foldr [of Map.empty] by auto
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    16
53013
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 51379
diff changeset
    17
context
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 51379
diff changeset
    18
begin
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    19
53013
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 51379
diff changeset
    20
interpretation lifting_syntax .
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 51379
diff changeset
    21
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    22
lemma empty_parametric:
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    23
  "(A ===> rel_option B) Map.empty Map.empty"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    24
  by transfer_prover
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    25
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    26
lemma lookup_parametric: "((A ===> B) ===> A ===> B) (\<lambda>m k. m k) (\<lambda>m k. m k)"
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    27
  by transfer_prover
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    28
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    29
lemma update_parametric:
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    30
  assumes [transfer_rule]: "bi_unique A"
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    31
  shows "(A ===> B ===> (A ===> rel_option B) ===> A ===> rel_option B)
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    32
    (\<lambda>k v m. m(k \<mapsto> v)) (\<lambda>k v m. m(k \<mapsto> v))"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    33
  by transfer_prover
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    34
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    35
lemma delete_parametric:
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    36
  assumes [transfer_rule]: "bi_unique A"
55525
70b7e91fa1f9 folded 'rel_option' into 'option_rel'
blanchet
parents: 55467
diff changeset
    37
  shows "(A ===> (A ===> rel_option B) ===> A ===> rel_option B) 
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    38
    (\<lambda>k m. m(k := None)) (\<lambda>k m. m(k := None))"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    39
  by transfer_prover
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    40
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    41
lemma is_none_parametric [transfer_rule]:
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    42
  "(rel_option A ===> HOL.eq) Option.is_none Option.is_none"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    43
  by (auto simp add: is_none_def rel_fun_def rel_option_iff split: option.split)
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    44
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    45
lemma dom_parametric:
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    46
  assumes [transfer_rule]: "bi_total A"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55525
diff changeset
    47
  shows "((A ===> rel_option B) ===> rel_set A) dom dom" 
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    48
  unfolding dom_def [abs_def] is_none_def [symmetric] by transfer_prover
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    49
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    50
lemma map_of_parametric [transfer_rule]:
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    51
  assumes [transfer_rule]: "bi_unique R1"
55944
7ab8f003fe41 renamed 'prod_rel' to 'rel_prod'
blanchet
parents: 55938
diff changeset
    52
  shows "(list_all2 (rel_prod R1 R2) ===> R1 ===> rel_option R2) map_of map_of"
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    53
  unfolding map_of_def by transfer_prover
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    54
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    55
lemma map_entry_parametric [transfer_rule]:
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    56
  assumes [transfer_rule]: "bi_unique A"
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    57
  shows "(A ===> (B ===> B) ===> (A ===> rel_option B) ===> A ===> rel_option B) 
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    58
    (\<lambda>k f m. (case m k of None \<Rightarrow> m
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    59
      | Some v \<Rightarrow> m (k \<mapsto> (f v)))) (\<lambda>k f m. (case m k of None \<Rightarrow> m
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    60
      | Some v \<Rightarrow> m (k \<mapsto> (f v))))"
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    61
  by transfer_prover
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    62
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    63
lemma tabulate_parametric: 
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    64
  assumes [transfer_rule]: "bi_unique A"
55525
70b7e91fa1f9 folded 'rel_option' into 'option_rel'
blanchet
parents: 55467
diff changeset
    65
  shows "(list_all2 A ===> (A ===> B) ===> A ===> rel_option B) 
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    66
    (\<lambda>ks f. (map_of (map (\<lambda>k. (k, f k)) ks))) (\<lambda>ks f. (map_of (map (\<lambda>k. (k, f k)) ks)))"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    67
  by transfer_prover
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    68
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    69
lemma bulkload_parametric: 
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    70
  "(list_all2 A ===> HOL.eq ===> rel_option A) 
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    71
    (\<lambda>xs k. if k < length xs then Some (xs ! k) else None) (\<lambda>xs k. if k < length xs then Some (xs ! k) else None)"
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    72
proof
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    73
  fix xs ys
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    74
  assume "list_all2 A xs ys"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    75
  then show "(HOL.eq ===> rel_option A)
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    76
    (\<lambda>k. if k < length xs then Some (xs ! k) else None)
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    77
    (\<lambda>k. if k < length ys then Some (ys ! k) else None)"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    78
    apply induct
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    79
    apply auto
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    80
    unfolding rel_fun_def
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    81
    apply clarsimp 
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    82
    apply (case_tac xa) 
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    83
    apply (auto dest: list_all2_lengthD list_all2_nthD)
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    84
    done
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    85
qed
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    86
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    87
lemma map_parametric: 
55525
70b7e91fa1f9 folded 'rel_option' into 'option_rel'
blanchet
parents: 55467
diff changeset
    88
  "((A ===> B) ===> (C ===> D) ===> (B ===> rel_option C) ===> A ===> rel_option D) 
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    89
     (\<lambda>f g m. (map_option g \<circ> m \<circ> f)) (\<lambda>f g m. (map_option g \<circ> m \<circ> f))"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    90
  by transfer_prover
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    91
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
    92
end
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
    93
53013
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 51379
diff changeset
    94
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    95
subsection {* Type definition and primitive operations *}
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
    96
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 45231
diff changeset
    97
typedef ('a, 'b) mapping = "UNIV :: ('a \<rightharpoonup> 'b) set"
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    98
  morphisms rep Mapping
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
    99
  ..
37700
bd90378b8171 refrain from using datatype declaration -- opens chance for quickcheck later on
haftmann
parents: 37299
diff changeset
   100
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   101
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
   102
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   103
lift_definition empty :: "('a, 'b) mapping"
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   104
  is Map.empty parametric empty_parametric .
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   105
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   106
lift_definition lookup :: "('a, 'b) mapping \<Rightarrow> 'a \<Rightarrow> 'b option"
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   107
  is "\<lambda>m k. m k" parametric lookup_parametric .
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   108
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   109
lift_definition update :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   110
  is "\<lambda>k v m. m(k \<mapsto> v)" parametric update_parametric .
37700
bd90378b8171 refrain from using datatype declaration -- opens chance for quickcheck later on
haftmann
parents: 37299
diff changeset
   111
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   112
lift_definition delete :: "'a \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   113
  is "\<lambda>k m. m(k := None)" parametric delete_parametric .
39380
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 39302
diff changeset
   114
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   115
lift_definition keys :: "('a, 'b) mapping \<Rightarrow> 'a set"
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   116
  is dom parametric dom_parametric .
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   117
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   118
lift_definition tabulate :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping"
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   119
  is "\<lambda>ks f. (map_of (List.map (\<lambda>k. (k, f k)) ks))" parametric tabulate_parametric .
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   120
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   121
lift_definition bulkload :: "'a list \<Rightarrow> (nat, 'a) mapping"
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   122
  is "\<lambda>xs k. if k < length xs then Some (xs ! k) else None" parametric bulkload_parametric .
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   123
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   124
lift_definition map :: "('c \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('c, 'd) mapping"
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   125
  is "\<lambda>f g m. (map_option g \<circ> m \<circ> f)" parametric map_parametric .
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   126
51161
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   127
40605
0bc28f978bcf mapper for mapping type
haftmann
parents: 39380
diff changeset
   128
subsection {* Functorial structure *}
0bc28f978bcf mapper for mapping type
haftmann
parents: 39380
diff changeset
   129
55467
a5c9002bc54d renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
blanchet
parents: 55466
diff changeset
   130
functor map: map
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 54853
diff changeset
   131
  by (transfer, auto simp add: fun_eq_iff option.map_comp option.map_id)+
40605
0bc28f978bcf mapper for mapping type
haftmann
parents: 39380
diff changeset
   132
51161
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   133
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   134
subsection {* Derived operations *}
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   135
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   136
definition ordered_keys :: "('a\<Colon>linorder, 'b) mapping \<Rightarrow> 'a list"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   137
where
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   138
  "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
   139
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   140
definition is_empty :: "('a, 'b) mapping \<Rightarrow> bool"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   141
where
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   142
  "is_empty m \<longleftrightarrow> keys m = {}"
35157
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   143
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   144
definition size :: "('a, 'b) mapping \<Rightarrow> nat"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   145
where
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   146
  "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
   147
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   148
definition replace :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   149
where
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   150
  "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
   151
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   152
definition default :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   153
where
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   154
  "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
   155
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   156
text {* Manual derivation of transfer rule is non-trivial *}
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   157
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   158
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
   159
  "\<lambda>k f m. (case m k of None \<Rightarrow> m
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   160
    | Some v \<Rightarrow> m (k \<mapsto> (f v)))" parametric map_entry_parametric .
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   161
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   162
lemma map_entry_code [code]:
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   163
  "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
   164
    | 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
   165
  by transfer rule
37026
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
   166
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   167
definition map_default :: "'a \<Rightarrow> 'b \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   168
where
37026
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
   169
  "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
   170
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   171
definition of_alist :: "('k \<times> 'v) list \<Rightarrow> ('k, 'v) mapping"
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   172
where
54853
a435932a9f12 prefer Y_of_X over X_to_Y;
haftmann
parents: 53026
diff changeset
   173
  "of_alist xs = foldr (\<lambda>(k, v) m. update k v m) xs empty"
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
   174
51161
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   175
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
   176
begin
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   177
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   178
definition
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   179
  "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
   180
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   181
instance proof
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   182
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
   183
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   184
end
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   185
53013
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 51379
diff changeset
   186
context
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 51379
diff changeset
   187
begin
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   188
53013
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 51379
diff changeset
   189
interpretation lifting_syntax .
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 51379
diff changeset
   190
51161
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   191
lemma [transfer_rule]:
51379
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
   192
  assumes [transfer_rule]: "bi_total A"
6dd83e007f56 convert mappings to parametric lifting
kuncar
parents: 51375
diff changeset
   193
  assumes [transfer_rule]: "bi_unique B"
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   194
  shows "(pcr_mapping A B ===> pcr_mapping A B ===> op=) HOL.eq HOL.equal"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   195
  by (unfold equal) transfer_prover
51161
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   196
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   197
lemma of_alist_transfer [transfer_rule]:
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   198
  assumes [transfer_rule]: "bi_unique R1"
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   199
  shows "(list_all2 (rel_prod R1 R2) ===> pcr_mapping R1 R2) map_of of_alist"
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   200
  unfolding of_alist_def [abs_def] map_of_foldr [abs_def] by transfer_prover
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   201
53013
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 51379
diff changeset
   202
end
51161
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 49975
diff changeset
   203
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   204
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   205
subsection {* Properties *}
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   206
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   207
lemma lookup_update:
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   208
  "lookup (update k v m) k = Some v" 
49973
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
   209
  by transfer simp
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
   210
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   211
lemma lookup_update_neq:
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   212
  "k \<noteq> k' \<Longrightarrow> lookup (update k v m) k' = lookup m k'" 
49973
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
   213
  by transfer simp
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
   214
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   215
lemma lookup_empty:
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   216
  "lookup empty k = None" 
49973
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
   217
  by transfer simp
e84baadd4963 new theorems
kuncar
parents: 49939
diff changeset
   218
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   219
lemma keys_is_none_rep [code_unfold]:
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   220
  "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
   221
  by transfer (auto simp add: is_none_def)
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   222
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   223
lemma update_update:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   224
  "update k v (update k w m) = update k v m"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   225
  "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
   226
  by (transfer, simp add: fun_upd_twist)+
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   227
35157
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   228
lemma update_delete [simp]:
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   229
  "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
   230
  by transfer simp
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   231
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   232
lemma delete_update:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   233
  "delete k (update k v m) = delete k m"
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   234
  "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
   235
  by (transfer, simp add: fun_upd_twist)+
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   236
35157
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   237
lemma delete_empty [simp]:
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   238
  "delete k empty = empty"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   239
  by transfer simp
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   240
35157
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   241
lemma replace_update:
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   242
  "k \<notin> keys m \<Longrightarrow> replace k v m = m"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   243
  "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
   244
  by (transfer, auto simp add: replace_def fun_upd_twist)+
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   245
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   246
lemma size_empty [simp]:
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   247
  "size empty = 0"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   248
  unfolding size_def by transfer simp
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   249
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   250
lemma size_update:
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   251
  "finite (keys m) \<Longrightarrow> size (update k v m) =
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   252
    (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
   253
  unfolding size_def by transfer (auto simp add: insert_dom)
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   254
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   255
lemma size_delete:
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   256
  "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
   257
  unfolding size_def by transfer simp
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   258
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   259
lemma size_tabulate [simp]:
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   260
  "size (tabulate ks f) = length (remdups ks)"
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   261
  unfolding size_def by transfer (auto simp add: map_of_map_restrict  card_set comp_def)
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents:
diff changeset
   262
29831
5dc920623bb1 Isar proof
haftmann
parents: 29828
diff changeset
   263
lemma bulkload_tabulate:
29826
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
   264
  "bulkload xs = tabulate [0..<length xs] (nth xs)"
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   265
  by transfer (auto simp add: map_of_map_restrict)
29826
5132da6ebca3 added bulkload
haftmann
parents: 29814
diff changeset
   266
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   267
lemma is_empty_empty [simp]:
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   268
  "is_empty empty"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   269
  unfolding is_empty_def by transfer simp 
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   270
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   271
lemma is_empty_update [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   272
  "\<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
   273
  unfolding is_empty_def by transfer simp
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   274
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   275
lemma is_empty_delete:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   276
  "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
   277
  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
   278
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   279
lemma is_empty_replace [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   280
  "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
   281
  unfolding is_empty_def replace_def by transfer auto
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   282
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   283
lemma is_empty_default [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   284
  "\<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
   285
  unfolding is_empty_def default_def by transfer auto
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   286
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   287
lemma is_empty_map_entry [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   288
  "is_empty (map_entry k f m) \<longleftrightarrow> is_empty m"
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   289
  unfolding is_empty_def by transfer (auto split: option.split)
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   290
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   291
lemma is_empty_map_default [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   292
  "\<not> is_empty (map_default k v f m)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   293
  by (simp add: map_default_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   294
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56529
diff changeset
   295
lemma keys_dom_lookup:
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56529
diff changeset
   296
  "keys m = dom (Mapping.lookup m)"
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56529
diff changeset
   297
  by transfer rule
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56529
diff changeset
   298
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   299
lemma keys_empty [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   300
  "keys empty = {}"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   301
  by transfer simp
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   302
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   303
lemma keys_update [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   304
  "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
   305
  by transfer simp
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   306
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   307
lemma keys_delete [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   308
  "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
   309
  by transfer simp
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   310
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   311
lemma keys_replace [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   312
  "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
   313
  unfolding replace_def by transfer (simp add: insert_absorb)
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   314
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   315
lemma keys_default [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   316
  "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
   317
  unfolding default_def by transfer (simp add: insert_absorb)
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   318
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   319
lemma keys_map_entry [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   320
  "keys (map_entry k f m) = keys m"
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   321
  by transfer (auto split: option.split)
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   322
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   323
lemma keys_map_default [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   324
  "keys (map_default k v f m) = insert k (keys m)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   325
  by (simp add: map_default_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   326
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   327
lemma keys_tabulate [simp]:
37026
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
   328
  "keys (tabulate ks f) = set ks"
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   329
  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
   330
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   331
lemma keys_bulkload [simp]:
37026
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
   332
  "keys (bulkload xs) = {0..<length xs}"
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   333
  by (simp add: bulkload_tabulate)
37026
7e8979a155ae operations default, map_entry, map_default; more lemmas
haftmann
parents: 36176
diff changeset
   334
37052
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   335
lemma distinct_ordered_keys [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   336
  "distinct (ordered_keys m)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   337
  by (simp add: ordered_keys_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   338
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   339
lemma ordered_keys_infinite [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   340
  "\<not> finite (keys m) \<Longrightarrow> ordered_keys m = []"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   341
  by (simp add: ordered_keys_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   342
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   343
lemma ordered_keys_empty [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   344
  "ordered_keys empty = []"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   345
  by (simp add: ordered_keys_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   346
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   347
lemma ordered_keys_update [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   348
  "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
   349
  "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
   350
  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
   351
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   352
lemma ordered_keys_delete [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   353
  "ordered_keys (delete k m) = remove1 k (ordered_keys m)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   354
proof (cases "finite (keys m)")
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   355
  case False then show ?thesis by simp
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   356
next
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   357
  case True note fin = True
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   358
  show ?thesis
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   359
  proof (cases "k \<in> keys m")
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   360
    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
   361
    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
   362
  next
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   363
    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
   364
  qed
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   365
qed
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   366
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   367
lemma ordered_keys_replace [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   368
  "ordered_keys (replace k v m) = ordered_keys m"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   369
  by (simp add: replace_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   370
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   371
lemma ordered_keys_default [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   372
  "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
   373
  "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
   374
  by (simp_all add: default_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   375
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   376
lemma ordered_keys_map_entry [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   377
  "ordered_keys (map_entry k f m) = ordered_keys m"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   378
  by (simp add: ordered_keys_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   379
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   380
lemma ordered_keys_map_default [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   381
  "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
   382
  "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
   383
  by (simp_all add: map_default_def)
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   384
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   385
lemma ordered_keys_tabulate [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   386
  "ordered_keys (tabulate ks f) = sort (remdups ks)"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   387
  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
   388
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   389
lemma ordered_keys_bulkload [simp]:
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   390
  "ordered_keys (bulkload ks) = [0..<length ks]"
80dd92673fca more lemmas about mappings, in particular keys
haftmann
parents: 37026
diff changeset
   391
  by (simp add: ordered_keys_def)
36110
4ab91a42666a lemma is_empty_empty
haftmann
parents: 35194
diff changeset
   392
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   393
lemma tabulate_fold:
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   394
  "tabulate xs f = fold (\<lambda>k m. update k (f k) m) xs empty"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   395
proof transfer
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   396
  fix f :: "'a \<Rightarrow> 'b" and xs
56529
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   397
  have "map_of (List.map (\<lambda>k. (k, f k)) xs) = foldr (\<lambda>k m. m(k \<mapsto> f k)) xs Map.empty"
aff193f53a64 restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents: 56528
diff changeset
   398
    by (simp add: foldr_map comp_def map_of_foldr)
56528
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   399
  also have "foldr (\<lambda>k m. m(k \<mapsto> f k)) xs = fold (\<lambda>k m. m(k \<mapsto> f k)) xs"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   400
    by (rule foldr_fold) (simp add: fun_eq_iff)
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   401
  ultimately show "map_of (List.map (\<lambda>k. (k, f k)) xs) = fold (\<lambda>k m. m(k \<mapsto> f k)) xs Map.empty"
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   402
    by simp
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   403
qed
f732e6f3bf7f removed duplication and tuned
haftmann
parents: 55945
diff changeset
   404
31459
ae39b7b2a68a added trees implementing mappings
haftmann
parents: 30663
diff changeset
   405
37700
bd90378b8171 refrain from using datatype declaration -- opens chance for quickcheck later on
haftmann
parents: 37299
diff changeset
   406
subsection {* Code generator setup *}
31459
ae39b7b2a68a added trees implementing mappings
haftmann
parents: 30663
diff changeset
   407
37701
411717732710 explicit code_datatype declaration prevents multiple instantiations later on
haftmann
parents: 37700
diff changeset
   408
code_datatype empty update
411717732710 explicit code_datatype declaration prevents multiple instantiations later on
haftmann
parents: 37700
diff changeset
   409
49929
70300f1b6835 update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents: 49834
diff changeset
   410
hide_const (open) empty is_empty rep lookup update delete ordered_keys keys size
54853
a435932a9f12 prefer Y_of_X over X_to_Y;
haftmann
parents: 53026
diff changeset
   411
  replace default map_entry map_default tabulate bulkload map of_alist
35157
73cd6f78c86d more close integration with theory Map
haftmann
parents: 33640
diff changeset
   412
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49973
diff changeset
   413
end