| author | wenzelm | 
| Tue, 23 May 2023 21:43:36 +0200 | |
| changeset 78099 | 4d9349989d94 | 
| parent 73526 | a3cc9fa1295d | 
| permissions | -rw-r--r-- | 
| 61789 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 62130 | 3 | section \<open>1-2 Brother Tree Implementation of Maps\<close> | 
| 61789 | 4 | |
| 5 | theory Brother12_Map | |
| 6 | imports | |
| 7 | Brother12_Set | |
| 67965 | 8 | Map_Specs | 
| 61789 | 9 | begin | 
| 10 | ||
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62130diff
changeset | 11 | fun lookup :: "('a \<times> 'b) bro \<Rightarrow> 'a::linorder \<Rightarrow> 'b option" where
 | 
| 61789 | 12 | "lookup N0 x = None" | | 
| 13 | "lookup (N1 t) x = lookup t x" | | |
| 14 | "lookup (N2 l (a,b) r) x = | |
| 15 | (case cmp x a of | |
| 16 | LT \<Rightarrow> lookup l x | | |
| 17 | EQ \<Rightarrow> Some b | | |
| 18 | GT \<Rightarrow> lookup r x)" | |
| 19 | ||
| 20 | locale update = insert | |
| 21 | begin | |
| 22 | ||
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62130diff
changeset | 23 | fun upd :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a\<times>'b) bro \<Rightarrow> ('a\<times>'b) bro" where
 | 
| 61789 | 24 | "upd x y N0 = L2 (x,y)" | | 
| 25 | "upd x y (N1 t) = n1 (upd x y t)" | | |
| 26 | "upd x y (N2 l (a,b) r) = | |
| 27 | (case cmp x a of | |
| 28 | LT \<Rightarrow> n2 (upd x y l) (a,b) r | | |
| 29 | EQ \<Rightarrow> N2 l (a,y) r | | |
| 30 | GT \<Rightarrow> n2 l (a,b) (upd x y r))" | |
| 31 | ||
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62130diff
changeset | 32 | definition update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a\<times>'b) bro \<Rightarrow> ('a\<times>'b) bro" where
 | 
| 61789 | 33 | "update x y t = tree(upd x y t)" | 
| 34 | ||
| 35 | end | |
| 36 | ||
| 37 | context delete | |
| 38 | begin | |
| 39 | ||
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62130diff
changeset | 40 | fun del :: "'a::linorder \<Rightarrow> ('a\<times>'b) bro \<Rightarrow> ('a\<times>'b) bro" where
 | 
| 61789 | 41 | "del _ N0 = N0" | | 
| 42 | "del x (N1 t) = N1 (del x t)" | | |
| 43 | "del x (N2 l (a,b) r) = | |
| 44 | (case cmp x a of | |
| 45 | LT \<Rightarrow> n2 (del x l) (a,b) r | | |
| 46 | GT \<Rightarrow> n2 l (a,b) (del x r) | | |
| 68020 | 47 | EQ \<Rightarrow> (case split_min r of | 
| 61789 | 48 | None \<Rightarrow> N1 l | | 
| 49 | Some (ab, r') \<Rightarrow> n2 l ab r'))" | |
| 50 | ||
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62130diff
changeset | 51 | definition delete :: "'a::linorder \<Rightarrow> ('a\<times>'b) bro \<Rightarrow> ('a\<times>'b) bro" where
 | 
| 61789 | 52 | "delete a t = tree (del a t)" | 
| 53 | ||
| 54 | end | |
| 55 | ||
| 56 | subsection "Functional Correctness Proofs" | |
| 57 | ||
| 58 | subsubsection "Proofs for lookup" | |
| 59 | ||
| 60 | lemma lookup_map_of: "t \<in> T h \<Longrightarrow> | |
| 61 | sorted1(inorder t) \<Longrightarrow> lookup t x = map_of (inorder t) x" | |
| 62 | by(induction h arbitrary: t) (auto simp: map_of_simps split: option.splits) | |
| 63 | ||
| 64 | subsubsection "Proofs for update" | |
| 65 | ||
| 66 | context update | |
| 67 | begin | |
| 68 | ||
| 69 | lemma inorder_upd: "t \<in> T h \<Longrightarrow> | |
| 70 | sorted1(inorder t) \<Longrightarrow> inorder(upd x y t) = upd_list x y (inorder t)" | |
| 71 | by(induction h arbitrary: t) (auto simp: upd_list_simps inorder_n1 inorder_n2) | |
| 72 | ||
| 73 | lemma inorder_update: "t \<in> T h \<Longrightarrow> | |
| 74 | sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)" | |
| 75 | by(simp add: update_def inorder_upd inorder_tree) | |
| 76 | ||
| 77 | end | |
| 78 | ||
| 79 | subsubsection \<open>Proofs for deletion\<close> | |
| 80 | ||
| 81 | context delete | |
| 82 | begin | |
| 83 | ||
| 84 | lemma inorder_del: | |
| 61792 | 85 | "t \<in> T h \<Longrightarrow> sorted1(inorder t) \<Longrightarrow> inorder(del x t) = del_list x (inorder t)" | 
| 73526 
a3cc9fa1295d
new automatic order prover: stateless, complete, verified
 nipkow parents: 
