| 61784 |      1 | (* Author: Tobias Nipkow *)
 | 
|  |      2 | 
 | 
|  |      3 | section \<open>A 1-2 Brother Tree Implementation of Sets\<close>
 | 
|  |      4 | 
 | 
|  |      5 | theory Brother12_Set
 | 
|  |      6 | imports
 | 
|  |      7 |   Cmp
 | 
|  |      8 |   Set_by_Ordered
 | 
|  |      9 | begin
 | 
|  |     10 | 
 | 
|  |     11 | subsection \<open>Data Type and Operations\<close>
 | 
|  |     12 | 
 | 
|  |     13 | datatype 'a bro =
 | 
|  |     14 |   N0 |
 | 
|  |     15 |   N1 "'a bro" |
 | 
|  |     16 |   N2 "'a bro" 'a "'a bro" |
 | 
|  |     17 |   (* auxiliary constructors: *)
 | 
|  |     18 |   L2 'a |
 | 
|  |     19 |   N3 "'a bro" 'a "'a bro" 'a "'a bro"
 | 
|  |     20 | 
 | 
|  |     21 | fun inorder :: "'a bro \<Rightarrow> 'a list" where
 | 
|  |     22 | "inorder N0 = []" |
 | 
|  |     23 | "inorder (N1 t) = inorder t" |
 | 
|  |     24 | "inorder (N2 l a r) = inorder l @ a # inorder r" |
 | 
|  |     25 | "inorder (L2 a) = [a]" |
 | 
|  |     26 | "inorder (N3 t1 a1 t2 a2 t3) = inorder t1 @ a1 # inorder t2 @ a2 # inorder t3"
 | 
|  |     27 | 
 | 
|  |     28 | fun isin :: "'a bro \<Rightarrow> 'a::cmp \<Rightarrow> bool" where
 | 
|  |     29 | "isin N0 x = False" |
 | 
|  |     30 | "isin (N1 t) x = isin t x" |
 | 
|  |     31 | "isin (N2 l a r) x =
 | 
|  |     32 |   (case cmp x a of
 | 
|  |     33 |      LT \<Rightarrow> isin l x |
 | 
|  |     34 |      EQ \<Rightarrow> True |
 | 
|  |     35 |      GT \<Rightarrow> isin r x)"
 | 
|  |     36 | 
 | 
|  |     37 | fun n1 :: "'a bro \<Rightarrow> 'a bro" where
 | 
|  |     38 | "n1 (L2 a) = N2 N0 a N0" |
 | 
|  |     39 | "n1 (N3 t1 a1 t2 a2 t3) = N2 (N2 t1 a1 t2) a2 (N1 t3)" |
 | 
|  |     40 | "n1 t = N1 t"
 | 
|  |     41 | 
 | 
|  |     42 | hide_const (open) insert
 | 
|  |     43 | 
 | 
|  |     44 | locale insert
 | 
|  |     45 | begin
 | 
|  |     46 | 
 | 
|  |     47 | fun n2 :: "'a bro \<Rightarrow> 'a \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where
 | 
|  |     48 | "n2 (L2 a1) a2 t = N3 N0 a1 N0 a2 t" |
 | 
|  |     49 | "n2 (N3 t1 a1 t2 a2 t3) a3 (N1 t4) = N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4)" |
 | 
|  |     50 | "n2 (N3 t1 a1 t2 a2 t3) a3 t4 = N3 (N2 t1 a1 t2) a2 (N1 t3) a3 t4" |
 | 
|  |     51 | "n2 t1 a1 (L2 a2) = N3 t1 a1 N0 a2 N0" |
 | 
|  |     52 | "n2 (N1 t1) a1 (N3 t2 a2 t3 a3 t4) = N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4)" |
 | 
|  |     53 | "n2 t1 a1 (N3 t2 a2 t3 a3 t4) = N3 t1 a1 (N1 t2) a2 (N2 t3 a3 t4)" |
 | 
|  |     54 | "n2 t1 a t2 = N2 t1 a t2"
 | 
|  |     55 | 
 | 
|  |     56 | fun ins :: "'a::cmp \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where
 | 
| 61789 |     57 | "ins x N0 = L2 x" |
 | 
|  |     58 | "ins x (N1 t) = n1 (ins x t)" |
 | 
|  |     59 | "ins x (N2 l a r) =
 | 
