# HG changeset patch # User nipkow # Date 1322219919 -3600 # Node ID e5fd20f232412bc3ff8de7ca252d8c81edd41cda # Parent ef08425dd2d5c72fae6eddb72f2deb0ed48848c7# Parent 5f85efcb50c19b77925f07d36b5784dd02beb91c merged diff -r 5f85efcb50c1 -r e5fd20f23241 src/HOL/Quotient_Examples/Lift_RBT.thy --- a/src/HOL/Quotient_Examples/Lift_RBT.thy Fri Nov 25 12:18:33 2011 +0100 +++ b/src/HOL/Quotient_Examples/Lift_RBT.thy Fri Nov 25 12:18:39 2011 +0100 @@ -53,14 +53,21 @@ declare lookup_def[unfolded map_fun_def comp_def id_def, code] -(* FIXME: some bug in quotient?*) -(* +(* FIXME: quotient_definition at the moment requires that types variables match exactly, +i.e., sort constraints must be annotated to the constant being lifted. +But, it should be possible to infer the necessary sort constraints, making the explicit +sort constraints unnecessary. + +Hence this unannotated quotient_definition fails: + quotient_definition empty where "empty :: ('a\linorder, 'b) rbt" is "RBT_Impl.Empty" + +Similar issue for quotient_definition for entries, keys, map, and fold. *) -definition empty :: "('a\linorder, 'b) rbt" where - "empty = RBT RBT_Impl.Empty" +quotient_definition empty where "empty :: ('a\linorder, 'b) rbt" +is "(RBT_Impl.Empty :: ('a\linorder, 'b) RBT_Impl.rbt)" lemma impl_of_empty [code abstract]: "impl_of empty = RBT_Impl.Empty" @@ -80,19 +87,16 @@ "impl_of (delete k t) = RBT_Impl.delete k (impl_of t)" by (simp add: delete_def RBT_inverse) -(* FIXME: bug in quotient? *) -(* -quotient_definition entries where "entries :: ('a\linorder, 'b) rbt \ ('a \ 'b) list" -is "RBT_Impl.entries" -*) -definition entries :: "('a\linorder, 'b) rbt \ ('a \ 'b) list" where - [code]: "entries t = RBT_Impl.entries (impl_of t)" +(* FIXME: unnecessary type annotations *) +quotient_definition "entries :: ('a\linorder, 'b) rbt \ ('a \ 'b) list" +is "RBT_Impl.entries :: ('a\linorder, 'b) RBT_Impl.rbt \ ('a \ 'b) list" -(* FIXME: bug in quotient? *) -(* -quotient_definition keys where "keys :: ('a\linorder, 'b) rbt \ 'a list" -is "RBT_Impl.keys" -*) +lemma [code]: "entries t = RBT_Impl.entries (impl_of t)" +unfolding entries_def map_fun_def comp_def id_def .. + +(* FIXME: unnecessary type annotations *) +quotient_definition "keys :: ('a\linorder, 'b) rbt \ 'a list" +is "RBT_Impl.keys :: ('a\linorder, 'b) RBT_Impl.rbt \ 'a list" quotient_definition "bulkload :: ('a\linorder \ 'b) list \ ('a, 'b) rbt" is "RBT_Impl.bulkload" @@ -108,28 +112,103 @@ "impl_of (map_entry k f t) = RBT_Impl.map_entry k f (impl_of t)" by (simp add: map_entry_def RBT_inverse) -(* Another bug: map is supposed to be a new definition, not using the old one. -quotient_definition "map :: ('a \ 'b \ 'b) \ ('a\linorder, 'b) rbt \ ('a, 'b) rbt" -is "RBT_Impl.map" -*) -(* But this works, and then shows the other bug... *) -(* +(* FIXME: unnecesary type annotations *) quotient_definition map where "map :: ('a \ 'b \ 'b) \ ('a\linorder, 'b) rbt \ ('a, 'b) rbt" -is "RBT_Impl.map" -*) -definition map :: "('a \ 'b \ 'b) \ ('a\linorder, 'b) rbt \ ('a, 'b) rbt" where - "map f t = RBT (RBT_Impl.map f (impl_of t))" +is "RBT_Impl.map :: ('a \ 'b \ 'b) \ ('a\linorder, 'b) RBT_Impl.rbt \ ('a, 'b) RBT_Impl.rbt" 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) -(* FIXME: bug? -quotient_definition fold where "fold :: ('a \ 'b \ 'c \ 'c) \ ('a\linorder, 'b) rbt \ 'c \ 'c" is "RBT_Impl.fold" -*) -(*FIXME: definition of fold should have the code attribute. *) + +(* FIXME: unnecessary type annotations *) +quotient_definition fold where "fold :: ('a \ 'b \ 'c \ 'c) \ ('a\linorder, 'b) rbt \ 'c \ 'c" is "RBT_Impl.fold :: ('a \ 'b \ 'c \ 'c) \ ('a\linorder, 'b) RBT_Impl.rbt \ 'c \ 'c" + +lemma [code]: "fold f t = RBT_Impl.fold f (impl_of t)" +unfolding fold_def map_fun_def comp_def id_def .. + + +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 *} + +(* TODO: obtain the following lemmas by lifting existing theorems. *) + +lemma lookup_RBT: + "is_rbt t \ lookup (RBT t) = RBT_Impl.lookup t" + by (simp add: lookup_def RBT_inverse) + +lemma lookup_impl_of: + "RBT_Impl.lookup (impl_of t) = lookup t" + by (simp add: lookup_def) + +lemma entries_impl_of: + "RBT_Impl.entries (impl_of t) = entries t" + by (simp add: entries_def) + +lemma keys_impl_of: + "RBT_Impl.keys (impl_of t) = keys t" + by (simp add: keys_def) + +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 (simp add: insert_def lookup_RBT lookup_insert lookup_impl_of) -definition fold :: "('a \ 'b \ 'c \ 'c) \ ('a\linorder, 'b) rbt \ 'c \ 'c" where - [code]: "fold f t = RBT_Impl.fold f (impl_of t)" +lemma lookup_delete [simp]: + "lookup (delete k t) = (lookup t)(k := None)" + by (simp add: delete_def lookup_RBT RBT_Impl.lookup_delete lookup_impl_of 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) + +lemma entries_lookup: + "entries t1 = entries t2 \ lookup t1 = lookup t2" + by (simp add: entries_def lookup_def entries_lookup) + +lemma lookup_bulkload [simp]: + "lookup (bulkload xs) = map_of xs" + by (simp add: bulkload_def lookup_RBT RBT_Impl.lookup_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_Impl.lookup_map_entry lookup_impl_of) + +lemma lookup_map [simp]: + "lookup (map f t) k = Option.map (f k) (lookup t k)" + by (simp add: map_def lookup_RBT RBT_Impl.lookup_map lookup_impl_of) + +lemma fold_fold: + "fold f t = More_List.fold (prod_case f) (entries t)" + by (simp add: fold_def fun_eq_iff RBT_Impl.fold_def entries_impl_of) + +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) + +lemma RBT_lookup_empty [simp]: (*FIXME*) + "RBT_Impl.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 (cases t) (simp add: empty_def lookup_def RBT_inject RBT_inverse) + +lemma sorted_keys [iff]: + "sorted (keys t)" + by (simp add: keys_def RBT_Impl.keys_def sorted_entries) + +lemma distinct_keys [iff]: + "distinct (keys t)" + by (simp add: keys_def RBT_Impl.keys_def distinct_entries) + end \ No newline at end of file