| 61793 |      1 | (*
 | 
| 62496 |      2 | Author: Tobias Nipkow and Daniel Stüwe
 | 
| 61793 |      3 | *)
 | 
|  |      4 | 
 | 
| 62130 |      5 | section \<open>AA Tree Implementation of Sets\<close>
 | 
| 61793 |      6 | 
 | 
|  |      7 | theory AA_Set
 | 
|  |      8 | imports
 | 
|  |      9 |   Isin2
 | 
|  |     10 |   Cmp
 | 
|  |     11 | begin
 | 
|  |     12 | 
 | 
|  |     13 | type_synonym 'a aa_tree = "('a,nat) tree"
 | 
|  |     14 | 
 | 
|  |     15 | fun lvl :: "'a aa_tree \<Rightarrow> nat" where
 | 
|  |     16 | "lvl Leaf = 0" |
 | 
|  |     17 | "lvl (Node lv _ _ _) = lv"
 | 
| 62496 |     18 | 
 | 
| 61793 |     19 | fun invar :: "'a aa_tree \<Rightarrow> bool" where
 | 
|  |     20 | "invar Leaf = True" |
 | 
|  |     21 | "invar (Node h l a r) =
 | 
|  |     22 |  (invar l \<and> invar r \<and>
 | 
|  |     23 |   h = lvl l + 1 \<and> (h = lvl r + 1 \<or> (\<exists>lr b rr. r = Node h lr b rr \<and> h = lvl rr + 1)))"
 | 
| 62496 |     24 | 
 | 
| 61793 |     25 | fun skew :: "'a aa_tree \<Rightarrow> 'a aa_tree" where
 | 
|  |     26 | "skew (Node lva (Node lvb t1 b t2) a t3) =
 | 
|  |     27 |   (if lva = lvb then Node lva t1 b (Node lva t2 a t3) else Node lva (Node lvb t1 b t2) a t3)" |
 | 
|  |     28 | "skew t = t"
 | 
|  |     29 | 
 | 
|  |     30 | fun split :: "'a aa_tree \<Rightarrow> 'a aa_tree" where
 | 
|  |     31 | "split (Node lva t1 a (Node lvb t2 b (Node lvc t3 c t4))) =
 | 
|  |     32 |    (if lva = lvb \<and> lvb = lvc (* lva = lvc suffices *)
 | 
|  |     33 |     then Node (lva+1) (Node lva t1 a t2) b (Node lva t3 c t4)
 | 
|  |     34 |     else Node lva t1 a (Node lvb t2 b (Node lvc t3 c t4)))" |
 | 
|  |     35 | "split t = t"
 | 
|  |     36 | 
 | 
|  |     37 | hide_const (open) insert
 | 
|  |     38 | 
 | 
|  |     39 | fun insert :: "'a::cmp \<Rightarrow> 'a aa_tree \<Rightarrow> 'a aa_tree" where
 | 
|  |     40 | "insert x Leaf = Node 1 Leaf x Leaf" |
 | 
|  |     41 | "insert x (Node lv t1 a t2) =
 | 
|  |     42 |   (case cmp x a of
 | 
|  |     43 |      LT \<Rightarrow> split (skew (Node lv (insert x t1) a t2)) |
 | 
|  |     44 |      GT \<Rightarrow> split (skew (Node lv t1 a (insert x t2))) |
 | 
|  |     45 |      EQ \<Rightarrow> Node lv t1 x t2)"
 | 
|  |     46 | 
 | 
|  |     47 | fun sngl :: "'a aa_tree \<Rightarrow> bool" where
 | 
|  |     48 | "sngl Leaf = False" |
 | 
|  |     49 | "sngl (Node _ _ _ Leaf) = True" |
 | 
|  |     50 | "sngl (Node lva _ _ (Node lvb _ _ _)) = (lva > lvb)"
 | 
|  |     51 | 
 | 
|  |     52 | definition adjust :: "'a aa_tree \<Rightarrow> 'a aa_tree" where
 | 
|  |     53 | "adjust t =
 | 
|  |     54 |  (case t of
 | 
|  |     55 |   Node lv l x r \<Rightarrow>
 | 
|  |     56 |    (if lvl l >= lv-1 \<and> lvl r >= lv-1 then t else
 | 
|  |     57 |     if lvl r < lv-1 \<and> sngl l then skew (Node (lv-1) l x r) else
 | 
|  |     58 |     if lvl r < lv-1
 | 
|  |     59 |     then case l of
 | 
|  |     60 |            Node lva t1 a (Node lvb t2 b t3)
 | 
| 62496 |     61 |              \<Rightarrow> Node (lvb+1) (Node lva t1 a t2) b (Node (lv-1) t3 x r) 
 | 
| 61793 |     62 |     else
 | 
|  |     63 |     if lvl r < lv then split (Node (lv-1) l x r)
 | 
|  |     64 |     else
 | 
|  |     65 |       case r of
 | 
| 62160 |     66 |         Node lvb t1 b t4 \<Rightarrow>
 | 
| 61793 |     67 |           (case t1 of
 | 
|  |     68 |              Node lva t2 a t3
 | 
|  |     69 |                \<Rightarrow> Node (lva+1) (Node (lv-1) l x t2) a
 | 
| 62496 |     70 |                     (split (Node (if sngl t1 then lva else lva+1) t3 b t4)))))"
 | 