|  |     60 |   (case cmp x a of
 | 
|  |     61 |      LT \<Rightarrow> n2 (ins x l) a r |
 | 
|  |     62 |      EQ \<Rightarrow> N2 l a r |
 | 
|  |     63 |      GT \<Rightarrow> n2 l a (ins x r))"
 | 
| 61784 |     64 | 
 | 
|  |     65 | fun tree :: "'a bro \<Rightarrow> 'a bro" where
 | 
|  |     66 | "tree (L2 a) = N2 N0 a N0" |
 | 
|  |     67 | "tree (N3 t1 a1 t2 a2 t3) = N2 (N2 t1 a1 t2) a2 (N1 t3)" |
 | 
|  |     68 | "tree t = t"
 | 
|  |     69 | 
 | 
|  |     70 | definition insert :: "'a::cmp \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where
 | 
|  |     71 | "insert x t = tree(ins x t)"
 | 
|  |     72 | 
 | 
|  |     73 | end
 | 
|  |     74 | 
 | 
|  |     75 | locale delete
 | 
|  |     76 | begin
 | 
|  |     77 | 
 | 
|  |     78 | fun n2 :: "'a bro \<Rightarrow> 'a \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where
 | 
|  |     79 | "n2 (N1 t1) a1 (N1 t2) = N1 (N2 t1 a1 t2)" |
 | 
|  |     80 | "n2 (N1 (N1 t1)) a1 (N2 (N1 t2) a2 (N2 t3 a3 t4)) =
 | 
|  |     81 |   N1 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4))" |
 | 
|  |     82 | "n2 (N1 (N1 t1)) a1 (N2 (N2 t2 a2 t3) a3 (N1 t4)) =
 | 
|  |     83 |   N1 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4))" |
 | 
|  |     84 | "n2 (N1 (N1 t1)) a1 (N2 (N2 t2 a2 t3) a3 (N2 t4 a4 t5)) =
 | 
|  |     85 |   N2 (N2 (N1 t1) a1 (N2 t2 a2 t3)) a3 (N1 (N2 t4 a4 t5))" |
 | 
|  |     86 | "n2 (N2 (N1 t1) a1 (N2 t2 a2 t3)) a3 (N1 (N1 t4)) =
 | 
|  |     87 |   N1 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4))" |
 | 
|  |     88 | "n2 (N2 (N2 t1 a1 t2) a2 (N1 t3)) a3 (N1 (N1 t4)) =
 | 
|  |     89 |   N1 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4))" |
 | 
|  |     90 | "n2 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4)) a5 (N1 (N1 t5)) =
 | 
|  |     91 |   N2 (N1 (N2 t1 a1 t2)) a2 (N2 (N2 t3 a3 t4) a5 (N1 t5))" |
 | 
|  |     92 | "n2 t1 a1 t2 = N2 t1 a1 t2"
 | 
|  |     93 | 
 | 
|  |     94 | fun del_min :: "'a bro \<Rightarrow> ('a \<times> 'a bro) option" where
 | 
|  |     95 | "del_min N0 = None" |
 | 
|  |     96 | "del_min (N1 t) =
 | 
|  |     97 |   (case del_min t of
 | 
|  |     98 |      None \<Rightarrow> None |
 | 
|  |     99 |      Some (a, t') \<Rightarrow> Some (a, N1 t'))" |
 | 
|  |    100 | "del_min (N2 t1 a t2) =
 | 
|  |    101 |   (case del_min t1 of
 | 
|  |    102 |      None \<Rightarrow> Some (a, N1 t2) |
 | 
|  |    103 |      Some (b, t1') \<Rightarrow> Some (b, n2 t1' a t2))"
 | 
|  |    104 | 
 | 
|  |    105 | fun del :: "'a::cmp \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where
 | 
|  |    106 | "del _ N0         = N0" |
 | 
|  |    107 | "del x (N1 t)     = N1 (del x t)" |
 | 
|  |    108 | "del x (N2 l a r) =
 | 
