| author | wenzelm | 
| Tue, 26 Jan 2021 23:34:40 +0100 | |
| changeset 73194 | c0d6d57a9a31 | 
| parent 72979 | e734cd65c926 | 
| child 77266 | 334015f9098e | 
| permissions | -rw-r--r-- | 
| 70250 | 1 | section "Tries via Search Trees" | 
| 2 | ||
| 3 | theory Trie_Map | |
| 4 | imports | |
| 5 | RBT_Map | |
| 6 | Trie_Fun | |
| 7 | begin | |
| 8 | ||
| 9 | text \<open>An implementation of tries based on maps implemented by red-black trees. | |
| 10 | Works for any kind of search tree.\<close> | |
| 11 | ||
| 12 | text \<open>Implementation of map:\<close> | |
| 13 | ||
| 14 | type_synonym 'a mapi = "'a rbt" | |
| 15 | ||
| 70262 | 16 | datatype 'a trie_map = Nd bool "('a * 'a trie_map) mapi"
 | 
| 70250 | 17 | |
| 18 | text \<open>In principle one should be able to given an implementation of tries | |
| 19 | once and for all for any map implementation and not just for a specific one (RBT) as done here. | |
| 20 | But because the map (@{typ "'a rbt"}) is used in a datatype, the HOL type system does not support this.
 | |
| 21 | ||
| 22 | However, the development below works verbatim for any map implementation, eg \<open>Tree_Map\<close>, | |
| 23 | and not just \<open>RBT_Map\<close>, except for the termination lemma \<open>lookup_size\<close>.\<close> | |
| 70755 
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
 nipkow parents: 
70266diff
changeset | 24 | term size_tree | 
| 70250 | 25 | lemma lookup_size[termination_simp]: | 
| 26 |   fixes t :: "('a::linorder * 'a trie_map) rbt"
 | |
| 70755 
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
 nipkow parents: 
70266diff
changeset | 27 | shows "lookup t a = Some b \<Longrightarrow> size b < Suc (size_tree (\<lambda>ab. Suc(Suc (size (snd(fst ab))))) t)" | 
| 70250 | 28 | apply(induction t a rule: lookup.induct) | 
| 29 | apply(auto split: if_splits) | |
| 30 | done | |
| 31 | ||
| 70266 | 32 | |
| 33 | definition empty :: "'a trie_map" where | |
| 34 | [simp]: "empty = Nd False Leaf" | |
| 35 | ||
| 70250 | 36 | fun isin :: "('a::linorder) trie_map \<Rightarrow> 'a list \<Rightarrow> bool" where
 | 
| 37 | "isin (Nd b m) [] = b" | | |
| 38 | "isin (Nd b m) (x # xs) = (case lookup m x of None \<Rightarrow> False | Some t \<Rightarrow> isin t xs)" | |
| 39 | ||
| 40 | fun insert :: "('a::linorder) list \<Rightarrow> 'a trie_map \<Rightarrow> 'a trie_map" where
 | |
| 41 | "insert [] (Nd b m) = Nd True m" | | |
| 42 | "insert (x#xs) (Nd b m) = | |
| 70266 | 43 | Nd b (update x (insert xs (case lookup m x of None \<Rightarrow> empty | Some t \<Rightarrow> t)) m)" | 
| 70250 | 44 | |
| 45 | fun delete :: "('a::linorder) list \<Rightarrow> 'a trie_map \<Rightarrow> 'a trie_map" where
 | |
| 46 | "delete [] (Nd b m) = Nd False m" | | |
| 47 | "delete (x#xs) (Nd b m) = Nd b | |
| 48 | (case lookup m x of | |
| 49 | None \<Rightarrow> m | | |
| 50 | Some t \<Rightarrow> update x (delete xs t) m)" | |
| 51 | ||
| 52 | ||
| 53 | subsection "Correctness" | |
| 54 | ||
| 55 | text \<open>Proof by stepwise refinement. First abstract to type @{typ "'a trie"}.\<close>
 | |
| 56 | ||
| 57 | fun abs :: "'a::linorder trie_map \<Rightarrow> 'a trie" where | |
| 70263 | 58 | "abs (Nd b t) = Trie_Fun.Nd b (\<lambda>a. map_option abs (lookup t a))" | 
| 70250 | 59 | |
| 60 | fun invar :: "('a::linorder)trie_map \<Rightarrow> bool" where
 | |
| 61 | "invar (Nd b m) = (M.invar m \<and> (\<forall>a t. lookup m a = Some t \<longrightarrow> invar t))" | |
| 62 | ||
| 70263 | 63 | lemma isin_abs: "isin t xs = Trie_Fun.isin (abs t) xs" | 
| 70250 | 64 | apply(induction t xs rule: isin.induct) | 
| 65 | apply(auto split: option.split) | |
| 66 | done | |
| 67 | ||
| 70263 | 68 | lemma abs_insert: "invar t \<Longrightarrow> abs(insert xs t) = Trie_Fun.insert xs (abs t)" | 
| 70250 | 69 | apply(induction xs t rule: insert.induct) | 
| 70262 | 70 | apply(auto simp: M.map_specs RBT_Set.empty_def[symmetric] split: option.split) | 
| 70250 | 71 | done | 
| 72 | ||
| 70263 | 73 | lemma abs_delete: "invar t \<Longrightarrow> abs(delete xs t) = Trie_Fun.delete xs (abs t)" | 
| 70250 | 74 | apply(induction xs t rule: delete.induct) | 
| 75 | apply(auto simp: M.map_specs split: option.split) | |
| 76 | done | |
| 77 | ||
| 78 | lemma invar_insert: "invar t \<Longrightarrow> invar (insert xs t)" | |
| 79 | apply(induction xs t rule: insert.induct) | |
| 70262 | 80 | apply(auto simp: M.map_specs RBT_Set.empty_def[symmetric] split: option.split) | 
| 70250 | 81 | done | 
| 82 | ||
| 83 | lemma invar_delete: "invar t \<Longrightarrow> invar (delete xs t)" | |
| 84 | apply(induction xs t rule: delete.induct) | |
| 85 | apply(auto simp: M.map_specs split: option.split) | |
| 86 | done | |
| 87 | ||
| 88 | text \<open>Overall correctness w.r.t. the \<open>Set\<close> ADT:\<close> | |
| 89 | ||
| 90 | interpretation S2: Set | |
| 70266 | 91 | where empty = empty and isin = isin and insert = insert and delete = delete | 
| 70250 | 92 | and set = "set o abs" and invar = invar | 
| 93 | proof (standard, goal_cases) | |
| 72979 | 94 | case 1 show ?case by (simp add: isin_case split: list.split) | 
| 70250 | 95 | next | 
| 72979 | 96 | case 2 thus ?case by (simp add: isin_abs) | 
| 70250 | 97 | next | 
| 72979 | 98 | case 3 thus ?case by (simp add: set_insert abs_insert del: set_def) | 
| 70250 | 99 | next | 
| 72979 | 100 | case 4 thus ?case by (simp add: set_delete abs_delete del: set_def) | 
| 70250 | 101 | next | 
| 70262 | 102 | case 5 thus ?case by (simp add: M.map_specs RBT_Set.empty_def[symmetric]) | 
| 70250 | 103 | next | 
| 104 | case 6 thus ?case by (simp add: invar_insert) | |
| 105 | next | |
| 106 | case 7 thus ?case by (simp add: invar_delete) | |
| 107 | qed | |
| 108 | ||
| 109 | ||
| 110 | end |