|  |     71 | 
 | 
|  |     72 | text{* In the paper, the last case of @{const adjust} is expressed with the help of an
 | 
|  |     73 | incorrect auxiliary function \texttt{nlvl}.
 | 
|  |     74 | 
 | 
|  |     75 | Function @{text del_max} below is called \texttt{dellrg} in the paper.
 | 
|  |     76 | The latter is incorrect for two reasons: \texttt{dellrg} is meant to delete the largest
 | 
|  |     77 | element but recurses on the left instead of the right subtree; the invariant
 | 
|  |     78 | is not restored.*}
 | 
|  |     79 | 
 | 
|  |     80 | fun del_max :: "'a aa_tree \<Rightarrow> 'a aa_tree * 'a" where
 | 
|  |     81 | "del_max (Node lv l a Leaf) = (l,a)" |
 | 
|  |     82 | "del_max (Node lv l a r) = (let (r',b) = del_max r in (adjust(Node lv l a r'), b))"
 | 
| 61793 |     83 | 
 | 
|  |     84 | fun delete :: "'a::cmp \<Rightarrow> 'a aa_tree \<Rightarrow> 'a aa_tree" where
 | 
|  |     85 | "delete _ Leaf = Leaf" |
 | 
|  |     86 | "delete x (Node lv l a r) =
 | 
|  |     87 |   (case cmp x a of
 | 
|  |     88 |      LT \<Rightarrow> adjust (Node lv (delete x l) a r) |
 | 
|  |     89 |      GT \<Rightarrow> adjust (Node lv l a (delete x r)) |
 | 
|  |     90 |      EQ \<Rightarrow> (if l = Leaf then r
 | 
|  |     91 |             else let (l',b) = del_max l in adjust (Node lv l' b r)))"
 | 
|  |     92 | 
 | 
| 62496 |     93 | fun pre_adjust where
 | 
|  |     94 | "pre_adjust (Node lv l a r) = (invar l \<and> invar r \<and>
 | 
|  |     95 |     ((lv = lvl l + 1 \<and> (lv = lvl r + 1 \<or> lv = lvl r + 2 \<or> lv = lvl r \<and> sngl r)) \<or>
 | 
|  |     96 |      (lv = lvl l + 2 \<and> (lv = lvl r + 1 \<or> lv = lvl r \<and> sngl r))))"
 | 
|  |     97 | 
 | 
|  |     98 | declare pre_adjust.simps [simp del]
 | 
|  |     99 | 
 | 
|  |    100 | subsection "Auxiliary Proofs"
 | 
|  |    101 | 
 | 
|  |    102 | lemma split_case: "split t = (case t of
 | 
|  |    103 |   Node lvx a x (Node lvy b y (Node lvz c z d)) \<Rightarrow>
 | 
|  |    104 |    (if lvx = lvy \<and> lvy = lvz
 | 
|  |    105 |     then Node (lvx+1) (Node lvx a x b) y (Node lvx c z d)
 | 
|  |    106 |     else t)
 | 
|  |    107 |   | t \<Rightarrow> t)"
 | 
|  |    108 | by(auto split: tree.split)
 | 
|  |    109 | 
 | 
|  |    110 | lemma skew_case: "skew t = (case t of
 | 
|  |    111 |   Node lvx (Node lvy a y b) x c \<Rightarrow>
 | 
|  |    112 |   (if lvx = lvy then Node lvx a y (Node lvx b x c) else t)
 | 
|  |    113 |  | t \<Rightarrow> t)"
 | 
|  |    114 | by(auto split: tree.split)
 | 
|  |    115 | 
 | 
|  |    116 | lemma lvl_0_iff: "invar t \<Longrightarrow> lvl t = 0 \<longleftrightarrow> t = Leaf"
 | 
|  |    117 | by(cases t) auto
 | 
|  |    118 | 
 | 
|  |    119 | lemma lvl_Suc_iff: "lvl t = Suc n \<longleftrightarrow> (\<exists> l a r. t = Node (Suc n) l a r)"
 | 
|  |    120 | by(cases t) auto
 | 
|  |    121 | 
 | 
|  |    122 | lemma lvl_skew: "lvl (skew t) = lvl t"
 | 
| 62526 |    123 | by(cases t rule: skew.cases) auto
 | 
| 62496 |    124 | 
 | 
