src/HOL/Library/RBT.thy
changeset 48622 caaa1a02c650
parent 47450 2ada2be850cb
child 49834 b27bbb021df1
equal deleted inserted replaced
48621:877df57629e3 48622:caaa1a02c650
     1 (* Author: Florian Haftmann, TU Muenchen *)
     1 (*  Title:      HOL/Library/RBT.thy
       
     2     Author:     Lukas Bulwahn and Ondrej Kuncar
       
     3 *)
     2 
     4 
     3 header {* Abstract type of Red-Black Trees *}
     5 header {* Abstract type of RBT trees *}
     4 
     6 
     5 (*<*)
     7 theory RBT 
     6 theory RBT
       
     7 imports Main RBT_Impl
     8 imports Main RBT_Impl
     8 begin
     9 begin
     9 
    10 
    10 subsection {* Type definition *}
    11 subsection {* Type definition *}
    11 
    12 
    12 typedef (open) ('a, 'b) rbt = "{t :: ('a\<Colon>linorder, 'b) RBT_Impl.rbt. is_rbt t}"
    13 typedef (open) ('a, 'b) rbt = "{t :: ('a\<Colon>linorder, 'b) RBT_Impl.rbt. is_rbt t}"
    13   morphisms impl_of RBT
    14   morphisms impl_of RBT
    14 proof
    15 proof -
    15   show "RBT_Impl.Empty \<in> {t. is_rbt t}" by simp
    16   have "RBT_Impl.Empty \<in> ?rbt" by simp
       
    17   then show ?thesis ..
    16 qed
    18 qed
    17 
    19 
    18 lemma rbt_eq_iff:
    20 lemma rbt_eq_iff:
    19   "t1 = t2 \<longleftrightarrow> impl_of t1 = impl_of t2"
    21   "t1 = t2 \<longleftrightarrow> impl_of t1 = impl_of t2"
    20   by (simp add: impl_of_inject)
    22   by (simp add: impl_of_inject)
    29 
    31 
    30 lemma RBT_impl_of [simp, code abstype]:
    32 lemma RBT_impl_of [simp, code abstype]:
    31   "RBT (impl_of t) = t"
    33   "RBT (impl_of t) = t"
    32   by (simp add: impl_of_inverse)
    34   by (simp add: impl_of_inverse)
    33 
    35 
    34 
       
    35 subsection {* Primitive operations *}
    36 subsection {* Primitive operations *}
    36 
    37 
    37 definition lookup :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b" where
    38 setup_lifting type_definition_rbt
    38   [code]: "lookup t = rbt_lookup (impl_of t)"
       
    39 
    39 
    40 definition empty :: "('a\<Colon>linorder, 'b) rbt" where
    40 lift_definition lookup :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b" is "rbt_lookup" 
    41   "empty = RBT RBT_Impl.Empty"
    41 by simp
    42 
    42 
    43 lemma impl_of_empty [code abstract]:
    43 lift_definition empty :: "('a\<Colon>linorder, 'b) rbt" is RBT_Impl.Empty 
    44   "impl_of empty = RBT_Impl.Empty"
    44 by (simp add: empty_def)
    45   by (simp add: empty_def RBT_inverse)
       
    46 
    45 
    47 definition insert :: "'a\<Colon>linorder \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
    46 lift_definition insert :: "'a\<Colon>linorder \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_insert" 
    48   "insert k v t = RBT (rbt_insert k v (impl_of t))"
    47 by simp
    49 
    48 
    50 lemma impl_of_insert [code abstract]:
    49 lift_definition delete :: "'a\<Colon>linorder \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_delete" 
    51   "impl_of (insert k v t) = rbt_insert k v (impl_of t)"
    50 by simp
    52   by (simp add: insert_def RBT_inverse)
       
    53 
    51 
    54 definition delete :: "'a\<Colon>linorder \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
    52 lift_definition entries :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a \<times> 'b) list" is RBT_Impl.entries
    55   "delete k t = RBT (rbt_delete k (impl_of t))"
    53 by simp
    56 
    54 
    57 lemma impl_of_delete [code abstract]:
    55 lift_definition keys :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a list" is RBT_Impl.keys 
    58   "impl_of (delete k t) = rbt_delete k (impl_of t)"
    56 by simp
    59   by (simp add: delete_def RBT_inverse)
       
    60 
    57 
    61 definition entries :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a \<times> 'b) list" where
    58 lift_definition bulkload :: "('a\<Colon>linorder \<times> 'b) list \<Rightarrow> ('a, 'b) rbt" is "rbt_bulkload" 
    62   [code]: "entries t = RBT_Impl.entries (impl_of t)"
    59 by simp
    63 
    60 
    64 definition keys :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a list" where
    61 lift_definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is rbt_map_entry 
    65   [code]: "keys t = RBT_Impl.keys (impl_of t)"
    62 by simp
    66 
    63 
    67 definition bulkload :: "('a\<Colon>linorder \<times> 'b) list \<Rightarrow> ('a, 'b) rbt" where
    64 lift_definition map :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is RBT_Impl.map
    68   "bulkload xs = RBT (rbt_bulkload xs)"
    65 by simp
    69 
    66 
    70 lemma impl_of_bulkload [code abstract]:
    67 lift_definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c"  is RBT_Impl.fold 
    71   "impl_of (bulkload xs) = rbt_bulkload xs"
    68 by simp
    72   by (simp add: bulkload_def RBT_inverse)
       
    73 
    69 
    74 definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
    70 lift_definition union :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_union"
    75   "map_entry k f t = RBT (rbt_map_entry k f (impl_of t))"
    71 by (simp add: rbt_union_is_rbt)
    76 
    72 
    77 lemma impl_of_map_entry [code abstract]:
    73 lift_definition foldi :: "('c \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a :: linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c"
    78   "impl_of (map_entry k f t) = rbt_map_entry k f (impl_of t)"
    74   is RBT_Impl.foldi by simp
    79   by (simp add: map_entry_def RBT_inverse)
       
    80 
       
    81 definition map :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
       
    82   "map f t = RBT (RBT_Impl.map f (impl_of t))"
       
    83 
       
    84 lemma impl_of_map [code abstract]:
       
    85   "impl_of (map f t) = RBT_Impl.map f (impl_of t)"
       
    86   by (simp add: map_def RBT_inverse)
       
    87 
       
    88 definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c" where
       
    89   [code]: "fold f t = RBT_Impl.fold f (impl_of t)"
       
    90 
       
    91 
    75 
    92 subsection {* Derived operations *}
    76 subsection {* Derived operations *}
    93 
    77 
    94 definition is_empty :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> bool" where
    78 definition is_empty :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> bool" where
    95   [code]: "is_empty t = (case impl_of t of RBT_Impl.Empty \<Rightarrow> True | _ \<Rightarrow> False)"
    79   [code]: "is_empty t = (case impl_of t of RBT_Impl.Empty \<Rightarrow> True | _ \<Rightarrow> False)"
   101   "is_rbt t \<Longrightarrow> lookup (RBT t) = rbt_lookup t"
    85   "is_rbt t \<Longrightarrow> lookup (RBT t) = rbt_lookup t"
   102   by (simp add: lookup_def RBT_inverse)
    86   by (simp add: lookup_def RBT_inverse)
   103 
    87 
   104 lemma lookup_impl_of:
    88 lemma lookup_impl_of:
   105   "rbt_lookup (impl_of t) = lookup t"
    89   "rbt_lookup (impl_of t) = lookup t"
   106   by (simp add: lookup_def)
    90   by transfer (rule refl)
   107 
    91 
   108 lemma entries_impl_of:
    92 lemma entries_impl_of:
   109   "RBT_Impl.entries (impl_of t) = entries t"
    93   "RBT_Impl.entries (impl_of t) = entries t"
   110   by (simp add: entries_def)
    94   by transfer (rule refl)
   111 
    95 
   112 lemma keys_impl_of:
    96 lemma keys_impl_of:
   113   "RBT_Impl.keys (impl_of t) = keys t"
    97   "RBT_Impl.keys (impl_of t) = keys t"
   114   by (simp add: keys_def)
    98   by transfer (rule refl)
   115 
    99 
   116 lemma lookup_empty [simp]:
   100 lemma lookup_empty [simp]:
   117   "lookup empty = Map.empty"
   101   "lookup empty = Map.empty"
   118   by (simp add: empty_def lookup_RBT fun_eq_iff)
   102   by (simp add: empty_def lookup_RBT fun_eq_iff)
   119 
   103 
   120 lemma lookup_insert [simp]:
   104 lemma lookup_insert [simp]:
   121   "lookup (insert k v t) = (lookup t)(k \<mapsto> v)"
   105   "lookup (insert k v t) = (lookup t)(k \<mapsto> v)"
   122   by (simp add: insert_def lookup_RBT rbt_lookup_rbt_insert lookup_impl_of)
   106   by transfer (rule rbt_lookup_rbt_insert)
   123 
   107 
   124 lemma lookup_delete [simp]:
   108 lemma lookup_delete [simp]:
   125   "lookup (delete k t) = (lookup t)(k := None)"
   109   "lookup (delete k t) = (lookup t)(k := None)"
   126   by (simp add: delete_def lookup_RBT rbt_lookup_rbt_delete lookup_impl_of restrict_complement_singleton_eq)
   110   by transfer (simp add: rbt_lookup_rbt_delete restrict_complement_singleton_eq)
   127 
   111 
   128 lemma map_of_entries [simp]:
   112 lemma map_of_entries [simp]:
   129   "map_of (entries t) = lookup t"
   113   "map_of (entries t) = lookup t"
   130   by (simp add: entries_def map_of_entries lookup_impl_of)
   114   by transfer (simp add: map_of_entries)
   131 
   115 
   132 lemma entries_lookup:
   116 lemma entries_lookup:
   133   "entries t1 = entries t2 \<longleftrightarrow> lookup t1 = lookup t2"
   117   "entries t1 = entries t2 \<longleftrightarrow> lookup t1 = lookup t2"
   134   by (simp add: entries_def lookup_def entries_rbt_lookup)
   118   by transfer (simp add: entries_rbt_lookup)
   135 
   119 
   136 lemma lookup_bulkload [simp]:
   120 lemma lookup_bulkload [simp]:
   137   "lookup (bulkload xs) = map_of xs"
   121   "lookup (bulkload xs) = map_of xs"
   138   by (simp add: bulkload_def lookup_RBT rbt_lookup_rbt_bulkload)
   122   by transfer (rule rbt_lookup_rbt_bulkload)
   139 
   123 
   140 lemma lookup_map_entry [simp]:
   124 lemma lookup_map_entry [simp]:
   141   "lookup (map_entry k f t) = (lookup t)(k := Option.map f (lookup t k))"
   125   "lookup (map_entry k f t) = (lookup t)(k := Option.map f (lookup t k))"
   142   by (simp add: map_entry_def lookup_RBT rbt_lookup_rbt_map_entry lookup_impl_of)
   126   by transfer (rule rbt_lookup_rbt_map_entry)
   143 
   127 
   144 lemma lookup_map [simp]:
   128 lemma lookup_map [simp]:
   145   "lookup (map f t) k = Option.map (f k) (lookup t k)"
   129   "lookup (map f t) k = Option.map (f k) (lookup t k)"
   146   by (simp add: map_def lookup_RBT rbt_lookup_map lookup_impl_of)
   130   by transfer (rule rbt_lookup_map)
   147 
   131 
   148 lemma fold_fold:
   132 lemma fold_fold:
   149   "fold f t = List.fold (prod_case f) (entries t)"
   133   "fold f t = List.fold (prod_case f) (entries t)"
   150   by (simp add: fold_def fun_eq_iff RBT_Impl.fold_def entries_impl_of)
   134   by transfer (rule RBT_Impl.fold_def)
       
   135 
       
   136 lemma impl_of_empty:
       
   137   "impl_of empty = RBT_Impl.Empty"
       
   138   by transfer (rule refl)
   151 
   139 
   152 lemma is_empty_empty [simp]:
   140 lemma is_empty_empty [simp]:
   153   "is_empty t \<longleftrightarrow> t = empty"
   141   "is_empty t \<longleftrightarrow> t = empty"
   154   by (simp add: rbt_eq_iff is_empty_def impl_of_empty split: rbt.split)
   142   unfolding is_empty_def by transfer (simp split: rbt.split)
   155 
   143 
   156 lemma RBT_lookup_empty [simp]: (*FIXME*)
   144 lemma RBT_lookup_empty [simp]: (*FIXME*)
   157   "rbt_lookup t = Map.empty \<longleftrightarrow> t = RBT_Impl.Empty"
   145   "rbt_lookup t = Map.empty \<longleftrightarrow> t = RBT_Impl.Empty"
   158   by (cases t) (auto simp add: fun_eq_iff)
   146   by (cases t) (auto simp add: fun_eq_iff)
   159 
   147 
   160 lemma lookup_empty_empty [simp]:
   148 lemma lookup_empty_empty [simp]:
   161   "lookup t = Map.empty \<longleftrightarrow> t = empty"
   149   "lookup t = Map.empty \<longleftrightarrow> t = empty"
   162   by (cases t) (simp add: empty_def lookup_def RBT_inject RBT_inverse)
   150   by transfer (rule RBT_lookup_empty)
   163 
   151 
   164 lemma sorted_keys [iff]:
   152 lemma sorted_keys [iff]:
   165   "sorted (keys t)"
   153   "sorted (keys t)"
   166   by (simp add: keys_def RBT_Impl.keys_def rbt_sorted_entries)
   154   by transfer (simp add: RBT_Impl.keys_def rbt_sorted_entries)
   167 
   155 
   168 lemma distinct_keys [iff]:
   156 lemma distinct_keys [iff]:
   169   "distinct (keys t)"
   157   "distinct (keys t)"
   170   by (simp add: keys_def RBT_Impl.keys_def distinct_entries)
   158   by transfer (simp add: RBT_Impl.keys_def distinct_entries)
       
   159 
       
   160 lemma finite_dom_lookup [simp, intro!]: "finite (dom (lookup t))"
       
   161   by transfer simp
       
   162 
       
   163 lemma lookup_union: "lookup (union s t) = lookup s ++ lookup t"
       
   164   by transfer (simp add: rbt_lookup_rbt_union)
       
   165 
       
   166 lemma lookup_in_tree: "(lookup t k = Some v) = ((k, v) \<in> set (entries t))"
       
   167   by transfer (simp add: rbt_lookup_in_tree)
       
   168 
       
   169 lemma keys_entries: "(k \<in> set (keys t)) = (\<exists>v. (k, v) \<in> set (entries t))"
       
   170   by transfer (simp add: keys_entries)
       
   171 
       
   172 lemma fold_def_alt:
       
   173   "fold f t = List.fold (prod_case f) (entries t)"
       
   174   by transfer (auto simp: RBT_Impl.fold_def)
       
   175 
       
   176 lemma distinct_entries: "distinct (List.map fst (entries t))"
       
   177   by transfer (simp add: distinct_entries)
       
   178 
       
   179 lemma non_empty_keys: "t \<noteq> empty \<Longrightarrow> keys t \<noteq> []"
       
   180   by transfer (simp add: non_empty_rbt_keys)
       
   181 
       
   182 lemma keys_def_alt:
       
   183   "keys t = List.map fst (entries t)"
       
   184   by transfer (simp add: RBT_Impl.keys_def)
   171 
   185 
   172 subsection {* Quickcheck generators *}
   186 subsection {* Quickcheck generators *}
   173 
   187 
   174 quickcheck_generator rbt predicate: is_rbt constructors: empty, insert
   188 quickcheck_generator rbt predicate: is_rbt constructors: empty, insert
   175 
   189