use lifting/transfer formalization of RBT from Lift_RBT
authorkuncar
Tue, 31 Jul 2012 13:55:39 +0200
changeset 49637caaa1a02c650
parent 49636 877df57629e3
child 49638 bea613f2543d
use lifting/transfer formalization of RBT from Lift_RBT
src/HOL/Library/RBT.thy
src/HOL/Quotient_Examples/Lift_RBT.thy
     1.1 --- a/src/HOL/Library/RBT.thy	Tue Jul 31 13:55:39 2012 +0200
     1.2 +++ b/src/HOL/Library/RBT.thy	Tue Jul 31 13:55:39 2012 +0200
     1.3 @@ -1,9 +1,10 @@
     1.4 -(* Author: Florian Haftmann, TU Muenchen *)
     1.5 +(*  Title:      HOL/Library/RBT.thy
     1.6 +    Author:     Lukas Bulwahn and Ondrej Kuncar
     1.7 +*)
     1.8  
     1.9 -header {* Abstract type of Red-Black Trees *}
    1.10 +header {* Abstract type of RBT trees *}
    1.11  
    1.12 -(*<*)
    1.13 -theory RBT
    1.14 +theory RBT 
    1.15  imports Main RBT_Impl
    1.16  begin
    1.17  
    1.18 @@ -11,8 +12,9 @@
    1.19  
    1.20  typedef (open) ('a, 'b) rbt = "{t :: ('a\<Colon>linorder, 'b) RBT_Impl.rbt. is_rbt t}"
    1.21    morphisms impl_of RBT
    1.22 -proof
    1.23 -  show "RBT_Impl.Empty \<in> {t. is_rbt t}" by simp
    1.24 +proof -
    1.25 +  have "RBT_Impl.Empty \<in> ?rbt" by simp
    1.26 +  then show ?thesis ..
    1.27  qed
    1.28  
    1.29  lemma rbt_eq_iff:
    1.30 @@ -31,63 +33,45 @@
    1.31    "RBT (impl_of t) = t"
    1.32    by (simp add: impl_of_inverse)
    1.33  
    1.34 -
    1.35  subsection {* Primitive operations *}
    1.36  
    1.37 -definition lookup :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b" where
    1.38 -  [code]: "lookup t = rbt_lookup (impl_of t)"
    1.39 +setup_lifting type_definition_rbt
    1.40  
    1.41 -definition empty :: "('a\<Colon>linorder, 'b) rbt" where
    1.42 -  "empty = RBT RBT_Impl.Empty"
    1.43 +lift_definition lookup :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b" is "rbt_lookup" 
    1.44 +by simp
    1.45  
    1.46 -lemma impl_of_empty [code abstract]:
    1.47 -  "impl_of empty = RBT_Impl.Empty"
    1.48 -  by (simp add: empty_def RBT_inverse)
    1.49 +lift_definition empty :: "('a\<Colon>linorder, 'b) rbt" is RBT_Impl.Empty 
    1.50 +by (simp add: empty_def)
    1.51  
    1.52 -definition insert :: "'a\<Colon>linorder \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
    1.53 -  "insert k v t = RBT (rbt_insert k v (impl_of t))"
    1.54 +lift_definition insert :: "'a\<Colon>linorder \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_insert" 
    1.55 +by simp
    1.56  
    1.57 -lemma impl_of_insert [code abstract]:
    1.58 -  "impl_of (insert k v t) = rbt_insert k v (impl_of t)"
    1.59 -  by (simp add: insert_def RBT_inverse)
    1.60 +lift_definition delete :: "'a\<Colon>linorder \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_delete" 
    1.61 +by simp
    1.62  
    1.63 -definition delete :: "'a\<Colon>linorder \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
    1.64 -  "delete k t = RBT (rbt_delete k (impl_of t))"
    1.65 +lift_definition entries :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a \<times> 'b) list" is RBT_Impl.entries
    1.66 +by simp
    1.67  
    1.68 -lemma impl_of_delete [code abstract]:
    1.69 -  "impl_of (delete k t) = rbt_delete k (impl_of t)"
    1.70 -  by (simp add: delete_def RBT_inverse)
    1.71 +lift_definition keys :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a list" is RBT_Impl.keys 
    1.72 +by simp
    1.73  
    1.74 -definition entries :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a \<times> 'b) list" where
    1.75 -  [code]: "entries t = RBT_Impl.entries (impl_of t)"
    1.76 +lift_definition bulkload :: "('a\<Colon>linorder \<times> 'b) list \<Rightarrow> ('a, 'b) rbt" is "rbt_bulkload" 
    1.77 +by simp
    1.78  
    1.79 -definition keys :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a list" where
    1.80 -  [code]: "keys t = RBT_Impl.keys (impl_of t)"
    1.81 +lift_definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is rbt_map_entry 
    1.82 +by simp
    1.83  
    1.84 -definition bulkload :: "('a\<Colon>linorder \<times> 'b) list \<Rightarrow> ('a, 'b) rbt" where
    1.85 -  "bulkload xs = RBT (rbt_bulkload xs)"
    1.86 +lift_definition map :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is RBT_Impl.map
    1.87 +by simp
    1.88  
    1.89 -lemma impl_of_bulkload [code abstract]:
    1.90 -  "impl_of (bulkload xs) = rbt_bulkload xs"
    1.91 -  by (simp add: bulkload_def RBT_inverse)
    1.92 +lift_definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c"  is RBT_Impl.fold 
    1.93 +by simp
    1.94  
    1.95 -definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
    1.96 -  "map_entry k f t = RBT (rbt_map_entry k f (impl_of t))"
    1.97 +lift_definition union :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_union"
    1.98 +by (simp add: rbt_union_is_rbt)
    1.99  
   1.100 -lemma impl_of_map_entry [code abstract]:
   1.101 -  "impl_of (map_entry k f t) = rbt_map_entry k f (impl_of t)"
   1.102 -  by (simp add: map_entry_def RBT_inverse)
   1.103 -
   1.104 -definition map :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
   1.105 -  "map f t = RBT (RBT_Impl.map f (impl_of t))"
   1.106 -
   1.107 -lemma impl_of_map [code abstract]:
   1.108 -  "impl_of (map f t) = RBT_Impl.map f (impl_of t)"
   1.109 -  by (simp add: map_def RBT_inverse)
   1.110 -
   1.111 -definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c" where
   1.112 -  [code]: "fold f t = RBT_Impl.fold f (impl_of t)"
   1.113 -
   1.114 +lift_definition foldi :: "('c \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a :: linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c"
   1.115 +  is RBT_Impl.foldi by simp
   1.116  
   1.117  subsection {* Derived operations *}
   1.118  
   1.119 @@ -103,15 +87,15 @@
   1.120  
   1.121  lemma lookup_impl_of:
   1.122    "rbt_lookup (impl_of t) = lookup t"
   1.123 -  by (simp add: lookup_def)
   1.124 +  by transfer (rule refl)
   1.125  
   1.126  lemma entries_impl_of:
   1.127    "RBT_Impl.entries (impl_of t) = entries t"
   1.128 -  by (simp add: entries_def)
   1.129 +  by transfer (rule refl)
   1.130  
   1.131  lemma keys_impl_of:
   1.132    "RBT_Impl.keys (impl_of t) = keys t"
   1.133 -  by (simp add: keys_def)
   1.134 +  by transfer (rule refl)
   1.135  
   1.136  lemma lookup_empty [simp]:
   1.137    "lookup empty = Map.empty"
   1.138 @@ -119,39 +103,43 @@
   1.139  
   1.140  lemma lookup_insert [simp]:
   1.141    "lookup (insert k v t) = (lookup t)(k \<mapsto> v)"
   1.142 -  by (simp add: insert_def lookup_RBT rbt_lookup_rbt_insert lookup_impl_of)
   1.143 +  by transfer (rule rbt_lookup_rbt_insert)
   1.144  
   1.145  lemma lookup_delete [simp]:
   1.146    "lookup (delete k t) = (lookup t)(k := None)"
   1.147 -  by (simp add: delete_def lookup_RBT rbt_lookup_rbt_delete lookup_impl_of restrict_complement_singleton_eq)
   1.148 +  by transfer (simp add: rbt_lookup_rbt_delete restrict_complement_singleton_eq)
   1.149  
   1.150  lemma map_of_entries [simp]:
   1.151    "map_of (entries t) = lookup t"
   1.152 -  by (simp add: entries_def map_of_entries lookup_impl_of)
   1.153 +  by transfer (simp add: map_of_entries)
   1.154  
   1.155  lemma entries_lookup:
   1.156    "entries t1 = entries t2 \<longleftrightarrow> lookup t1 = lookup t2"
   1.157 -  by (simp add: entries_def lookup_def entries_rbt_lookup)
   1.158 +  by transfer (simp add: entries_rbt_lookup)
   1.159  
   1.160  lemma lookup_bulkload [simp]:
   1.161    "lookup (bulkload xs) = map_of xs"
   1.162 -  by (simp add: bulkload_def lookup_RBT rbt_lookup_rbt_bulkload)
   1.163 +  by transfer (rule rbt_lookup_rbt_bulkload)
   1.164  
   1.165  lemma lookup_map_entry [simp]:
   1.166    "lookup (map_entry k f t) = (lookup t)(k := Option.map f (lookup t k))"
   1.167 -  by (simp add: map_entry_def lookup_RBT rbt_lookup_rbt_map_entry lookup_impl_of)
   1.168 +  by transfer (rule rbt_lookup_rbt_map_entry)
   1.169  
   1.170  lemma lookup_map [simp]:
   1.171    "lookup (map f t) k = Option.map (f k) (lookup t k)"
   1.172 -  by (simp add: map_def lookup_RBT rbt_lookup_map lookup_impl_of)
   1.173 +  by transfer (rule rbt_lookup_map)
   1.174  
   1.175  lemma fold_fold:
   1.176    "fold f t = List.fold (prod_case f) (entries t)"
   1.177 -  by (simp add: fold_def fun_eq_iff RBT_Impl.fold_def entries_impl_of)
   1.178 +  by transfer (rule RBT_Impl.fold_def)
   1.179 +
   1.180 +lemma impl_of_empty:
   1.181 +  "impl_of empty = RBT_Impl.Empty"
   1.182 +  by transfer (rule refl)
   1.183  
   1.184  lemma is_empty_empty [simp]:
   1.185    "is_empty t \<longleftrightarrow> t = empty"
   1.186 -  by (simp add: rbt_eq_iff is_empty_def impl_of_empty split: rbt.split)
   1.187 +  unfolding is_empty_def by transfer (simp split: rbt.split)
   1.188  
   1.189  lemma RBT_lookup_empty [simp]: (*FIXME*)
   1.190    "rbt_lookup t = Map.empty \<longleftrightarrow> t = RBT_Impl.Empty"
   1.191 @@ -159,15 +147,41 @@
   1.192  
   1.193  lemma lookup_empty_empty [simp]:
   1.194    "lookup t = Map.empty \<longleftrightarrow> t = empty"
   1.195 -  by (cases t) (simp add: empty_def lookup_def RBT_inject RBT_inverse)
   1.196 +  by transfer (rule RBT_lookup_empty)
   1.197  
   1.198  lemma sorted_keys [iff]:
   1.199    "sorted (keys t)"
   1.200 -  by (simp add: keys_def RBT_Impl.keys_def rbt_sorted_entries)
   1.201 +  by transfer (simp add: RBT_Impl.keys_def rbt_sorted_entries)
   1.202  
   1.203  lemma distinct_keys [iff]:
   1.204    "distinct (keys t)"
   1.205 -  by (simp add: keys_def RBT_Impl.keys_def distinct_entries)
   1.206 +  by transfer (simp add: RBT_Impl.keys_def distinct_entries)
   1.207 +
   1.208 +lemma finite_dom_lookup [simp, intro!]: "finite (dom (lookup t))"
   1.209 +  by transfer simp
   1.210 +
   1.211 +lemma lookup_union: "lookup (union s t) = lookup s ++ lookup t"
   1.212 +  by transfer (simp add: rbt_lookup_rbt_union)
   1.213 +
   1.214 +lemma lookup_in_tree: "(lookup t k = Some v) = ((k, v) \<in> set (entries t))"
   1.215 +  by transfer (simp add: rbt_lookup_in_tree)
   1.216 +
   1.217 +lemma keys_entries: "(k \<in> set (keys t)) = (\<exists>v. (k, v) \<in> set (entries t))"
   1.218 +  by transfer (simp add: keys_entries)
   1.219 +
   1.220 +lemma fold_def_alt:
   1.221 +  "fold f t = List.fold (prod_case f) (entries t)"
   1.222 +  by transfer (auto simp: RBT_Impl.fold_def)
   1.223 +
   1.224 +lemma distinct_entries: "distinct (List.map fst (entries t))"
   1.225 +  by transfer (simp add: distinct_entries)
   1.226 +
   1.227 +lemma non_empty_keys: "t \<noteq> empty \<Longrightarrow> keys t \<noteq> []"
   1.228 +  by transfer (simp add: non_empty_rbt_keys)
   1.229 +
   1.230 +lemma keys_def_alt:
   1.231 +  "keys t = List.map fst (entries t)"
   1.232 +  by transfer (simp add: RBT_Impl.keys_def)
   1.233  
   1.234  subsection {* Quickcheck generators *}
   1.235  
     2.1 --- a/src/HOL/Quotient_Examples/Lift_RBT.thy	Tue Jul 31 13:55:39 2012 +0200
     2.2 +++ b/src/HOL/Quotient_Examples/Lift_RBT.thy	Tue Jul 31 13:55:39 2012 +0200
     2.3 @@ -8,183 +8,6 @@
     2.4  imports Main "~~/src/HOL/Library/RBT_Impl"
     2.5  begin
     2.6  
     2.7 -(* TODO: Replace the ancient Library/RBT theory by this example of the lifting and transfer mechanism. *)
     2.8 -
     2.9 -subsection {* Type definition *}
    2.10 -
    2.11 -typedef (open) ('a, 'b) rbt = "{t :: ('a\<Colon>linorder, 'b) RBT_Impl.rbt. is_rbt t}"
    2.12 -  morphisms impl_of RBT
    2.13 -proof -
    2.14 -  have "RBT_Impl.Empty \<in> ?rbt" by simp
    2.15 -  then show ?thesis ..
    2.16 -qed
    2.17 -
    2.18 -lemma rbt_eq_iff:
    2.19 -  "t1 = t2 \<longleftrightarrow> impl_of t1 = impl_of t2"
    2.20 -  by (simp add: impl_of_inject)
    2.21 -
    2.22 -lemma rbt_eqI:
    2.23 -  "impl_of t1 = impl_of t2 \<Longrightarrow> t1 = t2"
    2.24 -  by (simp add: rbt_eq_iff)
    2.25 -
    2.26 -lemma is_rbt_impl_of [simp, intro]:
    2.27 -  "is_rbt (impl_of t)"
    2.28 -  using impl_of [of t] by simp
    2.29 -
    2.30 -lemma RBT_impl_of [simp, code abstype]:
    2.31 -  "RBT (impl_of t) = t"
    2.32 -  by (simp add: impl_of_inverse)
    2.33 -
    2.34 -subsection {* Primitive operations *}
    2.35 -
    2.36 -setup_lifting type_definition_rbt
    2.37 -
    2.38 -lift_definition lookup :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b" is "rbt_lookup" 
    2.39 -by simp
    2.40 -
    2.41 -lift_definition empty :: "('a\<Colon>linorder, 'b) rbt" is RBT_Impl.Empty 
    2.42 -by (simp add: empty_def)
    2.43 -
    2.44 -lift_definition insert :: "'a\<Colon>linorder \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_insert" 
    2.45 -by simp
    2.46 -
    2.47 -lift_definition delete :: "'a\<Colon>linorder \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_delete" 
    2.48 -by simp
    2.49 -
    2.50 -lift_definition entries :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a \<times> 'b) list" is RBT_Impl.entries
    2.51 -by simp
    2.52 -
    2.53 -lift_definition keys :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a list" is RBT_Impl.keys 
    2.54 -by simp
    2.55 -
    2.56 -lift_definition bulkload :: "('a\<Colon>linorder \<times> 'b) list \<Rightarrow> ('a, 'b) rbt" is "rbt_bulkload" 
    2.57 -by simp
    2.58 -
    2.59 -lift_definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is rbt_map_entry 
    2.60 -by simp
    2.61 -
    2.62 -lift_definition map :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is RBT_Impl.map
    2.63 -by simp
    2.64 -
    2.65 -lift_definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c"  is RBT_Impl.fold 
    2.66 -by simp
    2.67 -
    2.68 -lift_definition union :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_union"
    2.69 -by (simp add: rbt_union_is_rbt)
    2.70 -
    2.71 -lift_definition foldi :: "('c \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a :: linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c"
    2.72 -  is RBT_Impl.foldi by simp
    2.73 -
    2.74 -export_code lookup empty insert delete entries keys bulkload map_entry map fold union foldi in SML
    2.75 -
    2.76 -subsection {* Derived operations *}
    2.77 -
    2.78 -definition is_empty :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> bool" where
    2.79 -  [code]: "is_empty t = (case impl_of t of RBT_Impl.Empty \<Rightarrow> True | _ \<Rightarrow> False)"
    2.80 -
    2.81 -
    2.82 -subsection {* Abstract lookup properties *}
    2.83 -
    2.84 -lemma lookup_RBT:
    2.85 -  "is_rbt t \<Longrightarrow> lookup (RBT t) = rbt_lookup t"
    2.86 -  by (simp add: lookup_def RBT_inverse)
    2.87 -
    2.88 -lemma lookup_impl_of:
    2.89 -  "rbt_lookup (impl_of t) = lookup t"
    2.90 -  by transfer (rule refl)
    2.91 -
    2.92 -lemma entries_impl_of:
    2.93 -  "RBT_Impl.entries (impl_of t) = entries t"
    2.94 -  by transfer (rule refl)
    2.95 -
    2.96 -lemma keys_impl_of:
    2.97 -  "RBT_Impl.keys (impl_of t) = keys t"
    2.98 -  by transfer (rule refl)
    2.99 -
   2.100 -lemma lookup_empty [simp]:
   2.101 -  "lookup empty = Map.empty"
   2.102 -  by (simp add: empty_def lookup_RBT fun_eq_iff)
   2.103 -
   2.104 -lemma lookup_insert [simp]:
   2.105 -  "lookup (insert k v t) = (lookup t)(k \<mapsto> v)"
   2.106 -  by transfer (rule rbt_lookup_rbt_insert)
   2.107 -
   2.108 -lemma lookup_delete [simp]:
   2.109 -  "lookup (delete k t) = (lookup t)(k := None)"
   2.110 -  by transfer (simp add: rbt_lookup_rbt_delete restrict_complement_singleton_eq)
   2.111 -
   2.112 -lemma map_of_entries [simp]:
   2.113 -  "map_of (entries t) = lookup t"
   2.114 -  by transfer (simp add: map_of_entries)
   2.115 -
   2.116 -lemma entries_lookup:
   2.117 -  "entries t1 = entries t2 \<longleftrightarrow> lookup t1 = lookup t2"
   2.118 -  by transfer (simp add: entries_rbt_lookup)
   2.119 -
   2.120 -lemma lookup_bulkload [simp]:
   2.121 -  "lookup (bulkload xs) = map_of xs"
   2.122 -  by transfer (rule rbt_lookup_rbt_bulkload)
   2.123 -
   2.124 -lemma lookup_map_entry [simp]:
   2.125 -  "lookup (map_entry k f t) = (lookup t)(k := Option.map f (lookup t k))"
   2.126 -  by transfer (rule rbt_lookup_rbt_map_entry)
   2.127 -
   2.128 -lemma lookup_map [simp]:
   2.129 -  "lookup (map f t) k = Option.map (f k) (lookup t k)"
   2.130 -  by transfer (rule rbt_lookup_map)
   2.131 -
   2.132 -lemma fold_fold:
   2.133 -  "fold f t = List.fold (prod_case f) (entries t)"
   2.134 -  by transfer (rule RBT_Impl.fold_def)
   2.135 -
   2.136 -lemma impl_of_empty:
   2.137 -  "impl_of empty = RBT_Impl.Empty"
   2.138 -  by transfer (rule refl)
   2.139 -
   2.140 -lemma is_empty_empty [simp]:
   2.141 -  "is_empty t \<longleftrightarrow> t = empty"
   2.142 -  unfolding is_empty_def by transfer (simp split: rbt.split)
   2.143 -
   2.144 -lemma RBT_lookup_empty [simp]: (*FIXME*)
   2.145 -  "rbt_lookup t = Map.empty \<longleftrightarrow> t = RBT_Impl.Empty"
   2.146 -  by (cases t) (auto simp add: fun_eq_iff)
   2.147 -
   2.148 -lemma lookup_empty_empty [simp]:
   2.149 -  "lookup t = Map.empty \<longleftrightarrow> t = empty"
   2.150 -  by transfer (rule RBT_lookup_empty)
   2.151 -
   2.152 -lemma sorted_keys [iff]:
   2.153 -  "sorted (keys t)"
   2.154 -  by transfer (simp add: RBT_Impl.keys_def rbt_sorted_entries)
   2.155 -
   2.156 -lemma distinct_keys [iff]:
   2.157 -  "distinct (keys t)"
   2.158 -  by transfer (simp add: RBT_Impl.keys_def distinct_entries)
   2.159 -
   2.160 -lemma finite_dom_lookup [simp, intro!]: "finite (dom (lookup t))"
   2.161 -  by transfer simp
   2.162 -
   2.163 -lemma lookup_union: "lookup (union s t) = lookup s ++ lookup t"
   2.164 -  by transfer (simp add: rbt_lookup_rbt_union)
   2.165 -
   2.166 -lemma lookup_in_tree: "(lookup t k = Some v) = ((k, v) \<in> set (entries t))"
   2.167 -  by transfer (simp add: rbt_lookup_in_tree)
   2.168 -
   2.169 -lemma keys_entries: "(k \<in> set (keys t)) = (\<exists>v. (k, v) \<in> set (entries t))"
   2.170 -  by transfer (simp add: keys_entries)
   2.171 -
   2.172 -lemma fold_def_alt:
   2.173 -  "fold f t = List.fold (prod_case f) (entries t)"
   2.174 -  by transfer (auto simp: RBT_Impl.fold_def)
   2.175 -
   2.176 -lemma distinct_entries: "distinct (List.map fst (entries t))"
   2.177 -  by transfer (simp add: distinct_entries)
   2.178 -
   2.179 -lemma non_empty_keys: "t \<noteq> Lift_RBT.empty \<Longrightarrow> keys t \<noteq> []"
   2.180 -  by transfer (simp add: non_empty_rbt_keys)
   2.181 -
   2.182 -lemma keys_def_alt:
   2.183 -  "keys t = List.map fst (entries t)"
   2.184 -  by transfer (simp add: RBT_Impl.keys_def)
   2.185 +(* Moved to ~~/HOL/Library/RBT" *)
   2.186  
   2.187  end
   2.188 \ No newline at end of file