|  |    109 |   (case cmp x a of
 | 
|  |    110 |      LT \<Rightarrow> n2 (del x l) a r |
 | 
|  |    111 |      GT \<Rightarrow> n2 l a (del x r) |
 | 
|  |    112 |      EQ \<Rightarrow> (case del_min r of
 | 
|  |    113 |               None \<Rightarrow> N1 l |
 | 
|  |    114 |               Some (b, r') \<Rightarrow> n2 l b r'))"
 | 
|  |    115 | 
 | 
|  |    116 | fun tree :: "'a bro \<Rightarrow> 'a bro" where
 | 
|  |    117 | "tree (N1 t) = t" |
 | 
|  |    118 | "tree t = t"
 | 
|  |    119 | 
 | 
|  |    120 | definition delete :: "'a::cmp \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where
 | 
|  |    121 | "delete a t = tree (del a t)"
 | 
|  |    122 | 
 | 
|  |    123 | end
 | 
|  |    124 | 
 | 
|  |    125 | subsection \<open>Invariants\<close>
 | 
|  |    126 | 
 | 
|  |    127 | fun B :: "nat \<Rightarrow> 'a bro set"
 | 
|  |    128 | and U :: "nat \<Rightarrow> 'a bro set" where
 | 
|  |    129 | "B 0 = {N0}" |
 | 
|  |    130 | "B (Suc h) = { N2 t1 a t2 | t1 a t2. 
 | 
|  |    131 |   t1 \<in> B h \<union> U h \<and> t2 \<in> B h \<or> t1 \<in> B h \<and> t2 \<in> B h \<union> U h}" |
 | 
|  |    132 | "U 0 = {}" |
 | 
|  |    133 | "U (Suc h) = N1 ` B h"
 | 
|  |    134 | 
 | 
|  |    135 | abbreviation "T h \<equiv> B h \<union> U h"
 | 
|  |    136 | 
 | 
|  |    137 | fun Bp :: "nat \<Rightarrow> 'a bro set" where
 | 
|  |    138 | "Bp 0 = B 0 \<union> L2 ` UNIV" |
 | 
|  |    139 | "Bp (Suc 0) = B (Suc 0) \<union> {N3 N0 a N0 b N0|a b. True}" |
 | 
|  |    140 | "Bp (Suc(Suc h)) = B (Suc(Suc h)) \<union>
 | 
|  |    141 |   {N3 t1 a t2 b t3 | t1 a t2 b t3. t1 \<in> B (Suc h) \<and> t2 \<in> U (Suc h) \<and> t3 \<in> B (Suc h)}"
 | 
|  |    142 | 
 | 
|  |    143 | fun Um :: "nat \<Rightarrow> 'a bro set" where
 | 
|  |    144 | "Um 0 = {}" |
 | 
|  |    145 | "Um (Suc h) = N1 ` T h"
 | 
|  |    146 | 
 | 
|  |    147 | 
 | 
|  |    148 | subsection "Functional Correctness Proofs"
 | 
|  |    149 | 
 | 
|  |    150 | subsubsection "Proofs for isin"
 | 
|  |    151 | 
 | 
|  |    152 | lemma
 | 
|  |    153 |   "t \<in> T h \<Longrightarrow> sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems(inorder t))"
 | 
|  |    154 | by(induction h arbitrary: t) (fastforce simp: elems_simps1 split: if_splits)+
 | 
|  |    155 | 
 | 
|  |    156 | lemma isin_set: "t \<in> T h \<Longrightarrow>
 | 
|  |    157 |   sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems(inorder t))"
 | 
|  |    158 | by(induction h arbitrary: t) (auto simp: elems_simps2 split: if_splits)
 | 
|  |    159 | 
 | 
|  |    160 | subsubsection "Proofs for insertion"
 | 
|  |    161 | 
 | 
|  |    162 | lemma inorder_n1: "inorder(n1 t) = inorder t"
 | 
|  |    163 | by(induction t rule: n1.induct) (auto simp: sorted_lems)
 | 
|  |    164 | 
 | 
|  |    165 | context insert
 | 
|  |    166 | begin
 | 
|  |    167 | 
 | 
|  |    168 | lemma inorder_n2: "inorder(n2 l a r) = inorder l @ a # inorder r"
 | 
|  |    169 | by(cases "(l,a,r)" rule: n2.cases) (auto simp: sorted_lems)
 | 
|  |    170 | 
 | 
|  |    171 | lemma inorder_tree: "inorder(tree t) = inorder t"
 | 
|  |    172 | by(cases t) auto
 | 
|  |    173 | 
 | 
|  |    174 | lemma inorder_ins: "t \<in> T h \<Longrightarrow>
 | 
|  |    175 |   sorted(inorder t) \<Longrightarrow> inorder(ins a t) = ins_list a (inorder t)"
 | 