|  |    125 | lemma lvl_split: "lvl (split t) = lvl t \<or> lvl (split t) = lvl t + 1 \<and> sngl (split t)"
 | 
| 62526 |    126 | by(cases t rule: split.cases) auto
 | 
| 62496 |    127 | 
 | 
|  |    128 | lemma invar_2Nodes:"invar (Node lv l x (Node rlv rl rx rr)) =
 | 
|  |    129 |      (invar l \<and> invar \<langle>rlv, rl, rx, rr\<rangle> \<and> lv = Suc (lvl l) \<and>
 | 
|  |    130 |      (lv = Suc rlv \<or> rlv = lv \<and> lv = Suc (lvl rr)))"
 | 
|  |    131 | by simp
 | 
|  |    132 | 
 | 
|  |    133 | lemma invar_NodeLeaf[simp]:
 | 
|  |    134 |   "invar (Node lv l x Leaf) = (invar l \<and> lv = Suc (lvl l) \<and> lv = Suc 0)"
 | 
|  |    135 | by simp
 | 
|  |    136 | 
 | 
|  |    137 | lemma sngl_if_invar: "invar (Node n l a r) \<Longrightarrow> n = lvl r \<Longrightarrow> sngl r"
 | 
|  |    138 | by(cases r rule: sngl.cases) clarsimp+
 | 
|  |    139 | 
 | 
|  |    140 | 
 | 
|  |    141 | subsection "Invariance"
 | 
|  |    142 | 
 | 
|  |    143 | subsubsection "Proofs for insert"
 | 
|  |    144 | 
 | 
|  |    145 | lemma lvl_insert_aux:
 | 
|  |    146 |   "lvl (insert x t) = lvl t \<or> lvl (insert x t) = lvl t + 1 \<and> sngl (insert x t)"
 | 
|  |    147 | apply(induction t)
 | 
|  |    148 | apply (auto simp: lvl_skew)
 | 
|  |    149 | apply (metis Suc_eq_plus1 lvl.simps(2) lvl_split lvl_skew)+
 | 
|  |    150 | done
 | 
|  |    151 | 
 | 
|  |    152 | lemma lvl_insert: obtains
 | 
|  |    153 |   (Same) "lvl (insert x t) = lvl t" |
 | 
|  |    154 |   (Incr) "lvl (insert x t) = lvl t + 1" "sngl (insert x t)"
 | 
|  |    155 | using lvl_insert_aux by blast
 | 
|  |    156 | 
 | 
|  |    157 | lemma lvl_insert_sngl: "invar t \<Longrightarrow> sngl t \<Longrightarrow> lvl(insert x t) = lvl t"
 | 
| 62526 |    158 | proof (induction t rule: insert.induct)
 | 
| 62496 |    159 |   case (2 x lv t1 a t2)
 | 
|  |    160 |   consider (LT) "x < a" | (GT) "x > a" | (EQ) "x = a" 
 | 
|  |    161 |     using less_linear by blast 
 | 
|  |    162 |   thus ?case proof cases
 | 
|  |    163 |     case LT
 | 
|  |    164 |     thus ?thesis using 2 by (auto simp add: skew_case split_case split: tree.splits)
 | 
|  |    165 |   next
 | 
|  |    166 |     case GT
 | 
|  |    167 |     thus ?thesis using 2 proof (cases t1)
 | 
|  |    168 |       case Node
 | 
|  |    169 |       thus ?thesis using 2 GT  
 | 
|  |    170 |         apply (auto simp add: skew_case split_case split: tree.splits)
 | 
|  |    171 |         by (metis less_not_refl2 lvl.simps(2) lvl_insert_aux n_not_Suc_n sngl.simps(3))+ 
 | 
|  |    172 |     qed (auto simp add: lvl_0_iff)
 | 
|  |    173 |   qed simp
 | 
|  |    174 | qed simp
 | 
|  |    175 | 
 | 
|  |    176 | lemma skew_invar: "invar t \<Longrightarrow> skew t = t"
 | 
| 62526 |    177 | by(cases t rule: skew.cases) auto
 | 
| 62496 |    178 | 
 | 
|  |    179 | lemma split_invar: "invar t \<Longrightarrow> split t = t"
 | 
| 62526 |    180 | by(cases t rule: split.cases) clarsimp+
 | 
| 62496 |    181 | 
 | 
|  |    182 | lemma invar_NodeL:
 | 
|  |    183 |   "\<lbrakk> invar(Node n l x r); invar l'; lvl l' = lvl l \<rbrakk> \<Longrightarrow> invar(Node n l' x r)"
 | 
|  |    184 | by(auto)
 | 
|  |    185 | 
 | 
|  |    186 | lemma invar_NodeR:
 | 
|  |    187 |   "\<lbrakk> invar(Node n l x r); n = lvl r + 1; invar r'; lvl r' = lvl r \<rbrakk> \<Longrightarrow> invar(Node n l x r')"
 | 
