| 71844 |      1 | (* Author: Tobias Nipkow *)
 | 
| 71814 |      2 | 
 | 
| 71844 |      3 | section "AVL Tree with Balance Factors (2)"
 | 
| 71814 |      4 | 
 | 
| 71844 |      5 | theory AVL_Bal2_Set
 | 
| 71814 |      6 | imports
 | 
|  |      7 |   Cmp
 | 
|  |      8 |   Isin2
 | 
|  |      9 | begin
 | 
|  |     10 | 
 | 
| 71844 |     11 | text \<open>This version passes a flag (\<open>Same\<close>/\<open>Diff\<close>) back up to signal if the height changed.\<close>
 | 
|  |     12 | 
 | 
| 71814 |     13 | datatype bal = Lh | Bal | Rh
 | 
|  |     14 | 
 | 
|  |     15 | type_synonym 'a tree_bal = "('a * bal) tree"
 | 
|  |     16 | 
 | 
|  |     17 | text \<open>Invariant:\<close>
 | 
|  |     18 | 
 | 
|  |     19 | fun avl :: "'a tree_bal \<Rightarrow> bool" where
 | 
|  |     20 | "avl Leaf = True" |
 | 
|  |     21 | "avl (Node l (a,b) r) =
 | 
|  |     22 |   ((case b of
 | 
|  |     23 |     Bal \<Rightarrow> height r = height l |
 | 
|  |     24 |     Lh \<Rightarrow> height l = height r + 1 |
 | 
|  |     25 |     Rh \<Rightarrow> height r = height l + 1)
 | 
|  |     26 |   \<and> avl l \<and> avl r)"
 | 
|  |     27 | 
 | 
|  |     28 | 
 | 
|  |     29 | subsection \<open>Code\<close>
 | 
|  |     30 | 
 | 
| 71828 |     31 | datatype 'a alt = Same 'a | Diff 'a
 | 
| 71814 |     32 | 
 | 
| 71828 |     33 | type_synonym 'a tree_bal2 = "'a tree_bal alt"
 | 
|  |     34 | 
 | 
|  |     35 | fun tree :: "'a alt \<Rightarrow> 'a" where
 | 
| 71814 |     36 | "tree(Same t) = t" |
 | 
|  |     37 | "tree(Diff t) = t"
 | 
|  |     38 | 
 | 
| 71820 |     39 | fun rot2 where
 | 
|  |     40 | "rot2 A a B c C = (case B of
 | 
|  |     41 |   (Node B\<^sub>1 (b, bb) B\<^sub>2) \<Rightarrow>
 | 
|  |     42 |     let b\<^sub>1 = if bb = Rh then Lh else Bal;
 | 
|  |     43 |         b\<^sub>2 = if bb = Lh then Rh else Bal
 | 
|  |     44 |     in Node (Node A (a,b\<^sub>1) B\<^sub>1) (b,Bal) (Node B\<^sub>2 (c,b\<^sub>2) C))"
 | 
| 71814 |     45 | 
 | 
| 71818 |     46 | fun balL :: "'a tree_bal2 \<Rightarrow> 'a \<Rightarrow> bal \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal2" where
 | 
| 71815 |     47 | "balL AB' c bc C = (case AB' of
 | 
|  |     48 |    Same AB \<Rightarrow> Same (Node AB (c,bc) C) |
 | 
|  |     49 |    Diff AB \<Rightarrow> (case bc of
 | 
|  |     50 |      Bal \<Rightarrow> Diff (Node AB (c,Lh) C) |
 | 
|  |     51 |      Rh \<Rightarrow> Same (Node AB (c,Bal) C) |
 | 
| 71824 |     52 |      Lh \<Rightarrow> (case AB of
 | 
|  |     53 |        Node A (a,Lh) B \<Rightarrow> Same(Node A (a,Bal) (Node B (c,Bal) C)) |
 | 
|  |     54 |        Node A (a,Rh) B \<Rightarrow> Same(rot2 A a B c C))))"
 | 
| 71814 |     55 | 
 | 
| 71818 |     56 | fun balR :: "'a tree_bal \<Rightarrow> 'a \<Rightarrow> bal \<Rightarrow> 'a tree_bal2 \<Rightarrow> 'a tree_bal2" where
 | 
| 71815 |     57 | "balR A a ba BC' = (case BC' of
 | 
|  |     58 |    Same BC \<Rightarrow> Same (Node A (a,ba) BC) |
 | 