|  |    176 | by(induction h arbitrary: t) (auto simp: ins_list_simps inorder_n1 inorder_n2)
 | 
|  |    177 | 
 | 
|  |    178 | lemma inorder_insert: "t \<in> T h \<Longrightarrow>
 | 
|  |    179 |   sorted(inorder t) \<Longrightarrow> inorder(insert a t) = ins_list a (inorder t)"
 | 
|  |    180 | by(simp add: insert_def inorder_ins inorder_tree)
 | 
|  |    181 | 
 | 
|  |    182 | end
 | 
|  |    183 | 
 | 
|  |    184 | subsubsection \<open>Proofs for deletion\<close>
 | 
|  |    185 | 
 | 
|  |    186 | context delete
 | 
|  |    187 | begin
 | 
|  |    188 | 
 | 
|  |    189 | lemma inorder_tree: "inorder(tree t) = inorder t"
 | 
|  |    190 | by(cases t) auto
 | 
|  |    191 | 
 | 
|  |    192 | lemma inorder_n2: "inorder(n2 l a r) = inorder l @ a # inorder r"
 | 
|  |    193 | by(induction l a r rule: n2.induct) (auto)
 | 
|  |    194 | 
 | 
|  |    195 | lemma inorder_del_min:
 | 
| 61792 |    196 |   "t \<in> T h \<Longrightarrow> (del_min t = None \<longleftrightarrow> inorder t = []) \<and>
 | 
