# HG changeset patch # User kuncar # Date 1343735739 -7200 # Node ID caaa1a02c650e6e12da388eef4d5c7230cc64558 # Parent 877df57629e3c008d280155844ba5f7f7fc5623e use lifting/transfer formalization of RBT from Lift_RBT diff -r 877df57629e3 -r caaa1a02c650 src/HOL/Library/RBT.thy --- a/src/HOL/Library/RBT.thy Tue Jul 31 13:55:39 2012 +0200 +++ b/src/HOL/Library/RBT.thy Tue Jul 31 13:55:39 2012 +0200 @@ -1,9 +1,10 @@ -(* Author: Florian Haftmann, TU Muenchen *) +(* Title: HOL/Library/RBT.thy + Author: Lukas Bulwahn and Ondrej Kuncar +*) -header {* Abstract type of Red-Black Trees *} +header {* Abstract type of RBT trees *} -(*<*) -theory RBT +theory RBT imports Main RBT_Impl begin @@ -11,8 +12,9 @@ typedef (open) ('a, 'b) rbt = "{t :: ('a\linorder, 'b) RBT_Impl.rbt. is_rbt t}" morphisms impl_of RBT -proof - show "RBT_Impl.Empty \ {t. is_rbt t}" by simp +proof - + have "RBT_Impl.Empty \ ?rbt" by simp + then show ?thesis .. qed lemma rbt_eq_iff: @@ -31,63 +33,45 @@ "RBT (impl_of t) = t" by (simp add: impl_of_inverse) - subsection {* Primitive operations *} -definition lookup :: "('a\linorder, 'b) rbt \ 'a \ 'b" where - [code]: "lookup t = rbt_lookup (impl_of t)" - -definition empty :: "('a\linorder, 'b) rbt" where - "empty = RBT RBT_Impl.Empty" +setup_lifting type_definition_rbt -lemma impl_of_empty [code abstract]: - "impl_of empty = RBT_Impl.Empty" - by (simp add: empty_def RBT_inverse) +lift_definition lookup :: "('a\linorder, 'b) rbt \ 'a \ 'b" is "rbt_lookup" +by simp -definition insert :: "'a\linorder \ 'b \ ('a, 'b) rbt \ ('a, 'b) rbt" where - "insert k v t = RBT (rbt_insert k v (impl_of t))" +lift_definition empty :: "('a\linorder, 'b) rbt" is RBT_Impl.Empty +by (simp add: empty_def) -lemma impl_of_insert [code abstract]: - "impl_of (insert k v t) = rbt_insert k v (impl_of t)" - by (simp add: insert_def RBT_inverse) - -definition delete :: "'a\linorder \ ('a, 'b) rbt \ ('a, 'b) rbt" where - "delete k t = RBT (rbt_delete k (impl_of t))" +lift_definition insert :: "'a\linorder \ 'b \ ('a, 'b) rbt \ ('a, 'b) rbt" is "rbt_insert" +by simp -lemma impl_of_delete [code abstract]: - "impl_of (delete k t) = rbt_delete k (impl_of t)" - by (simp add: delete_def RBT_inverse) +lift_definition delete :: "'a\linorder \ ('a, 'b) rbt \ ('a, 'b) rbt" is "rbt_delete" +by simp -definition entries :: "('a\linorder, 'b) rbt \ ('a \ 'b) list" where - [code]: "entries t = RBT_Impl.entries (impl_of t)" +lift_definition entries :: "('a\linorder, 'b) rbt \ ('a \ 'b) list" is RBT_Impl.entries +by simp -definition keys :: "('a\linorder, 'b) rbt \ 'a list" where - [code]: "keys t = RBT_Impl.keys (impl_of t)" - -definition bulkload :: "('a\linorder \ 'b) list \ ('a, 'b) rbt" where - "bulkload xs = RBT (rbt_bulkload xs)" +lift_definition keys :: "('a\linorder, 'b) rbt \ 'a list" is RBT_Impl.keys +by simp -lemma impl_of_bulkload [code abstract]: - "impl_of (bulkload xs) = rbt_bulkload xs" - by (simp add: bulkload_def RBT_inverse) +lift_definition bulkload :: "('a\linorder \ 'b) list \ ('a, 'b) rbt" is "rbt_bulkload" +by simp -definition map_entry :: "'a \ ('b \ 'b) \ ('a\linorder, 'b) rbt \ ('a, 'b) rbt" where - "map_entry k f t = RBT (rbt_map_entry k f (impl_of t))" +lift_definition map_entry :: "'a \ ('b \ 'b) \ ('a\linorder, 'b) rbt \ ('a, 'b) rbt" is rbt_map_entry +by simp -lemma impl_of_map_entry [code abstract]: - "impl_of (map_entry k f t) = rbt_map_entry k f (impl_of t)" - by (simp add: map_entry_def RBT_inverse) +lift_definition map :: "('a \ 'b \ 'b) \ ('a\linorder, 'b) rbt \ ('a, 'b) rbt" is RBT_Impl.map +by simp -definition map :: "('a \ 'b \ 'b) \ ('a\linorder, 'b) rbt \ ('a, 'b) rbt" where - "map f t = RBT (RBT_Impl.map f (impl_of t))" +lift_definition fold :: "('a \ 'b \ 'c \ 'c) \ ('a\linorder, 'b) rbt \ 'c \ 'c" is RBT_Impl.fold +by simp -lemma impl_of_map [code abstract]: - "impl_of (map f t) = RBT_Impl.map f (impl_of t)" - by (simp add: map_def RBT_inverse) +lift_definition union :: "('a\linorder, 'b) rbt \ ('a, 'b) rbt \ ('a, 'b) rbt" is "rbt_union" +by (simp add: rbt_union_is_rbt) -definition fold :: "('a \ 'b \ 'c \ 'c) \ ('a\linorder, 'b) rbt \ 'c \ 'c" where - [code]: "fold f t = RBT_Impl.fold f (impl_of t)" - +lift_definition foldi :: "('c \ bool) \ ('a \ 'b \ 'c \ 'c) \ ('a :: linorder, 'b) rbt \ 'c \ 'c" + is RBT_Impl.foldi by simp subsection {* Derived operations *} @@ -103,15 +87,15 @@ lemma lookup_impl_of: "rbt_lookup (impl_of t) = lookup t" - by (simp add: lookup_def) + by transfer (rule refl) lemma entries_impl_of: "RBT_Impl.entries (impl_of t) = entries t" - by (simp add: entries_def) + by transfer (rule refl) lemma keys_impl_of: "RBT_Impl.keys (impl_of t) = keys t" - by (simp add: keys_def) + by transfer (rule refl) lemma lookup_empty [simp]: "lookup empty = Map.empty" @@ -119,39 +103,43 @@ lemma lookup_insert [simp]: "lookup (insert k v t) = (lookup t)(k \ v)" - by (simp add: insert_def lookup_RBT rbt_lookup_rbt_insert lookup_impl_of) + by transfer (rule rbt_lookup_rbt_insert) lemma lookup_delete [simp]: "lookup (delete k t) = (lookup t)(k := None)" - by (simp add: delete_def lookup_RBT rbt_lookup_rbt_delete lookup_impl_of restrict_complement_singleton_eq) + by transfer (simp add: rbt_lookup_rbt_delete restrict_complement_singleton_eq) lemma map_of_entries [simp]: "map_of (entries t) = lookup t" - by (simp add: entries_def map_of_entries lookup_impl_of) + by transfer (simp add: map_of_entries) lemma entries_lookup: "entries t1 = entries t2 \ lookup t1 = lookup t2" - by (simp add: entries_def lookup_def entries_rbt_lookup) + by transfer (simp add: entries_rbt_lookup) lemma lookup_bulkload [simp]: "lookup (bulkload xs) = map_of xs" - by (simp add: bulkload_def lookup_RBT rbt_lookup_rbt_bulkload) + by transfer (rule rbt_lookup_rbt_bulkload) lemma lookup_map_entry [simp]: "lookup (map_entry k f t) = (lookup t)(k := Option.map f (lookup t k))" - by (simp add: map_entry_def lookup_RBT rbt_lookup_rbt_map_entry lookup_impl_of) + by transfer (rule rbt_lookup_rbt_map_entry) lemma lookup_map [simp]: "lookup (map f t) k = Option.map (f k) (lookup t k)" - by (simp add: map_def lookup_RBT rbt_lookup_map lookup_impl_of) + by transfer (rule rbt_lookup_map) lemma fold_fold: "fold f t = List.fold (prod_case f) (entries t)" - by (simp add: fold_def fun_eq_iff RBT_Impl.fold_def entries_impl_of) + by transfer (rule RBT_Impl.fold_def) + +lemma impl_of_empty: + "impl_of empty = RBT_Impl.Empty" + by transfer (rule refl) lemma is_empty_empty [simp]: "is_empty t \ t = empty" - by (simp add: rbt_eq_iff is_empty_def impl_of_empty split: rbt.split) + unfolding is_empty_def by transfer (simp split: rbt.split) lemma RBT_lookup_empty [simp]: (*FIXME*) "rbt_lookup t = Map.empty \ t = RBT_Impl.Empty" @@ -159,15 +147,41 @@ lemma lookup_empty_empty [simp]: "lookup t = Map.empty \ t = empty" - by (cases t) (simp add: empty_def lookup_def RBT_inject RBT_inverse) + by transfer (rule RBT_lookup_empty) lemma sorted_keys [iff]: "sorted (keys t)" - by (simp add: keys_def RBT_Impl.keys_def rbt_sorted_entries) + by transfer (simp add: RBT_Impl.keys_def rbt_sorted_entries) lemma distinct_keys [iff]: "distinct (keys t)" - by (simp add: keys_def RBT_Impl.keys_def distinct_entries) + by transfer (simp add: RBT_Impl.keys_def distinct_entries) + +lemma finite_dom_lookup [simp, intro!]: "finite (dom (lookup t))" + by transfer simp + +lemma lookup_union: "lookup (union s t) = lookup s ++ lookup t" + by transfer (simp add: rbt_lookup_rbt_union) + +lemma lookup_in_tree: "(lookup t k = Some v) = ((k, v) \ set (entries t))" + by transfer (simp add: rbt_lookup_in_tree) + +lemma keys_entries: "(k \ set (keys t)) = (\v. (k, v) \ set (entries t))" + by transfer (simp add: keys_entries) + +lemma fold_def_alt: + "fold f t = List.fold (prod_case f) (entries t)" + by transfer (auto simp: RBT_Impl.fold_def) + +lemma distinct_entries: "distinct (List.map fst (entries t))" + by transfer (simp add: distinct_entries) + +lemma non_empty_keys: "t \ empty \ keys t \ []" + by transfer (simp add: non_empty_rbt_keys) + +lemma keys_def_alt: + "keys t = List.map fst (entries t)" + by transfer (simp add: RBT_Impl.keys_def) subsection {* Quickcheck generators *} diff -r 877df57629e3 -r caaa1a02c650 src/HOL/Quotient_Examples/Lift_RBT.thy --- a/src/HOL/Quotient_Examples/Lift_RBT.thy Tue Jul 31 13:55:39 2012 +0200 +++ b/src/HOL/Quotient_Examples/Lift_RBT.thy Tue Jul 31 13:55:39 2012 +0200 @@ -8,183 +8,6 @@ imports Main "~~/src/HOL/Library/RBT_Impl" begin -(* TODO: Replace the ancient Library/RBT theory by this example of the lifting and transfer mechanism. *) - -subsection {* Type definition *} - -typedef (open) ('a, 'b) rbt = "{t :: ('a\linorder, 'b) RBT_Impl.rbt. is_rbt t}" - morphisms impl_of RBT -proof - - have "RBT_Impl.Empty \ ?rbt" by simp - then show ?thesis .. -qed - -lemma rbt_eq_iff: - "t1 = t2 \ impl_of t1 = impl_of t2" - by (simp add: impl_of_inject) - -lemma rbt_eqI: - "impl_of t1 = impl_of t2 \ t1 = t2" - by (simp add: rbt_eq_iff) - -lemma is_rbt_impl_of [simp, intro]: - "is_rbt (impl_of t)" - using impl_of [of t] by simp - -lemma RBT_impl_of [simp, code abstype]: - "RBT (impl_of t) = t" - by (simp add: impl_of_inverse) - -subsection {* Primitive operations *} - -setup_lifting type_definition_rbt - -lift_definition lookup :: "('a\linorder, 'b) rbt \ 'a \ 'b" is "rbt_lookup" -by simp - -lift_definition empty :: "('a\linorder, 'b) rbt" is RBT_Impl.Empty -by (simp add: empty_def) - -lift_definition insert :: "'a\linorder \ 'b \ ('a, 'b) rbt \ ('a, 'b) rbt" is "rbt_insert" -by simp - -lift_definition delete :: "'a\linorder \ ('a, 'b) rbt \ ('a, 'b) rbt" is "rbt_delete" -by simp - -lift_definition entries :: "('a\linorder, 'b) rbt \ ('a \ 'b) list" is RBT_Impl.entries -by simp - -lift_definition keys :: "('a\linorder, 'b) rbt \ 'a list" is RBT_Impl.keys -by simp - -lift_definition bulkload :: "('a\linorder \ 'b) list \ ('a, 'b) rbt" is "rbt_bulkload" -by simp - -lift_definition map_entry :: "'a \ ('b \ 'b) \ ('a\linorder, 'b) rbt \ ('a, 'b) rbt" is rbt_map_entry -by simp - -lift_definition map :: "('a \ 'b \ 'b) \ ('a\linorder, 'b) rbt \ ('a, 'b) rbt" is RBT_Impl.map -by simp - -lift_definition fold :: "('a \ 'b \ 'c \ 'c) \ ('a\linorder, 'b) rbt \ 'c \ 'c" is RBT_Impl.fold -by simp - -lift_definition union :: "('a\linorder, 'b) rbt \ ('a, 'b) rbt \ ('a, 'b) rbt" is "rbt_union" -by (simp add: rbt_union_is_rbt) - -lift_definition foldi :: "('c \ bool) \ ('a \ 'b \ 'c \ 'c) \ ('a :: linorder, 'b) rbt \ 'c \ 'c" - is RBT_Impl.foldi by simp - -export_code lookup empty insert delete entries keys bulkload map_entry map fold union foldi in SML - -subsection {* Derived operations *} - -definition is_empty :: "('a\linorder, 'b) rbt \ bool" where - [code]: "is_empty t = (case impl_of t of RBT_Impl.Empty \ True | _ \ False)" - - -subsection {* Abstract lookup properties *} - -lemma lookup_RBT: - "is_rbt t \ lookup (RBT t) = rbt_lookup t" - by (simp add: lookup_def RBT_inverse) - -lemma lookup_impl_of: - "rbt_lookup (impl_of t) = lookup t" - by transfer (rule refl) - -lemma entries_impl_of: - "RBT_Impl.entries (impl_of t) = entries t" - by transfer (rule refl) - -lemma keys_impl_of: - "RBT_Impl.keys (impl_of t) = keys t" - by transfer (rule refl) - -lemma lookup_empty [simp]: - "lookup empty = Map.empty" - by (simp add: empty_def lookup_RBT fun_eq_iff) - -lemma lookup_insert [simp]: - "lookup (insert k v t) = (lookup t)(k \ v)" - by transfer (rule rbt_lookup_rbt_insert) - -lemma lookup_delete [simp]: - "lookup (delete k t) = (lookup t)(k := None)" - by transfer (simp add: rbt_lookup_rbt_delete restrict_complement_singleton_eq) - -lemma map_of_entries [simp]: - "map_of (entries t) = lookup t" - by transfer (simp add: map_of_entries) - -lemma entries_lookup: - "entries t1 = entries t2 \ lookup t1 = lookup t2" - by transfer (simp add: entries_rbt_lookup) - -lemma lookup_bulkload [simp]: - "lookup (bulkload xs) = map_of xs" - by transfer (rule rbt_lookup_rbt_bulkload) - -lemma lookup_map_entry [simp]: - "lookup (map_entry k f t) = (lookup t)(k := Option.map f (lookup t k))" - by transfer (rule rbt_lookup_rbt_map_entry) - -lemma lookup_map [simp]: - "lookup (map f t) k = Option.map (f k) (lookup t k)" - by transfer (rule rbt_lookup_map) - -lemma fold_fold: - "fold f t = List.fold (prod_case f) (entries t)" - by transfer (rule RBT_Impl.fold_def) - -lemma impl_of_empty: - "impl_of empty = RBT_Impl.Empty" - by transfer (rule refl) - -lemma is_empty_empty [simp]: - "is_empty t \ t = empty" - unfolding is_empty_def by transfer (simp split: rbt.split) - -lemma RBT_lookup_empty [simp]: (*FIXME*) - "rbt_lookup t = Map.empty \ t = RBT_Impl.Empty" - by (cases t) (auto simp add: fun_eq_iff) - -lemma lookup_empty_empty [simp]: - "lookup t = Map.empty \ t = empty" - by transfer (rule RBT_lookup_empty) - -lemma sorted_keys [iff]: - "sorted (keys t)" - by transfer (simp add: RBT_Impl.keys_def rbt_sorted_entries) - -lemma distinct_keys [iff]: - "distinct (keys t)" - by transfer (simp add: RBT_Impl.keys_def distinct_entries) - -lemma finite_dom_lookup [simp, intro!]: "finite (dom (lookup t))" - by transfer simp - -lemma lookup_union: "lookup (union s t) = lookup s ++ lookup t" - by transfer (simp add: rbt_lookup_rbt_union) - -lemma lookup_in_tree: "(lookup t k = Some v) = ((k, v) \ set (entries t))" - by transfer (simp add: rbt_lookup_in_tree) - -lemma keys_entries: "(k \ set (keys t)) = (\v. (k, v) \ set (entries t))" - by transfer (simp add: keys_entries) - -lemma fold_def_alt: - "fold f t = List.fold (prod_case f) (entries t)" - by transfer (auto simp: RBT_Impl.fold_def) - -lemma distinct_entries: "distinct (List.map fst (entries t))" - by transfer (simp add: distinct_entries) - -lemma non_empty_keys: "t \ Lift_RBT.empty \ keys t \ []" - by transfer (simp add: non_empty_rbt_keys) - -lemma keys_def_alt: - "keys t = List.map fst (entries t)" - by transfer (simp add: RBT_Impl.keys_def) +(* Moved to ~~/HOL/Library/RBT" *) end \ No newline at end of file