|  |     59 |    Diff BC \<Rightarrow> (case ba of
 | 
|  |     60 |      Bal \<Rightarrow> Diff (Node A (a,Rh) BC) |
 | 
|  |     61 |      Lh \<Rightarrow> Same (Node A (a,Bal) BC) |
 | 
| 71824 |     62 |      Rh \<Rightarrow> (case BC of
 | 
|  |     63 |        Node B (c,Rh) C \<Rightarrow> Same(Node (Node A (a,Bal) B) (c,Bal) C) |
 | 
|  |     64 |        Node B (c,Lh) C \<Rightarrow> Same(rot2 A a B c C))))"
 | 
| 71814 |     65 | 
 | 
| 71828 |     66 | fun ins :: "'a::linorder \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal2" where
 | 
|  |     67 | "ins x Leaf = Diff(Node Leaf (x, Bal) Leaf)" |
 | 
|  |     68 | "ins x (Node l (a, b) r) = (case cmp x a of
 | 
| 71814 |     69 |    EQ \<Rightarrow> Same(Node l (a, b) r) |
 | 
| 71828 |     70 |    LT \<Rightarrow> balL (ins x l) a b r |
 | 
|  |     71 |    GT \<Rightarrow> balR l a b (ins x r))"
 | 
|  |     72 | 
 | 
|  |     73 | definition insert :: "'a::linorder \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal" where
 | 
|  |     74 | "insert x t = tree(ins x t)"
 | 
| 71814 |     75 | 
 | 
| 71818 |     76 | fun baldR :: "'a tree_bal \<Rightarrow> 'a \<Rightarrow> bal \<Rightarrow> 'a tree_bal2 \<Rightarrow> 'a tree_bal2" where
 | 
| 71816 |     77 | "baldR AB c bc C' = (case C' of
 | 
|  |     78 |    Same C \<Rightarrow> Same (Node AB (c,bc) C) |
 | 
|  |     79 |    Diff C \<Rightarrow> (case bc of
 | 
|  |     80 |      Bal \<Rightarrow> Same (Node AB (c,Lh) C) |
 | 
|  |     81 |      Rh \<Rightarrow> Diff (Node AB (c,Bal) C) |
 | 
| 71814 |     82 |      Lh \<Rightarrow> (case AB of
 | 
| 71816 |     83 |        Node A (a,Lh) B \<Rightarrow> Diff(Node A (a,Bal) (Node B (c,Bal) C)) |
 | 
|  |     84 |        Node A (a,Bal) B \<Rightarrow> Same(Node A (a,Rh) (Node B (c,Lh) C)) |
 | 
| 71820 |     85 |        Node A (a,Rh) B \<Rightarrow> Diff(rot2 A a B c C))))"
 | 
| 71814 |     86 | 
 | 
| 71818 |     87 | fun baldL :: "'a tree_bal2 \<Rightarrow> 'a \<Rightarrow> bal \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal2" where
 | 
| 71816 |     88 | "baldL A' a ba BC = (case A' of
 | 
|  |     89 |    Same A \<Rightarrow> Same (Node A (a,ba) BC) |
 | 
|  |     90 |    Diff A \<Rightarrow> (case ba of
 | 
|  |     91 |      Bal \<Rightarrow> Same (Node A (a,Rh) BC) |
 | 
|  |     92 |      Lh \<Rightarrow> Diff (Node A (a,Bal) BC) |
 | 
| 71814 |     93 |      Rh \<Rightarrow> (case BC of
 | 
| 71816 |     94 |        Node B (c,Rh) C \<Rightarrow> Diff(Node (Node A (a,Bal) B) (c,Bal) C) |
 | 
|  |     95 |        Node B (c,Bal) C \<Rightarrow> Same(Node (Node A (a,Rh) B) (c,Lh) C) |
 | 
| 71820 |     96 |        Node B (c,Lh) C \<Rightarrow> Diff(rot2 A a B c C))))"
 | 
| 71814 |     97 | 
 | 
| 71818 |     98 | fun split_max :: "'a tree_bal \<Rightarrow> 'a tree_bal2 * 'a" where
 | 
| 71814 |     99 | "split_max (Node l (a, ba) r) =
 | 
|  |    100 |   (if r = Leaf then (Diff l,a) else let (r',a') = split_max r in (baldR l a ba r', a'))"
 | 
|  |    101 | 
 | 
| 71828 |    102 | fun del :: "'a::linorder \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal2" where
 | 
|  |    103 | "del _ Leaf = Same Leaf" |
 | 
|  |    104 | "del x (Node l (a, ba) r) =
 | 