|  |    188 | by(auto)
 | 
|  |    189 | 
 | 
|  |    190 | lemma invar_NodeR2:
 | 
|  |    191 |   "\<lbrakk> invar(Node n l x r); sngl r'; n = lvl r + 1; invar r'; lvl r' = n \<rbrakk> \<Longrightarrow> invar(Node n l x r')"
 | 
|  |    192 | by(cases r' rule: sngl.cases) clarsimp+
 | 
|  |    193 | 
 | 
|  |    194 | 
 | 
|  |    195 | lemma lvl_insert_incr_iff: "(lvl(insert a t) = lvl t + 1) \<longleftrightarrow>
 | 
|  |    196 |   (EX l x r. insert a t = Node (lvl t + 1) l x r \<and> lvl l = lvl r)"
 | 
|  |    197 | apply(cases t)
 | 
|  |    198 | apply(auto simp add: skew_case split_case split: if_splits)
 | 
|  |    199 | apply(auto split: tree.splits if_splits)
 | 
|  |    200 | done
 | 
|  |    201 | 
 | 
|  |    202 | lemma invar_insert: "invar t \<Longrightarrow> invar(insert a t)"
 | 
|  |    203 | proof(induction t)
 | 
|  |    204 |   case (Node n l x r)
 | 
|  |    205 |   hence il: "invar l" and ir: "invar r" by auto
 | 
|  |    206 |   note N = Node
 | 
|  |    207 |   let ?t = "Node n l x r"
 | 
|  |    208 |   have "a < x \<or> a = x \<or> x < a" by auto
 | 
|  |    209 |   moreover
 | 
|  |    210 |   { assume "a < x"
 | 
|  |    211 |     note iil = Node.IH(1)[OF il]
 | 
|  |    212 |     have ?case
 | 
|  |    213 |     proof (cases rule: lvl_insert[of a l])
 | 
|  |    214 |       case (Same) thus ?thesis
 | 
|  |    215 |         using \<open>a<x\<close> invar_NodeL[OF Node.prems iil Same]
 | 
|  |    216 |         by (simp add: skew_invar split_invar del: invar.simps)
 | 
|  |    217 |     next
 | 
|  |    218 |       case (Incr)
 | 
|  |    219 |       then obtain t1 w t2 where ial[simp]: "insert a l = Node n t1 w t2"
 | 
|  |    220 |         using Node.prems by (auto simp: lvl_Suc_iff)
 | 
|  |    221 |       have l12: "lvl t1 = lvl t2"
 | 
|  |    222 |         by (metis Incr(1) ial lvl_insert_incr_iff tree.inject)
 | 
|  |    223 |       have "insert a ?t = split(skew(Node n (insert a l) x r))"
 | 
|  |    224 |         by(simp add: \<open>a<x\<close>)
 | 
|  |    225 |       also have "skew(Node n (insert a l) x r) = Node n t1 w (Node n t2 x r)"
 | 
|  |    226 |         by(simp)
 | 
|  |    227 |       also have "invar(split \<dots>)"
 | 
|  |    228 |       proof (cases r)
 | 
|  |    229 |         case Leaf
 | 
|  |    230 |         hence "l = Leaf" using Node.prems by(auto simp: lvl_0_iff)
 | 
|  |    231 |         thus ?thesis using Leaf ial by simp
 | 
|  |    232 |       next
 | 
|  |    233 |         case [simp]: (Node m t3 y t4)
 | 
|  |    234 |         show ?thesis (*using N(3) iil l12 by(auto)*)
 | 
|  |    235 |         proof cases
 | 
|  |    236 |           assume "m = n" thus ?thesis using N(3) iil by(auto)
 | 
|  |    237 |         next
 | 
|  |    238 |           assume "m \<noteq> n" thus ?thesis using N(3) iil l12 by(auto)
 | 
|  |    239 |         qed
 | 
|  |    240 |       qed
 | 
|  |    241 |       finally show ?thesis .
 | 
|  |    242 |     qed
 | 
|  |    243 |   }
 | 
|  |    244 |   moreover
 | 
|  |    245 |   { assume "x < a"
 | 
|  |    246 |     note iir = Node.IH(2)[OF ir]
 | 
|  |    247 |     from \<open>invar ?t\<close> have "n = lvl r \<or> n = lvl r + 1" by auto
 | 
|  |    248 |     hence ?case
 | 
|  |    249 |     proof
 | 
|  |    250 |       assume 0: "n = lvl r"
 | 
|  |    251 |       have "insert a ?t = split(skew(Node n l x (insert a r)))"
 | 
|  |    252 |         using \<open>a>x\<close> by(auto)
 | 
|  |    253 |       also have "skew(Node n l x (insert a r)) = Node n l x (insert a r)"
 | 
|  |    254 |         using Node.prems by(simp add: skew_case split: tree.split)
 | 
|  |    255 |       also have "invar(split \<dots>)"
 | 
|  |    256 |       proof -
 | 
|  |    257 |         from lvl_insert_sngl[OF ir sngl_if_invar[OF \<open>invar ?t\<close> 0], of a]
 | 
|  |    258 |         obtain t1 y t2 where iar: "insert a r = Node n t1 y t2"
 | 
|  |    259 |           using Node.prems 0 by (auto simp: lvl_Suc_iff)
 | 
|  |    260 |         from Node.prems iar 0 iir
 | 
|  |    261 |         show ?thesis by (auto simp: split_case split: tree.splits)
 | 
|  |    262 |       qed
 | 
|  |    263 |       finally show ?thesis .
 | 
|  |    264 |     next
 | 
|  |    265 |       assume 1: "n = lvl r + 1"
 | 
|  |    266 |       hence "sngl ?t" by(cases r) auto
 | 
|  |    267 |       show ?thesis
 | 
|  |    268 |       proof (cases rule: lvl_insert[of a r])
 | 
|  |    269 |         case (Same)
 | 
|  |    270 |         show ?thesis using \<open>x<a\<close> il ir invar_NodeR[OF Node.prems 1 iir Same]
 | 
|  |    271 |           by (auto simp add: skew_invar split_invar)
 | 
|  |    272 |       next
 | 
|  |    273 |         case (Incr)
 | 
|  |    274 |         thus ?thesis using invar_NodeR2[OF `invar ?t` Incr(2) 1 iir] 1 \<open>x < a\<close>
 | 
|  |    275 |           by (auto simp add: skew_invar split_invar split: if_splits)
 | 
|  |    276 |       qed
 | 
|  |    277 |     qed
 | 
|  |    278 |   }
 | 
|  |    279 |   moreover { assume "a = x" hence ?case using Node.prems by auto }
 | 
|  |    280 |   ultimately show ?case by blast
 | 
|  |    281 | qed simp
 | 
|  |    282 | 
 | 
|  |    283 | 
 | 
|  |    284 | subsubsection "Proofs for delete"
 | 
|  |    285 | 
 | 
|  |    286 | lemma invarL: "ASSUMPTION(invar \<langle>lv, l, a, r\<rangle>) \<Longrightarrow> invar l"
 | 
|  |    287 | by(simp add: ASSUMPTION_def)
 | 
|  |    288 | 
 | 
|  |    289 | lemma invarR: "ASSUMPTION(invar \<langle>lv, l, a, r\<rangle>) \<Longrightarrow> invar r"
 | 
|  |    290 | by(simp add: ASSUMPTION_def)
 | 
|  |    291 | 
 | 
|  |    292 | lemma sngl_NodeI:
 | 
|  |    293 |   "sngl (Node lv l a r) \<Longrightarrow> sngl (Node lv l' a' r)"
 | 
|  |    294 | by(cases r) (simp_all)
 | 
|  |    295 | 
 | 
|  |    296 | 
 | 
|  |    297 | declare invarL[simp] invarR[simp]
 | 
|  |    298 | 
 | 
|  |    299 | lemma pre_cases:
 | 
|  |    300 | assumes "pre_adjust (Node lv l x r)"
 | 
|  |    301 | obtains
 | 
|  |    302 |  (tSngl) "invar l \<and> invar r \<and>
 | 
|  |    303 |     lv = Suc (lvl r) \<and> lvl l = lvl r" |
 | 
|  |    304 |  (tDouble) "invar l \<and> invar r \<and>
 | 
|  |    305 |     lv = lvl r \<and> Suc (lvl l) = lvl r \<and> sngl r " |
 | 
|  |    306 |  (rDown) "invar l \<and> invar r \<and>
 | 
|  |    307 |     lv = Suc (Suc (lvl r)) \<and>  lv = Suc (lvl l)" |
 | 
|  |    308 |  (lDown_tSngl) "invar l \<and> invar r \<and>
 | 
|  |    309 |     lv = Suc (lvl r) \<and> lv = Suc (Suc (lvl l))" |
 | 
|  |    310 |  (lDown_tDouble) "invar l \<and> invar r \<and>
 | 
|  |    311 |     lv = lvl r \<and> lv = Suc (Suc (lvl l)) \<and> sngl r"
 | 
|  |    312 | using assms unfolding pre_adjust.simps
 | 
|  |    313 | by auto
 | 
|  |    314 | 
 | 
|  |    315 | declare invar.simps(2)[simp del] invar_2Nodes[simp add]
 | 
|  |    316 | 
 | 
|  |    317 | lemma invar_adjust:
 | 
|  |    318 |   assumes pre: "pre_adjust (Node lv l a r)"
 | 
|  |    319 |   shows  "invar(adjust (Node lv l a r))"
 | 
|  |    320 | using pre proof (cases rule: pre_cases)
 | 
|  |    321 |   case (tDouble) thus ?thesis unfolding adjust_def by (cases r) (auto simp: invar.simps(2)) 
 | 
|  |    322 | next 
 | 
|  |    323 |   case (rDown)
 | 
|  |    324 |   from rDown obtain llv ll la lr where l: "l = Node llv ll la lr" by (cases l) auto
 | 
|  |    325 |   from rDown show ?thesis unfolding adjust_def by (auto simp: l invar.simps(2) split: tree.splits)
 | 
|  |    326 | next
 | 
|  |    327 |   case (lDown_tDouble)
 | 
|  |    328 |   from lDown_tDouble obtain rlv rr ra rl where r: "r = Node rlv rl ra rr" by (cases r) auto
 | 
|  |    329 |   from lDown_tDouble and r obtain rrlv rrr rra rrl where
 | 
|  |    330 |     rr :"rr = Node rrlv rrr rra rrl" by (cases rr) auto
 | 
|  |    331 |   from  lDown_tDouble show ?thesis unfolding adjust_def r rr
 | 
|  |    332 |     apply (cases rl) apply (auto simp add: invar.simps(2))
 | 
|  |    333 |     using lDown_tDouble by (auto simp: split_case lvl_0_iff  elim:lvl.elims split: tree.split)
 | 
|  |    334 | qed (auto simp: split_case invar.simps(2) adjust_def split: tree.splits)
 | 
|  |    335 | 
 | 
|  |    336 | lemma lvl_adjust:
 | 
|  |    337 |   assumes "pre_adjust (Node lv l a r)"
 | 
|  |    338 |   shows "lv = lvl (adjust(Node lv l a r)) \<or> lv = lvl (adjust(Node lv l a r)) + 1"
 | 
|  |    339 | using assms(1) proof(cases rule: pre_cases)
 | 
|  |    340 |   case lDown_tSngl thus ?thesis
 | 
|  |    341 |     using lvl_split[of "\<langle>lvl r, l, a, r\<rangle>"] by (auto simp: adjust_def)
 | 
|  |    342 | next
 | 
|  |    343 |   case lDown_tDouble thus ?thesis
 | 
|  |    344 |     by (auto simp: adjust_def invar.simps(2) split: tree.split)
 | 
|  |    345 | qed (auto simp: adjust_def split: tree.splits)
 | 
|  |    346 | 
 | 
|  |    347 | lemma sngl_adjust: assumes "pre_adjust (Node lv l a r)"
 | 
|  |    348 |   "sngl \<langle>lv, l, a, r\<rangle>" "lv = lvl (adjust \<langle>lv, l, a, r\<rangle>)"
 | 
|  |    349 |   shows "sngl (adjust \<langle>lv, l, a, r\<rangle>)" 
 | 
|  |    350 | using assms proof (cases rule: pre_cases)
 | 
|  |    351 |   case rDown
 | 
|  |    352 |   thus ?thesis using assms(2,3) unfolding adjust_def
 | 
|  |    353 |     by (auto simp add: skew_case) (auto split: tree.split)
 | 
|  |    354 | qed (auto simp: adjust_def skew_case split_case split: tree.split)
 | 
|  |    355 | 
 | 
|  |    356 | definition "post_del t t' ==
 | 
|  |    357 |   invar t' \<and>
 | 
|  |    358 |   (lvl t' = lvl t \<or> lvl t' + 1 = lvl t) \<and>
 | 
|  |    359 |   (lvl t' = lvl t \<and> sngl t \<longrightarrow> sngl t')"
 | 
|  |    360 | 
 | 
|  |    361 | lemma pre_adj_if_postR:
 | 
|  |    362 |   "invar\<langle>lv, l, a, r\<rangle> \<Longrightarrow> post_del r r' \<Longrightarrow> pre_adjust \<langle>lv, l, a, r'\<rangle>"
 | 
|  |    363 | by(cases "sngl r")
 | 
|  |    364 |   (auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims)
 | 
|  |    365 | 
 | 
|  |    366 | lemma pre_adj_if_postL:
 | 
|  |    367 |   "invar\<langle>lv, l, a, r\<rangle> \<Longrightarrow> post_del l l' \<Longrightarrow> pre_adjust \<langle>lv, l', b, r\<rangle>"
 | 
|  |    368 | by(cases "sngl r")
 | 
|  |    369 |   (auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims)
 | 
|  |    370 | 
 | 
|  |    371 | lemma post_del_adjL:
 | 
|  |    372 |   "\<lbrakk> invar\<langle>lv, l, a, r\<rangle>; pre_adjust \<langle>lv, l', b, r\<rangle> \<rbrakk>
 | 
|  |    373 |   \<Longrightarrow> post_del \<langle>lv, l, a, r\<rangle> (adjust \<langle>lv, l', b, r\<rangle>)"
 | 
|  |    374 | unfolding post_del_def
 | 
|  |    375 | by (metis invar_adjust lvl_adjust sngl_NodeI sngl_adjust lvl.simps(2))
 | 
|  |    376 | 
 | 
|  |    377 | lemma post_del_adjR:
 | 
|  |    378 | assumes "invar\<langle>lv, l, a, r\<rangle>" "pre_adjust \<langle>lv, l, a, r'\<rangle>" "post_del r r'"
 | 
|  |    379 | shows "post_del \<langle>lv, l, a, r\<rangle> (adjust \<langle>lv, l, a, r'\<rangle>)"
 | 
|  |    380 | proof(unfold post_del_def, safe del: disjCI)
 | 
|  |    381 |   let ?t = "\<langle>lv, l, a, r\<rangle>"
 | 
|  |    382 |   let ?t' = "adjust \<langle>lv, l, a, r'\<rangle>"
 | 
|  |    383 |   show "invar ?t'" by(rule invar_adjust[OF assms(2)])
 | 
|  |    384 |   show "lvl ?t' = lvl ?t \<or> lvl ?t' + 1 = lvl ?t"
 | 
|  |    385 |     using lvl_adjust[OF assms(2)] by auto
 | 
|  |    386 |   show "sngl ?t'" if as: "lvl ?t' = lvl ?t" "sngl ?t"
 | 
|  |    387 |   proof -
 | 
|  |    388 |     have s: "sngl \<langle>lv, l, a, r'\<rangle>"
 | 
|  |    389 |     proof(cases r')
 | 
|  |    390 |       case Leaf thus ?thesis by simp
 | 
|  |    391 |     next
 | 
|  |    392 |       case Node thus ?thesis using as(2) assms(1,3)
 | 
|  |    393 |       by (cases r) (auto simp: post_del_def)
 | 
|  |    394 |     qed
 | 
|  |    395 |     show ?thesis using as(1) sngl_adjust[OF assms(2) s] by simp
 | 
|  |    396 |   qed
 | 
|  |    397 | qed
 | 
|  |    398 | 
 | 
|  |    399 | declare prod.splits[split]
 | 
|  |    400 | 
 | 
|  |    401 | theorem post_del_max:
 | 
|  |    402 |  "\<lbrakk> invar t; (t', x) = del_max t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow> post_del t t'"
 | 
|  |    403 | proof (induction t arbitrary: t' rule: del_max.induct)
 | 
|  |    404 |   case (2 lv l a lvr rl ra rr)
 | 
|  |    405 |   let ?r =  "\<langle>lvr, rl, ra, rr\<rangle>"
 | 
|  |    406 |   let ?t = "\<langle>lv, l, a, ?r\<rangle>"
 | 
|  |    407 |   from "2.prems"(2) obtain r' where r': "(r', x) = del_max ?r"
 | 
|  |    408 |     and [simp]: "t' = adjust \<langle>lv, l, a, r'\<rangle>" by auto
 | 
|  |    409 |   from  "2.IH"[OF _ r'] \<open>invar ?t\<close> have post: "post_del ?r r'" by simp
 | 
|  |    410 |   note preR = pre_adj_if_postR[OF \<open>invar ?t\<close> post]
 | 
|  |    411 |   show ?case by (simp add: post_del_adjR[OF "2.prems"(1) preR post])
 | 
|  |    412 | qed (auto simp: post_del_def)
 | 
|  |    413 | 
 | 
|  |    414 | theorem post_delete: "invar t \<Longrightarrow> post_del t (delete x t)"
 | 
|  |    415 | proof (induction t)
 | 
|  |    416 |   case (Node lv l a r)
 | 
|  |    417 | 
 | 
|  |    418 |   let ?l' = "delete x l" and ?r' = "delete x r"
 | 
|  |    419 |   let ?t = "Node lv l a r" let ?t' = "delete x ?t"
 | 
|  |    420 | 
 | 
|  |    421 |   from Node.prems have inv_l: "invar l" and inv_r: "invar r" by (auto)
 | 
|  |    422 | 
 | 
|  |    423 |   note post_l' = Node.IH(1)[OF inv_l]
 | 
|  |    424 |   note preL = pre_adj_if_postL[OF Node.prems post_l']
 | 
|  |    425 | 
 | 
|  |    426 |   note post_r' = Node.IH(2)[OF inv_r]
 | 
|  |    427 |   note preR = pre_adj_if_postR[OF Node.prems post_r']
 | 
|  |    428 | 
 | 
|  |    429 |   show ?case
 | 
|  |    430 |   proof (cases rule: linorder_cases[of x a])
 | 
|  |    431 |     case less
 | 
|  |    432 |     thus ?thesis using Node.prems by (simp add: post_del_adjL preL)
 | 
|  |    433 |   next
 | 
|  |    434 |     case greater
 | 
|  |    435 |     thus ?thesis using Node.prems by (simp add: post_del_adjR preR post_r')
 | 
|  |    436 |   next
 | 
|  |    437 |     case equal
 | 
|  |    438 |     show ?thesis
 | 
|  |    439 |     proof cases
 | 
|  |    440 |       assume "l = Leaf" thus ?thesis using equal Node.prems
 | 
|  |    441 |         by(auto simp: post_del_def invar.simps(2))
 | 
|  |    442 |     next
 | 
|  |    443 |       assume "l \<noteq> Leaf" thus ?thesis using equal
 | 
|  |    444 |         by simp (metis Node.prems inv_l post_del_adjL post_del_max pre_adj_if_postL)
 | 
|  |    445 |     qed
 | 
|  |    446 |   qed
 | 
|  |    447 | qed (simp add: post_del_def)
 | 
|  |    448 | 
 | 
|  |    449 | declare invar_2Nodes[simp del]
 | 
|  |    450 | 
 | 
| 61793 |    451 | 
 | 
|  |    452 | subsection "Functional Correctness"
 | 
|  |    453 | 
 | 
| 62496 |    454 | 
 | 
| 61793 |    455 | subsubsection "Proofs for insert"
 | 
|  |    456 | 
 | 
|  |    457 | lemma inorder_split: "inorder(split t) = inorder t"
 | 
|  |    458 | by(cases t rule: split.cases) (auto)
 | 
|  |    459 | 
 | 
|  |    460 | lemma inorder_skew: "inorder(skew t) = inorder t"
 | 
|  |    461 | by(cases t rule: skew.cases) (auto)
 | 
|  |    462 | 
 | 
|  |    463 | lemma inorder_insert:
 | 
|  |    464 |   "sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)"
 | 
|  |    465 | by(induction t) (auto simp: ins_list_simps inorder_split inorder_skew)
 | 
|  |    466 | 
 | 
| 62496 |    467 | 
 | 
| 61793 |    468 | subsubsection "Proofs for delete"
 | 
|  |    469 | 
 | 
| 62496 |    470 | lemma inorder_adjust: "t \<noteq> Leaf \<Longrightarrow> pre_adjust t \<Longrightarrow> inorder(adjust t) = inorder t"
 | 
| 62526 |    471 | by(cases t)
 | 
| 62496 |    472 |   (auto simp: adjust_def inorder_skew inorder_split invar.simps(2) pre_adjust.simps
 | 
|  |    473 |      split: tree.splits)
 | 
|  |    474 | 
 | 
| 61793 |    475 | lemma del_maxD:
 | 
| 62496 |    476 |   "\<lbrakk> del_max t = (t',x); t \<noteq> Leaf; invar t \<rbrakk> \<Longrightarrow> inorder t' @ [x] = inorder t"
 | 
| 61793 |    477 | by(induction t arbitrary: t' rule: del_max.induct)
 | 
| 62496 |    478 |   (auto simp: sorted_lems inorder_adjust pre_adj_if_postR post_del_max split: prod.splits)
 | 
| 61793 |    479 | 
 | 
|  |    480 | lemma inorder_delete:
 | 
| 62496 |    481 |   "invar t \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
 | 
| 61793 |    482 | by(induction t)
 | 
| 62496 |    483 |   (auto simp: del_list_simps inorder_adjust pre_adj_if_postL pre_adj_if_postR 
 | 
|  |    484 |               post_del_max post_delete del_maxD split: prod.splits)
 | 
| 61793 |    485 | 
 | 
| 62496 |    486 | interpretation I: Set_by_Ordered
 | 
| 61793 |    487 | where empty = Leaf and isin = isin and insert = insert and delete = delete
 | 
| 62496 |    488 | and inorder = inorder and inv = invar
 | 
| 61793 |    489 | proof (standard, goal_cases)
 | 
|  |    490 |   case 1 show ?case by simp
 | 
|  |    491 | next
 | 
|  |    492 |   case 2 thus ?case by(simp add: isin_set)
 | 
|  |    493 | next
 | 
|  |    494 |   case 3 thus ?case by(simp add: inorder_insert)
 | 
|  |    495 | next
 | 
|  |    496 |   case 4 thus ?case by(simp add: inorder_delete)
 | 
| 62496 |    497 | next
 | 
|  |    498 |   case 5 thus ?case by(simp)
 | 
|  |    499 | next
 | 
|  |    500 |   case 6 thus ?case by(simp add: invar_insert)
 | 
|  |    501 | next
 | 
|  |    502 |   case 7 thus ?case using post_delete by(auto simp: post_del_def)
 | 
|  |    503 | qed
 | 
| 61793 |    504 | 
 | 
| 62390 |    505 | end
 |