| 71795 |      1 | (*
 | 
|  |      2 | Author:     Tobias Nipkow, Daniel Stüwe
 | 
|  |      3 | Based on the AFP entry AVL.
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | section "AVL Tree Implementation of Sets"
 | 
|  |      7 | 
 | 
|  |      8 | theory AVL_Set_Code
 | 
|  |      9 | imports
 | 
|  |     10 |   Cmp
 | 
|  |     11 |   Isin2
 | 
|  |     12 | begin
 | 
|  |     13 | 
 | 
|  |     14 | subsection \<open>Code\<close>
 | 
|  |     15 | 
 | 
|  |     16 | type_synonym 'a avl_tree = "('a*nat) tree"
 | 
|  |     17 | 
 | 
|  |     18 | definition empty :: "'a avl_tree" where
 | 
|  |     19 | "empty = Leaf"
 | 
|  |     20 | 
 | 
|  |     21 | fun ht :: "'a avl_tree \<Rightarrow> nat" where
 | 
|  |     22 | "ht Leaf = 0" |
 | 
|  |     23 | "ht (Node l (a,n) r) = n"
 | 
|  |     24 | 
 | 
|  |     25 | definition node :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
 | 
|  |     26 | "node l a r = Node l (a, max (ht l) (ht r) + 1) r"
 | 
|  |     27 | 
 | 
|  |     28 | definition balL :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
 | 
|  |     29 | "balL AB b C =
 | 
|  |     30 |   (if ht AB = ht C + 2 then
 | 
|  |     31 |      case AB of 
 | 
|  |     32 |        Node A (a, _) B \<Rightarrow>
 | 
|  |     33 |          if ht A \<ge> ht B then node A a (node B b C)
 | 
|  |     34 |          else
 | 
|  |     35 |            case B of
 | 
|  |     36 |              Node B\<^sub>1 (ab, _) B\<^sub>2 \<Rightarrow> node (node A a B\<^sub>1) ab (node B\<^sub>2 b C)
 | 
|  |     37 |    else node AB b C)"
 | 
|  |     38 | 
 | 
|  |     39 | definition balR :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
 | 
|  |     40 | "balR A a BC =
 | 
|  |     41 |    (if ht BC = ht A + 2 then
 | 
|  |     42 |       case BC of
 | 
|  |     43 |         Node B (b, _) C \<Rightarrow>
 | 
|  |     44 |           if ht B \<le> ht C then node (node A a B) b C
 | 
|  |     45 |           else
 | 
|  |     46 |             case B of
 | 
|  |     47 |               Node B\<^sub>1 (ab, _) B\<^sub>2 \<Rightarrow> node (node A a B\<^sub>1) ab (node B\<^sub>2 b C)
 | 
|  |     48 |   else node A a BC)"
 | 
|  |     49 | 
 | 
|  |     50 | fun insert :: "'a::linorder \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
 | 
|  |     51 | "insert x Leaf = Node Leaf (x, 1) Leaf" |
 | 
|  |     52 | "insert x (Node l (a, n) r) = (case cmp x a of
 | 
|  |     53 |    EQ \<Rightarrow> Node l (a, n) r |
 | 
|  |     54 |    LT \<Rightarrow> balL (insert x l) a r |
 | 
|  |     55 |    GT \<Rightarrow> balR l a (insert x r))"
 | 
|  |     56 | 
 | 
|  |     57 | fun split_max :: "'a avl_tree \<Rightarrow> 'a avl_tree * 'a" where
 | 
|  |     58 | "split_max (Node l (a, _) r) =
 | 
|  |     59 |   (if r = Leaf then (l,a) else let (r',a') = split_max r in (balL l a r', a'))"
 | 
|  |     60 | 
 | 
|  |     61 | lemmas split_max_induct = split_max.induct[case_names Node Leaf]
 | 
|  |     62 | 
 | 
|  |     63 | fun delete :: "'a::linorder \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
 | 
|  |     64 | "delete _ Leaf = Leaf" |
 | 
|  |     65 | "delete x (Node l (a, n) r) =
 | 
|  |     66 |   (case cmp x a of
 | 
|  |     67 |      EQ \<Rightarrow> if l = Leaf then r
 | 
|  |     68 |            else let (l', a') = split_max l in balR l' a' r |
 | 
|  |     69 |      LT \<Rightarrow> balR (delete x l) a r |
 | 
|  |     70 |      GT \<Rightarrow> balL l a (delete x r))"
 | 
|  |     71 | 
 | 
|  |     72 | 
 | 
|  |     73 | subsection \<open>Functional Correctness Proofs\<close>
 | 
|  |     74 | 
 | 
|  |     75 | text\<open>Very different from the AFP/AVL proofs\<close>
 | 
|  |     76 | 
 | 
|  |     77 | 
 | 
|  |     78 | subsubsection "Proofs for insert"
 | 
|  |     79 | 
 | 
|  |     80 | lemma inorder_balL:
 | 
|  |     81 |   "inorder (balL l a r) = inorder l @ a # inorder r"
 | 
|  |     82 | by (auto simp: node_def balL_def split:tree.splits)
 | 
|  |     83 | 
 | 
|  |     84 | lemma inorder_balR:
 | 
|  |     85 |   "inorder (balR l a r) = inorder l @ a # inorder r"
 | 
|  |     86 | by (auto simp: node_def balR_def split:tree.splits)
 | 
|  |     87 | 
 | 
|  |     88 | theorem inorder_insert:
 | 
|  |     89 |   "sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)"
 | 
|  |     90 | by (induct t) 
 | 
|  |     91 |    (auto simp: ins_list_simps inorder_balL inorder_balR)
 | 
|  |     92 | 
 | 
|  |     93 | 
 | 
|  |     94 | subsubsection "Proofs for delete"
 | 
|  |     95 | 
 | 
|  |     96 | lemma inorder_split_maxD:
 | 
|  |     97 |   "\<lbrakk> split_max t = (t',a); t \<noteq> Leaf \<rbrakk> \<Longrightarrow>
 | 
|  |     98 |    inorder t' @ [a] = inorder t"
 | 
|  |     99 | by(induction t arbitrary: t' rule: split_max.induct)
 | 
|  |    100 |   (auto simp: inorder_balL split: if_splits prod.splits tree.split)
 | 
|  |    101 | 
 | 
|  |    102 | theorem inorder_delete:
 | 
|  |    103 |   "sorted(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
 | 
|  |    104 | by(induction t)
 | 
|  |    105 |   (auto simp: del_list_simps inorder_balL inorder_balR inorder_split_maxD split: prod.splits)
 | 
|  |    106 | 
 | 
|  |    107 | end
 |