| 71814 |    105 |   (case cmp x a of
 | 
|  |    106 |      EQ \<Rightarrow> if l = Leaf then Diff r
 | 
|  |    107 |            else let (l', a') = split_max l in baldL l' a' ba r |
 | 
| 71828 |    108 |      LT \<Rightarrow> baldL (del x l) a ba r |
 | 
|  |    109 |      GT \<Rightarrow> baldR l a ba (del x r))"
 | 
|  |    110 | 
 | 
|  |    111 | definition delete :: "'a::linorder \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal" where
 | 
|  |    112 | "delete x t = tree(del x t)"
 | 
| 71814 |    113 | 
 | 
|  |    114 | lemmas split_max_induct = split_max.induct[case_names Node Leaf]
 | 
|  |    115 | 
 | 
| 71828 |    116 | lemmas splits = if_splits tree.splits alt.splits bal.splits
 | 
| 71814 |    117 | 
 | 
|  |    118 | subsection \<open>Proofs\<close>
 | 
|  |    119 | 
 | 
| 71828 |    120 | subsubsection "Proofs about insertion"
 | 
| 71814 |    121 | 
 | 
| 71828 |    122 | lemma avl_ins_case: "avl t \<Longrightarrow> case ins x t of
 | 
| 71814 |    123 |    Same t' \<Rightarrow> avl t' \<and> height t' = height t |
 | 
| 71824 |    124 |    Diff t' \<Rightarrow> avl t' \<and> height t' = height t + 1 \<and>
 | 
