(* Title: Isabelle Collections Library Author: Andreas Lochbihler <andreas dot lochbihler at kit.edu> Maintainer: Andreas Lochbihler <andreas dot lochbihler at kit.edu> *) section ‹\isaheader{Array-based hash map implementation}› theory ArrayHashMap_Impl imports "../../Lib/HashCode" "../../Lib/Code_Target_ICF" "../../Lib/Diff_Array" "../gen_algo/ListGA" ListMapImpl "../../Iterator/Array_Iterator" begin text ‹Misc.› setup Locale_Code.open_block interpretation a_idx_it: idx_iteratei_loc list_of_array "λ_. True" array_length array_get apply unfold_locales apply (case_tac [!] s) [2] apply auto done setup Locale_Code.close_block (* lemma idx_iteratei_aux_array_get_Array_conv_nth: "idx_iteratei_aux array_get sz i (Array xs) c f σ = idx_iteratei_aux (!) sz i xs c f σ" apply(induct get≡"(!) :: 'b list ⇒ nat ⇒ 'b" sz i xs c f σ rule: idx_iteratei_aux.induct) apply(subst (1 2) idx_iteratei_aux.simps) apply simp done lemma idx_iteratei_array_get_Array_conv_nth: "idx_iteratei array_get array_length (Array xs) = idx_iteratei nth length xs" by(simp add: idx_iteratei_def fun_eq_iff idx_iteratei_aux_array_get_Array_conv_nth) lemma idx_iteratei_aux_nth_conv_foldli_drop: fixes xs :: "'b list" assumes "i ≤ length xs" shows "idx_iteratei_aux (!) (length xs) i xs c f σ = foldli (drop (length xs - i) xs) c f σ" using assms proof(induct get≡"(!) :: 'b list ⇒ nat ⇒ 'b" sz≡"length xs" i xs c f σ rule: idx_iteratei_aux.induct) case (1 i l c f σ) show ?case proof(cases "i = 0 ∨ ¬ c σ") case True thus ?thesis by(subst idx_iteratei_aux.simps)(auto) next case False hence i: "i > 0" and c: "c σ" by auto hence "idx_iteratei_aux (!) (length l) i l c f σ = idx_iteratei_aux (!) (length l) (i - 1) l c f (f (l ! (length l - i)) σ)" by(subst idx_iteratei_aux.simps) simp also have "… = foldli (drop (length l - (i - 1)) l) c f (f (l ! (length l - i)) σ)" using `i ≤ length l` i c by -(rule 1, auto) also from `i ≤ length l` i have "drop (length l - i) l = (l ! (length l - i)) # drop (length l - (i - 1)) l" by(subst Cons_nth_drop_Suc[symmetric])(simp_all, metis Suc_eq_plus1_left add_diff_assoc) hence "foldli (drop (length l - (i - 1)) l) c f (f (l ! (length l - i)) σ) = foldli (drop (length l - i) l) c f σ" using c by simp finally show ?thesis . qed qed lemma idx_iteratei_nth_length_conv_foldli: "idx_iteratei nth length = foldli" by(rule ext)+(simp add: idx_iteratei_def idx_iteratei_aux_nth_conv_foldli_drop) *) subsection ‹Type definition and primitive operations› definition load_factor :: nat ― ‹in percent› where "load_factor = 75" text ‹ We do not use @{typ "('k, 'v) assoc_list"} for the buckets but plain lists of key-value pairs. This speeds up rehashing because we then do not have to go through the abstract operations. ›