1(* Author: Tobias Nipkow *)
2
3section "AA Tree Implementation of Maps"
4
5theory AA_Map
6imports
7  AA_Set
8  Lookup2
9begin
10
11fun update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) aa_tree \<Rightarrow> ('a*'b) aa_tree" where
12"update x y Leaf = Node Leaf (x,y) 1 Leaf" |
13"update x y (Node t1 (a,b) lv t2) =
14  (case cmp x a of
15     LT \<Rightarrow> split (skew (Node (update x y t1) (a,b) lv t2)) |
16     GT \<Rightarrow> split (skew (Node t1 (a,b) lv (update x y t2))) |
17     EQ \<Rightarrow> Node t1 (x,y) lv t2)"
18
19fun delete :: "'a::linorder \<Rightarrow> ('a*'b) aa_tree \<Rightarrow> ('a*'b) aa_tree" where
20"delete _ Leaf = Leaf" |
21"delete x (Node l (a,b) lv r) =
22  (case cmp x a of
23     LT \<Rightarrow> adjust (Node (delete x l) (a,b) lv r) |
24     GT \<Rightarrow> adjust (Node l (a,b) lv (delete x r)) |
25     EQ \<Rightarrow> (if l = Leaf then r
26            else let (l',ab') = split_max l in adjust (Node l' ab' lv r)))"
27
28
29subsection "Invariance"
30
31subsubsection "Proofs for insert"
32
33lemma lvl_update_aux:
34  "lvl (update x y t) = lvl t \<or> lvl (update x y t) = lvl t + 1 \<and> sngl (update x y t)"
35apply(induction t)
36apply (auto simp: lvl_skew)
37apply (metis Suc_eq_plus1 lvl.simps(2) lvl_split lvl_skew)+
38done
39
40lemma lvl_update: obtains
41  (Same) "lvl (update x y t) = lvl t" |
42  (Incr) "lvl (update x y t) = lvl t + 1" "sngl (update x y t)"
43using lvl_update_aux by fastforce
44
45declare invar.simps(2)[simp]
46
47lemma lvl_update_sngl: "invar t \<Longrightarrow> sngl t \<Longrightarrow> lvl(update x y t) = lvl t"
48proof (induction t rule: update.induct)
49  case (2 x y t1 a b lv t2)
50  consider (LT) "x < a" | (GT) "x > a" | (EQ) "x = a" 
51    using less_linear by blast 
52  thus ?case proof cases
53    case LT
54    thus ?thesis using 2 by (auto simp add: skew_case split_case split: tree.splits)
55  next
56    case GT
57    thus ?thesis using 2 proof (cases t1)
58      case Node
59      thus ?thesis using 2 GT  
60        apply (auto simp add: skew_case split_case split: tree.splits)
61        by (metis less_not_refl2 lvl.simps(2) lvl_update_aux n_not_Suc_n sngl.simps(3))+ 
62    qed (auto simp add: lvl_0_iff)
63  qed simp
64qed simp
65
66lemma lvl_update_incr_iff: "(lvl(update a b t) = lvl t + 1) \<longleftrightarrow>
67  (\<exists>l x r. update a b t = Node l x (lvl t + 1) r \<and> lvl l = lvl r)"
68apply(cases t)
69apply(auto simp add: skew_case split_case split: if_splits)
70apply(auto split: tree.splits if_splits)
71done
72
73lemma invar_update: "invar t \<Longrightarrow> invar(update a b t)"
74proof(induction t)
75  case N: (Node l xy n r)
76  hence il: "invar l" and ir: "invar r" by auto
77  note iil = N.IH(1)[OF il]
78  note iir = N.IH(2)[OF ir]
79  obtain x y where [simp]: "xy = (x,y)" by fastforce
80  let ?t = "Node l xy n r"
81  have "a < x \<or> a = x \<or> x < a" by auto
82  moreover
83  have ?case if "a < x"
84  proof (cases rule: lvl_update[of a b l])
85    case (Same) thus ?thesis
86      using \<open>a<x\<close> invar_NodeL[OF N.prems iil Same]
87      by (simp add: skew_invar split_invar del: invar.simps)
88  next
89    case (Incr)
90    then obtain t1 w t2 where ial[simp]: "update a b l = Node t1 w n t2"
91      using N.prems by (auto simp: lvl_Suc_iff)
92    have l12: "lvl t1 = lvl t2"
93      by (metis Incr(1) ial lvl_update_incr_iff tree.inject)
94    have "update a b ?t = split(skew(Node (update a b l) xy n r))"
95      by(simp add: \<open>a<x\<close>)
96    also have "skew(Node (update a b l) xy n r) = Node t1 w n (Node t2 xy n r)"
97      by(simp)
98    also have "invar(split \<dots>)"
99    proof (cases r)
100      case Leaf
101      hence "l = Leaf" using N.prems by(auto simp: lvl_0_iff)
102      thus ?thesis using Leaf ial by simp
103    next
104      case [simp]: (Node t3 y m t4)
105      show ?thesis (*using N(3) iil l12 by(auto)*)
106      proof cases
107        assume "m = n" thus ?thesis using N(3) iil by(auto)
108      next
109        assume "m \<noteq> n" thus ?thesis using N(3) iil l12 by(auto)
110      qed
111    qed
112    finally show ?thesis .
113  qed
114  moreover
115  have ?case if "x < a"
116  proof -
117    from \<open>invar ?t\<close> have "n = lvl r \<or> n = lvl r + 1" by auto
118    thus ?case
119    proof
120      assume 0: "n = lvl r"
121      have "update a b ?t = split(skew(Node l xy n (update a b r)))"
122        using \<open>a>x\<close> by(auto)
123      also have "skew(Node l xy n (update a b r)) = Node l xy n (update a b r)"
124        using N.prems by(simp add: skew_case split: tree.split)
125      also have "invar(split \<dots>)"
126      proof -
127        from lvl_update_sngl[OF ir sngl_if_invar[OF \<open>invar ?t\<close> 0], of a b]
128        obtain t1 p t2 where iar: "update a b r = Node t1 p n t2"
129          using N.prems 0 by (auto simp: lvl_Suc_iff)
130        from N.prems iar 0 iir
131        show ?thesis by (auto simp: split_case split: tree.splits)
132      qed
133      finally show ?thesis .
134    next
135      assume 1: "n = lvl r + 1"
136      hence "sngl ?t" by(cases r) auto
137      show ?thesis
138      proof (cases rule: lvl_update[of a b r])
139        case (Same)
140        show ?thesis using \<open>x<a\<close> il ir invar_NodeR[OF N.prems 1 iir Same]
141          by (auto simp add: skew_invar split_invar)
142      next
143        case (Incr)
144        thus ?thesis using invar_NodeR2[OF \<open>invar ?t\<close> Incr(2) 1 iir] 1 \<open>x < a\<close>
145          by (auto simp add: skew_invar split_invar split: if_splits)
146      qed
147    qed
148  qed
149  moreover
150  have "a = x \<Longrightarrow> ?case" using N.prems by auto
151  ultimately show ?case by blast
152qed simp
153
154subsubsection "Proofs for delete"
155
156declare invar.simps(2)[simp del]
157
158theorem post_delete: "invar t \<Longrightarrow> post_del t (delete x t)"
159proof (induction t)
160  case (Node l ab lv r)
161
162  obtain a b where [simp]: "ab = (a,b)" by fastforce
163
164  let ?l' = "delete x l" and ?r' = "delete x r"
165  let ?t = "Node l ab lv r" let ?t' = "delete x ?t"
166
167  from Node.prems have inv_l: "invar l" and inv_r: "invar r" by (auto)
168
169  note post_l' = Node.IH(1)[OF inv_l]
170  note preL = pre_adj_if_postL[OF Node.prems post_l']
171
172  note post_r' = Node.IH(2)[OF inv_r]
173  note preR = pre_adj_if_postR[OF Node.prems post_r']
174
175  show ?case
176  proof (cases rule: linorder_cases[of x a])
177    case less
178    thus ?thesis using Node.prems by (simp add: post_del_adjL preL)
179  next
180    case greater
181    thus ?thesis using Node.prems preR by (simp add: post_del_adjR post_r')
182  next
183    case equal
184    show ?thesis
185    proof cases
186      assume "l = Leaf" thus ?thesis using equal Node.prems
187        by(auto simp: post_del_def invar.simps(2))
188    next
189      assume "l \<noteq> Leaf" thus ?thesis using equal Node.prems
190        by simp (metis inv_l post_del_adjL post_split_max pre_adj_if_postL)
191    qed
192  qed
193qed (simp add: post_del_def)
194
195
196subsection \<open>Functional Correctness Proofs\<close>
197
198theorem inorder_update:
199  "sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
200by (induct t) (auto simp: upd_list_simps inorder_split inorder_skew)
201
202theorem inorder_delete:
203  "\<lbrakk>invar t; sorted1(inorder t)\<rbrakk> \<Longrightarrow>
204  inorder (delete x t) = del_list x (inorder t)"
205by(induction t)
206  (auto simp: del_list_simps inorder_adjust pre_adj_if_postL pre_adj_if_postR 
207              post_split_max post_delete split_maxD split: prod.splits)
208
209interpretation I: Map_by_Ordered
210where empty = empty and lookup = lookup and update = update and delete = delete
211and inorder = inorder and inv = invar
212proof (standard, goal_cases)
213  case 1 show ?case by (simp add: empty_def)
214next
215  case 2 thus ?case by(simp add: lookup_map_of)
216next
217  case 3 thus ?case by(simp add: inorder_update)
218next
219  case 4 thus ?case by(simp add: inorder_delete)
220next
221  case 5 thus ?case by(simp add: empty_def)
222next
223  case 6 thus ?case by(simp add: invar_update)
224next
225  case 7 thus ?case using post_delete by(auto simp: post_del_def)
226qed
227
228end
229