| 61784 |    197 |   (del_min t = Some(a,t') \<longrightarrow> inorder t = a # inorder t')"
 | 
|  |    198 | by(induction h arbitrary: t a t') (auto simp: inorder_n2 split: option.splits)
 | 
|  |    199 | 
 | 
|  |    200 | lemma inorder_del:
 | 
| 61792 |    201 |   "t \<in> T h \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(del x t) = del_list x (inorder t)"
 | 
|  |    202 | by(induction h arbitrary: t) (auto simp: del_list_simps inorder_n2
 | 
|  |    203 |      inorder_del_min[OF UnI1] inorder_del_min[OF UnI2] split: option.splits)
 | 
|  |    204 | 
 | 
|  |    205 | lemma inorder_delete:
 | 
|  |    206 |   "t \<in> T h \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
 | 
|  |    207 | by(simp add: delete_def inorder_del inorder_tree)
 | 
| 61784 |    208 | 
 | 
|  |    209 | end
 | 
|  |    210 | 
 | 
|  |    211 | 
 | 
|  |    212 | subsection \<open>Invariant Proofs\<close>
 | 
|  |    213 | 
 | 
| 61789 |    214 | subsubsection \<open>Proofs for insertion\<close>
 | 
| 61784 |    215 | 
 | 
|  |    216 | lemma n1_type: "t \<in> Bp h \<Longrightarrow> n1 t \<in> T (Suc h)"
 | 
|  |    217 | by(cases h rule: Bp.cases) auto
 | 
|  |    218 | 
 | 
|  |    219 | context insert
 | 
|  |    220 | begin
 | 
|  |    221 | 
 | 
|  |    222 | lemma tree_type1: "t \<in> Bp h \<Longrightarrow> tree t \<in> B h \<union> B (Suc h)"
 | 
|  |    223 | by(cases h rule: Bp.cases) auto
 | 
|  |    224 | 
 | 
|  |    225 | lemma tree_type2: "t \<in> T h \<Longrightarrow> tree t \<in> T h"
 | 
|  |    226 | by(cases h) auto
 | 
|  |    227 | 
 | 
|  |    228 | lemma n2_type:
 | 
|  |    229 |   "(t1 \<in> Bp h \<and> t2 \<in> T h \<longrightarrow> n2 t1 a t2 \<in> Bp (Suc h)) \<and>
 | 
|  |    230 |    (t1 \<in> T h \<and> t2 \<in> Bp h \<longrightarrow> n2 t1 a t2 \<in> Bp (Suc h))"
 | 
|  |    231 | apply(cases h rule: Bp.cases)
 | 
|  |    232 | apply (auto)[2]
 | 
|  |    233 | apply(rule conjI impI | erule conjE exE imageE | simp | erule disjE)+
 | 
|  |    234 | done
 | 
|  |    235 | 
 | 
|  |    236 | lemma Bp_if_B: "t \<in> B h \<Longrightarrow> t \<in> Bp h"
 | 
|  |    237 | by (cases h rule: Bp.cases) simp_all
 | 
|  |    238 | 
 | 
|  |    239 | text{* An automatic proof: *}
 | 
|  |    240 | 
 | 
|  |    241 | lemma
 | 
|  |    242 |   "(t \<in> B h \<longrightarrow> ins x t \<in> Bp h) \<and> (t \<in> U h \<longrightarrow> ins x t \<in> T h)"
 | 
|  |    243 | apply(induction h arbitrary: t)
 | 
|  |    244 |  apply (simp)
 | 
|  |    245 | apply (fastforce simp: Bp_if_B n2_type dest: n1_type)
 | 
|  |    246 | done
 | 
|  |    247 | 
 | 
|  |    248 | text{* A detailed proof: *}
 | 
|  |    249 | 
 | 
|  |    250 | lemma ins_type:
 | 
|  |    251 | shows "t \<in> B h \<Longrightarrow> ins x t \<in> Bp h" and "t \<in> U h \<Longrightarrow> ins x t \<in> T h"
 | 
|  |    252 | proof(induction h arbitrary: t)
 | 
|  |    253 |   case 0
 | 
|  |    254 |   { case 1 thus ?case by simp
 | 
|  |    255 |   next
 | 
|  |    256 |     case 2 thus ?case by simp }
 | 
|  |    257 | next
 | 
|  |    258 |   case (Suc h)
 | 
|  |    259 |   { case 1
 | 
|  |    260 |     then obtain t1 a t2 where [simp]: "t = N2 t1 a t2" and
 | 
|  |    261 |       t1: "t1 \<in> T h" and t2: "t2 \<in> T h" and t12: "t1 \<in> B h \<or> t2 \<in> B h"
 | 
|  |    262 |       by auto
 | 
|  |    263 |     { assume "x < a"
 | 
|  |    264 |       hence "?case \<longleftrightarrow> n2 (ins x t1) a t2 \<in> Bp (Suc h)" by simp
 | 
|  |    265 |       also have "\<dots>"
 | 
|  |    266 |       proof cases
 | 
|  |    267 |         assume "t1 \<in> B h"
 | 
|  |    268 |         with t2 show ?thesis by (simp add: Suc.IH(1) n2_type)
 | 
|  |    269 |       next
 | 
|  |    270 |         assume "t1 \<notin> B h"
 | 
|  |    271 |         hence 1: "t1 \<in> U h" and 2: "t2 \<in> B h" using t1 t12 by auto
 | 
|  |    272 |         show ?thesis by (metis Suc.IH(2)[OF 1] Bp_if_B[OF 2] n2_type)
 | 
|  |    273 |       qed
 | 
|  |    274 |       finally have ?case .
 | 
|  |    275 |     }
 | 
|  |    276 |     moreover
 | 
|  |    277 |     { assume "a < x"
 | 
|  |    278 |       hence "?case \<longleftrightarrow> n2 t1 a (ins x t2) \<in> Bp (Suc h)" by simp
 | 
|  |    279 |       also have "\<dots>"
 | 
|  |    280 |       proof cases
 | 
|  |    281 |         assume "t2 \<in> B h"
 | 
|  |    282 |         with t1 show ?thesis by (simp add: Suc.IH(1) n2_type)
 | 
|  |    283 |       next
 | 
|  |    284 |         assume "t2 \<notin> B h"
 | 
|  |    285 |         hence 1: "t1 \<in> B h" and 2: "t2 \<in> U h" using t2 t12 by auto
 | 
|  |    286 |         show ?thesis by (metis Bp_if_B[OF 1] Suc.IH(2)[OF 2] n2_type)
 | 
|  |    287 |       qed
 | 
|  |    288 |     }
 | 
|  |    289 |     moreover 
 | 
|  |    290 |     { assume "x = a"
 | 
|  |    291 |       from 1 have "t \<in> Bp (Suc h)" by(rule Bp_if_B)
 | 
|  |    292 |       hence "?case" using `x = a` by simp
 | 
|  |    293 |     }
 | 
|  |    294 |     ultimately show ?case by auto
 | 
|  |    295 |   next
 | 
|  |    296 |     case 2 thus ?case using Suc(1) n1_type by fastforce }
 | 
|  |    297 | qed
 | 
|  |    298 | 
 | 
|  |    299 | lemma insert_type:
 | 
|  |    300 |   "t \<in> T h \<Longrightarrow> insert x t \<in> T h \<union> T (Suc h)"
 | 
|  |    301 | unfolding insert_def by (metis Un_iff ins_type tree_type1 tree_type2)
 | 
|  |    302 | 
 | 
|  |    303 | end
 | 
|  |    304 | 
 | 
| 61789 |    305 | subsubsection "Proofs for deletion"
 | 
| 61784 |    306 | 
 | 
|  |    307 | lemma B_simps[simp]: 
 | 
|  |    308 |   "N1 t \<in> B h = False"
 | 
|  |    309 |   "L2 y \<in> B h = False"
 | 
|  |    310 |   "(N3 t1 a1 t2 a2 t3) \<in> B h = False"
 | 
|  |    311 |   "N0 \<in> B h \<longleftrightarrow> h = 0"
 | 
|  |    312 | by (cases h, auto)+
 | 
|  |    313 | 
 | 
|  |    314 | context delete
 | 
|  |    315 | begin
 | 
|  |    316 | 
 | 
|  |    317 | lemma n2_type1:
 | 
|  |    318 |   "\<lbrakk>t1 \<in> Um h; t2 \<in> B h\<rbrakk> \<Longrightarrow> n2 t1 a t2 \<in> T (Suc h)"
 | 
|  |    319 | apply(cases h rule: Bp.cases)
 | 
|  |    320 | apply auto[2]
 | 
|  |    321 | apply(erule exE bexE conjE imageE | simp | erule disjE)+
 | 
|  |    322 | done
 | 
|  |    323 | 
 | 
|  |    324 | lemma n2_type2:
 | 
|  |    325 |   "\<lbrakk>t1 \<in> B h ; t2 \<in> Um h \<rbrakk> \<Longrightarrow> n2 t1 a t2 \<in> T (Suc h)"
 | 
|  |    326 | apply(cases h rule: Bp.cases)
 | 
|  |    327 | apply auto[2]
 | 
|  |    328 | apply(erule exE bexE conjE imageE | simp | erule disjE)+
 | 
|  |    329 | done
 | 
|  |    330 | 
 | 
|  |    331 | lemma n2_type3:
 | 
|  |    332 |   "\<lbrakk>t1 \<in> T h ; t2 \<in> T h \<rbrakk> \<Longrightarrow> n2 t1 a t2 \<in> T (Suc h)"
 | 
|  |    333 | apply(cases h rule: Bp.cases)
 | 
|  |    334 | apply auto[2]
 | 
|  |    335 | apply(erule exE bexE conjE imageE | simp | erule disjE)+
 | 
|  |    336 | done
 | 
|  |    337 | 
 | 
|  |    338 | lemma del_minNoneN0: "\<lbrakk>t \<in> B h; del_min t = None\<rbrakk> \<Longrightarrow>  t = N0"
 | 
|  |    339 | by (cases t) (auto split: option.splits)
 | 
|  |    340 | 
 | 
|  |    341 | lemma del_minNoneN1 : "\<lbrakk>t \<in> U h; del_min t = None\<rbrakk> \<Longrightarrow> t = N1 N0"
 | 
|  |    342 | by (cases h) (auto simp: del_minNoneN0  split: option.splits)
 | 
|  |    343 | 
 | 
|  |    344 | lemma del_min_type:
 | 
|  |    345 |   "t \<in> B h \<Longrightarrow> del_min t = Some (a, t') \<Longrightarrow> t' \<in> T h"
 | 
|  |    346 |   "t \<in> U h \<Longrightarrow> del_min t = Some (a, t') \<Longrightarrow> t' \<in> Um h"
 | 
|  |    347 | proof (induction h arbitrary: t a t')
 | 
|  |    348 |   case (Suc h)
 | 
|  |    349 |   { case 1
 | 
|  |    350 |     then obtain t1 a t2 where [simp]: "t = N2 t1 a t2" and
 | 
|  |    351 |       t12: "t1 \<in> T h" "t2 \<in> T h" "t1 \<in> B h \<or> t2 \<in> B h"
 | 
|  |    352 |       by auto
 | 
|  |    353 |     show ?case
 | 
|  |    354 |     proof (cases "del_min t1")
 | 
|  |    355 |       case None
 | 
|  |    356 |       show ?thesis
 | 
|  |    357 |       proof cases
 | 
|  |    358 |         assume "t1 \<in> B h"
 | 
|  |    359 |         with del_minNoneN0[OF this None] 1 show ?thesis by(auto)
 | 
|  |    360 |       next
 | 
|  |    361 |         assume "t1 \<notin> B h"
 | 
|  |    362 |         thus ?thesis using 1 None by (auto)
 | 
|  |    363 |       qed
 | 
|  |    364 |     next
 | 
|  |    365 |       case [simp]: (Some bt')
 | 
|  |    366 |       obtain b t1' where [simp]: "bt' = (b,t1')" by fastforce
 | 
|  |    367 |       show ?thesis
 | 
|  |    368 |       proof cases
 | 
|  |    369 |         assume "t1 \<in> B h"
 | 
|  |    370 |         from Suc.IH(1)[OF this] 1 have "t1' \<in> T h" by simp
 | 
|  |    371 |         from n2_type3[OF this t12(2)] 1 show ?thesis by auto
 | 
|  |    372 |       next
 | 
|  |    373 |         assume "t1 \<notin> B h"
 | 
|  |    374 |         hence t1: "t1 \<in> U h" and t2: "t2 \<in> B h" using t12 by auto
 | 
|  |    375 |         from Suc.IH(2)[OF t1] have "t1' \<in> Um h" by simp
 | 
|  |    376 |         from n2_type1[OF this t2] 1 show ?thesis by auto
 | 
|  |    377 |       qed
 | 
|  |    378 |     qed
 | 
|  |    379 |   }
 | 
|  |    380 |   { case 2
 | 
|  |    381 |     then obtain t1 where [simp]: "t = N1 t1" and t1: "t1 \<in> B h" by auto
 | 
|  |    382 |     show ?case
 | 
|  |    383 |     proof (cases "del_min t1")
 | 
|  |    384 |       case None
 | 
|  |    385 |       with del_minNoneN0[OF t1 None] 2 show ?thesis by(auto)
 | 
|  |    386 |     next
 | 
|  |    387 |       case [simp]: (Some bt')
 | 
|  |    388 |       obtain b t1' where [simp]: "bt' = (b,t1')" by fastforce
 | 
|  |    389 |       from Suc.IH(1)[OF t1] have "t1' \<in> T h" by simp
 | 
|  |    390 |       thus ?thesis using 2 by auto
 | 
|  |    391 |     qed
 | 
|  |    392 |   }
 | 
|  |    393 | qed auto
 | 
|  |    394 | 
 | 
|  |    395 | lemma del_type:
 | 
|  |    396 |   "t \<in> B h \<Longrightarrow> del x t \<in> T h"
 | 
|  |    397 |   "t \<in> U h \<Longrightarrow> del x t \<in> Um h"
 | 
|  |    398 | proof (induction h arbitrary: x t)
 | 
|  |    399 |   case (Suc h)
 | 
|  |    400 |   { case 1
 | 
|  |    401 |     then obtain l a r where [simp]: "t = N2 l a r" and
 | 
|  |    402 |       lr: "l \<in> T h" "r \<in> T h" "l \<in> B h \<or> r \<in> B h" by auto
 | 
|  |    403 |     { assume "x < a"
 | 
|  |    404 |       have ?case
 | 
|  |    405 |       proof cases
 | 
|  |    406 |         assume "l \<in> B h"
 | 
|  |    407 |         from n2_type3[OF Suc.IH(1)[OF this] lr(2)]
 | 
|  |    408 |         show ?thesis using `x<a` by(simp)
 | 
|  |    409 |       next
 | 
|  |    410 |         assume "l \<notin> B h"
 | 
|  |    411 |         hence "l \<in> U h" "r \<in> B h" using lr by auto
 | 
|  |    412 |         from n2_type1[OF Suc.IH(2)[OF this(1)] this(2)]
 | 
|  |    413 |         show ?thesis using `x<a` by(simp)
 | 
|  |    414 |       qed
 | 
|  |    415 |     } moreover
 | 
|  |    416 |     { assume "x > a"
 | 
|  |    417 |       have ?case
 | 
|  |    418 |       proof cases
 | 
|  |    419 |         assume "r \<in> B h"
 | 
|  |    420 |         from n2_type3[OF lr(1) Suc.IH(1)[OF this]]
 | 
|  |    421 |         show ?thesis using `x>a` by(simp)
 | 
|  |    422 |       next
 | 
|  |    423 |         assume "r \<notin> B h"
 | 
|  |    424 |         hence "l \<in> B h" "r \<in> U h" using lr by auto
 | 
|  |    425 |         from n2_type2[OF this(1) Suc.IH(2)[OF this(2)]]
 | 
|  |    426 |         show ?thesis using `x>a` by(simp)
 | 
|  |    427 |       qed
 | 
|  |    428 |     } moreover
 | 
|  |    429 |     { assume [simp]: "x=a"
 | 
|  |    430 |       have ?case
 | 
|  |    431 |       proof (cases "del_min r")
 | 
|  |    432 |         case None
 | 
|  |    433 |         show ?thesis
 | 
|  |    434 |         proof cases
 | 
|  |    435 |           assume "r \<in> B h"
 | 
|  |    436 |           with del_minNoneN0[OF this None] lr show ?thesis by(simp)
 | 
|  |    437 |         next
 | 
|  |    438 |           assume "r \<notin> B h"
 | 
|  |    439 |           hence "r \<in> U h" using lr by auto
 | 
|  |    440 |           with del_minNoneN1[OF this None] lr(3) show ?thesis by (simp)
 | 
|  |    441 |         qed
 | 
|  |    442 |       next
 | 
|  |    443 |         case [simp]: (Some br')
 | 
|  |    444 |         obtain b r' where [simp]: "br' = (b,r')" by fastforce
 | 
|  |    445 |         show ?thesis
 | 
|  |    446 |         proof cases
 | 
|  |    447 |           assume "r \<in> B h"
 | 
|  |    448 |           from del_min_type(1)[OF this] n2_type3[OF lr(1)]
 | 
|  |    449 |           show ?thesis by simp
 | 
|  |    450 |         next
 | 
|  |    451 |           assume "r \<notin> B h"
 | 
|  |    452 |           hence "l \<in> B h" and "r \<in> U h" using lr by auto
 | 
|  |    453 |           from del_min_type(2)[OF this(2)] n2_type2[OF this(1)]
 | 
|  |    454 |           show ?thesis by simp
 | 
|  |    455 |         qed
 | 
|  |    456 |       qed
 | 
|  |    457 |     } ultimately show ?case by auto
 | 
|  |    458 |   }
 | 
|  |    459 |   { case 2 with Suc.IH(1) show ?case by auto }
 | 
|  |    460 | qed auto
 | 
|  |    461 | 
 | 
|  |    462 | lemma tree_type:
 | 
|  |    463 |   "t \<in> Um (Suc h) \<Longrightarrow> tree t : T h"
 | 
|  |    464 |   "t \<in> T (Suc h) \<Longrightarrow> tree t : T h \<union> T(h+1)"
 | 
|  |    465 | by(auto)
 | 
|  |    466 | 
 | 
|  |    467 | lemma delete_type:
 | 
|  |    468 |   "t \<in> T h \<Longrightarrow> delete x t \<in> T h \<union> T(h-1)"
 | 
|  |    469 | unfolding delete_def
 | 
|  |    470 | by (cases h) (simp, metis del_type tree_type Un_iff Suc_eq_plus1 diff_Suc_1)
 | 
|  |    471 | 
 | 
|  |    472 | end
 | 
|  |    473 | 
 | 
| 61789 |    474 | 
 | 
| 61784 |    475 | subsection "Overall correctness"
 | 
|  |    476 | 
 | 
|  |    477 | interpretation Set_by_Ordered
 | 
| 61789 |    478 | where empty = N0 and isin = isin and insert = insert.insert
 | 
|  |    479 | and delete = delete.delete and inorder = inorder and inv = "\<lambda>t. \<exists>h. t \<in> T h"
 | 
| 61784 |    480 | proof (standard, goal_cases)
 | 
|  |    481 |   case 2 thus ?case by(auto intro!: isin_set)
 | 
|  |    482 | next
 | 
|  |    483 |   case 3 thus ?case by(auto intro!: insert.inorder_insert)
 | 
|  |    484 | next
 | 
| 61792 |    485 |   case 4 thus ?case by(auto intro!: delete.inorder_delete)
 | 
| 61784 |    486 | next
 | 
|  |    487 |   case 6 thus ?case using insert.insert_type by blast
 | 
|  |    488 | next
 | 
|  |    489 |   case 7 thus ?case using delete.delete_type by blast
 | 
|  |    490 | qed auto
 | 
|  |    491 | 
 | 
|  |    492 | end
 |