|  |    125 |       (\<forall>l a r. t' = Node l (a,Bal) r \<longrightarrow> a = x \<and> l = Leaf \<and> r = Leaf)"
 | 
| 71828 |    126 | apply(induction x t rule: ins.induct)
 | 
| 71814 |    127 | apply(auto simp: max_absorb1 split!: splits)
 | 
|  |    128 | done
 | 
|  |    129 | 
 | 
| 71828 |    130 | corollary avl_insert: "avl t \<Longrightarrow> avl(insert x t)"
 | 
|  |    131 | using avl_ins_case[of t x] by (simp add: insert_def split: splits)
 | 
| 71814 |    132 | 
 | 
| 71828 |    133 | (* The following aux lemma simplifies the inorder_ins proof *)
 | 
| 71824 |    134 | 
 | 
| 71828 |    135 | lemma ins_Diff[simp]: "avl t \<Longrightarrow>
 | 
|  |    136 |   ins x t \<noteq> Diff Leaf \<and>
 | 
|  |    137 |   (ins x t = Diff (Node l (a,Bal) r) \<longleftrightarrow> t = Leaf \<and> a = x \<and> l=Leaf \<and> r=Leaf) \<and>
 | 
|  |    138 |   ins x t \<noteq> Diff (Node l (a,Rh) Leaf) \<and>
 | 
|  |    139 |   ins x t \<noteq> Diff (Node Leaf (a,Lh) r)"
 | 
|  |    140 | by(drule avl_ins_case[of _ x]) (auto split: splits)
 | 
| 71824 |    141 | 
 | 
| 71828 |    142 | theorem inorder_ins:
 | 
|  |    143 |   "\<lbrakk> avl t;  sorted(inorder t) \<rbrakk> \<Longrightarrow> inorder(tree(ins x t)) = ins_list x (inorder t)"
 | 
| 71824 |    144 | apply(induction t)
 | 
|  |    145 | apply (auto simp: ins_list_simps  split!: splits)
 | 
|  |    146 | done
 | 
|  |    147 | 
 | 
| 71814 |    148 | 
 | 
| 71828 |    149 | subsubsection "Proofs about deletion"
 | 
| 71814 |    150 | 
 | 
|  |    151 | lemma inorder_baldL:
 | 
|  |    152 |   "\<lbrakk> ba = Rh \<longrightarrow> r \<noteq> Leaf; avl r \<rbrakk>
 | 
|  |    153 |   \<Longrightarrow> inorder (tree(baldL l a ba r)) = inorder (tree l) @ a # inorder r"
 | 
|  |    154 | by (auto split: splits)
 | 
|  |    155 | 
 | 
|  |    156 | lemma inorder_baldR:
 | 
|  |    157 |   "\<lbrakk> ba = Lh \<longrightarrow> l \<noteq> Leaf; avl l \<rbrakk>
 | 
|  |    158 |    \<Longrightarrow> inorder (tree(baldR l a ba r)) = inorder l @ a # inorder (tree r)"
 | 
|  |    159 | by (auto split: splits)
 | 
|  |    160 | 
 | 
|  |    161 | lemma avl_split_max:
 | 
|  |    162 |   "\<lbrakk> split_max t = (t',a); avl t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow> case t' of
 | 
|  |    163 |    Same t' \<Rightarrow> avl t' \<and> height t = height t' |
 | 
|  |    164 |    Diff t' \<Rightarrow> avl t' \<and> height t = height t' + 1"
 | 
|  |    165 | apply(induction t arbitrary: t' a rule: split_max_induct)
 | 
|  |    166 |  apply(fastforce simp: max_absorb1 max_absorb2 split!: splits prod.splits)
 | 
|  |    167 | apply simp
 | 
|  |    168 | done
 | 
|  |    169 | 
 | 
| 71828 |    170 | lemma avl_del_case: "avl t \<Longrightarrow> case del x t of
 | 
| 71814 |    171 |    Same t' \<Rightarrow> avl t' \<and> height t = height t' |
 | 
|  |    172 |    Diff t' \<Rightarrow> avl t' \<and> height t = height t' + 1"
 | 
| 71828 |    173 | apply(induction x t rule: del.induct)
 | 
| 71814 |    174 |  apply(auto simp: max_absorb1 max_absorb2 dest: avl_split_max split!: splits prod.splits)
 | 
|  |    175 | done
 | 
|  |    176 | 
 | 
| 71828 |    177 | corollary avl_delete: "avl t \<Longrightarrow> avl(delete x t)"
 | 
|  |    178 | using avl_del_case[of t x] by(simp add: delete_def split: splits)
 | 
| 71814 |    179 | 
 | 
|  |    180 | lemma inorder_split_maxD:
 | 
|  |    181 |   "\<lbrakk> split_max t = (t',a); t \<noteq> Leaf; avl t \<rbrakk> \<Longrightarrow>
 | 
|  |    182 |    inorder (tree t') @ [a] = inorder t"
 | 
|  |    183 | apply(induction t arbitrary: t' rule: split_max.induct)
 | 
|  |    184 |  apply(fastforce split!: splits prod.splits)
 | 
|  |    185 | apply simp
 | 
|  |    186 | done
 | 
|  |    187 | 
 | 
|  |    188 | lemma neq_Leaf_if_height_neq_0[simp]: "height t \<noteq> 0 \<Longrightarrow> t \<noteq> Leaf"
 | 
|  |    189 | by auto
 | 
|  |    190 | 
 | 
| 71828 |    191 | theorem inorder_del:
 | 
|  |    192 |   "\<lbrakk> avl t; sorted(inorder t) \<rbrakk>  \<Longrightarrow> inorder (tree(del x t)) = del_list x (inorder t)"
 | 
| 71814 |    193 | apply(induction t rule: tree2_induct)
 | 
|  |    194 | apply(auto simp: del_list_simps inorder_baldL inorder_baldR avl_delete inorder_split_maxD
 | 
|  |    195 |            simp del: baldR.simps baldL.simps split!: splits prod.splits)
 | 
|  |    196 | done
 | 
|  |    197 | 
 | 
|  |    198 | 
 | 
|  |    199 | subsubsection \<open>Set Implementation\<close>
 | 
|  |    200 | 
 | 
|  |    201 | interpretation S: Set_by_Ordered
 | 
|  |    202 | where empty = Leaf and isin = isin
 | 
| 71828 |    203 |   and insert = insert
 | 
|  |    204 |   and delete = delete
 | 
| 71814 |    205 |   and inorder = inorder and inv = avl
 | 
|  |    206 | proof (standard, goal_cases)
 | 
|  |    207 |   case 1 show ?case by (simp)
 | 
|  |    208 | next
 | 
|  |    209 |   case 2 thus ?case by(simp add: isin_set_inorder)
 | 
|  |    210 | next
 | 
| 71828 |    211 |   case 3 thus ?case by(simp add: inorder_ins insert_def)
 | 
| 71814 |    212 | next
 | 
| 71828 |    213 |   case 4 thus ?case by(simp add: inorder_del delete_def)
 | 
| 71814 |    214 | next
 | 
| 71828 |    215 |   case 5 thus ?case by (simp)
 | 
| 71814 |    216 | next
 | 
|  |    217 |   case 6 thus ?case by (simp add: avl_insert)
 | 
|  |    218 | next
 | 
|  |    219 |   case 7 thus ?case by (simp add: avl_delete)
 | 
|  |    220 | qed
 | 
|  |    221 | 
 | 
|  |    222 | end
 |