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