68431diff
changeset | 86 | apply (induction h arbitrary: t) | 
| 
a3cc9fa1295d
new automatic order prover: stateless, complete, verified
 nipkow parents: 
68431diff
changeset | 87 | apply (auto simp: del_list_simps inorder_n2) | 
| 
a3cc9fa1295d
new automatic order prover: stateless, complete, verified
 nipkow parents: 
68431diff
changeset | 88 | apply (auto simp: del_list_simps inorder_n2 | 
| 68020 | 89 | inorder_split_min[OF UnI1] inorder_split_min[OF UnI2] split: option.splits) | 
| 73526 
a3cc9fa1295d
new automatic order prover: stateless, complete, verified
 nipkow parents: 
68431diff
changeset | 90 | done | 
| 61792 | 91 | |
| 92 | lemma inorder_delete: | |
| 93 | "t \<in> T h \<Longrightarrow> sorted1(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)" | |
| 94 | by(simp add: delete_def inorder_del inorder_tree) | |
| 61789 | 95 | |
| 96 | end | |
| 97 | ||
| 98 | ||
| 99 | subsection \<open>Invariant Proofs\<close> | |
| 100 | ||
| 101 | subsubsection \<open>Proofs for update\<close> | |
| 102 | ||
| 103 | context update | |
| 104 | begin | |
| 105 | ||
| 106 | lemma upd_type: | |
| 107 | "(t \<in> B h \<longrightarrow> upd x y t \<in> Bp h) \<and> (t \<in> U h \<longrightarrow> upd x y t \<in> T h)" | |
| 108 | apply(induction h arbitrary: t) | |
| 109 | apply (simp) | |
| 110 | apply (fastforce simp: Bp_if_B n2_type dest: n1_type) | |
| 111 | done | |
| 112 | ||
| 113 | lemma update_type: | |
| 61809 | 114 | "t \<in> B h \<Longrightarrow> update x y t \<in> B h \<union> B (Suc h)" | 
| 115 | unfolding update_def by (metis upd_type tree_type) | |
| 61789 | 116 | |
| 117 | end | |
| 118 | ||
| 119 | subsubsection "Proofs for deletion" | |
| 120 | ||
| 121 | context delete | |
| 122 | begin | |
| 123 | ||
| 124 | lemma del_type: | |
| 125 | "t \<in> B h \<Longrightarrow> del x t \<in> T h" | |
| 126 | "t \<in> U h \<Longrightarrow> del x t \<in> Um h" | |
| 127 | proof (induction h arbitrary: x t) | |
| 128 | case (Suc h) | |
| 129 |   { case 1
 | |
| 130 | then obtain l a b r where [simp]: "t = N2 l (a,b) r" and | |
| 131 | lr: "l \<in> T h" "r \<in> T h" "l \<in> B h \<or> r \<in> B h" by auto | |
| 67040 | 132 | have ?case if "x < a" | 
| 133 | proof cases | |
| 134 | assume "l \<in> B h" | |
| 135 | from n2_type3[OF Suc.IH(1)[OF this] lr(2)] | |
| 67406 | 136 | show ?thesis using \<open>x<a\<close> by(simp) | 
| 67040 | 137 | next | 
| 138 | assume "l \<notin> B h" | |
| 139 | hence "l \<in> U h" "r \<in> B h" using lr by auto | |
| 140 | from n2_type1[OF Suc.IH(2)[OF this(1)] this(2)] | |
| 67406 | 141 | show ?thesis using \<open>x<a\<close> by(simp) | 
| 67040 | 142 | qed | 
| 143 | moreover | |
| 144 | have ?case if "x > a" | |
| 145 | proof cases | |
| 146 | assume "r \<in> B h" | |
| 147 | from n2_type3[OF lr(1) Suc.IH(1)[OF this]] | |
| 67406 | 148 | show ?thesis using \<open>x>a\<close> by(simp) | 
| 67040 | 149 | next | 
| 150 | assume "r \<notin> B h" | |
| 151 | hence "l \<in> B h" "r \<in> U h" using lr by auto | |
| 152 | from n2_type2[OF this(1) Suc.IH(2)[OF this(2)]] | |
| 67406 | 153 | show ?thesis using \<open>x>a\<close> by(simp) | 
| 67040 | 154 | qed | 
| 155 | moreover | |
| 156 | have ?case if [simp]: "x=a" | |
| 68020 | 157 | proof (cases "split_min r") | 
| 67040 | 158 | case None | 
| 159 | show ?thesis | |
| 61789 | 160 | proof cases | 
| 161 | assume "r \<in> B h" | |
| 68020 | 162 | with split_minNoneN0[OF this None] lr show ?thesis by(simp) | 
| 61789 | 163 | next | 
| 164 | assume "r \<notin> B h" | |
| 67040 | 165 | hence "r \<in> U h" using lr by auto | 
| 68020 | 166 | with split_minNoneN1[OF this None] lr(3) show ?thesis by (simp) | 
| 61789 | 167 | qed | 
| 67040 | 168 | next | 
| 169 | case [simp]: (Some br') | |
| 170 | obtain b r' where [simp]: "br' = (b,r')" by fastforce | |
| 171 | show ?thesis | |
| 172 | proof cases | |
| 173 | assume "r \<in> B h" | |
| 68020 | 174 | from split_min_type(1)[OF this] n2_type3[OF lr(1)] | 
| 67040 | 175 | show ?thesis by simp | 
| 61789 | 176 | next | 
| 67040 | 177 | assume "r \<notin> B h" | 
| 178 | hence "l \<in> B h" and "r \<in> U h" using lr by auto | |
| 68020 | 179 | from split_min_type(2)[OF this(2)] n2_type2[OF this(1)] | 
| 67040 | 180 | show ?thesis by simp | 
| 61789 | 181 | qed | 
| 67040 | 182 | qed | 
| 183 | ultimately show ?case by auto | |
| 61789 | 184 | } | 
| 185 |   { case 2 with Suc.IH(1) show ?case by auto }
 | |
| 186 | qed auto | |
| 187 | ||
| 188 | lemma delete_type: | |
| 61809 | 189 | "t \<in> B h \<Longrightarrow> delete x t \<in> B h \<union> B(h-1)" | 
| 61789 | 190 | unfolding delete_def | 
| 61809 | 191 | by (cases h) (simp, metis del_type(1) tree_type Suc_eq_plus1 diff_Suc_1) | 
| 61789 | 192 | |
| 193 | end | |
| 194 | ||
| 195 | subsection "Overall correctness" | |
| 196 | ||
| 197 | interpretation Map_by_Ordered | |
| 68431 | 198 | where empty = empty and lookup = lookup and update = update.update | 
| 61809 | 199 | and delete = delete.delete and inorder = inorder and inv = "\<lambda>t. \<exists>h. t \<in> B h" | 
| 61789 | 200 | proof (standard, goal_cases) | 
| 201 | case 2 thus ?case by(auto intro!: lookup_map_of) | |
| 202 | next | |
| 203 | case 3 thus ?case by(auto intro!: update.inorder_update) | |
| 204 | next | |
| 61792 | 205 | case 4 thus ?case by(auto intro!: delete.inorder_delete) | 
| 61789 | 206 | next | 
| 207 | case 6 thus ?case using update.update_type by (metis Un_iff) | |
| 208 | next | |
| 209 | case 7 thus ?case using delete.delete_type by blast | |
| 68431 | 210 | qed (auto simp: empty_def) | 
| 61789 | 211 | |
| 212 | end |