| author | wenzelm | 
| Sun, 03 Sep 2023 16:44:07 +0200 | |
| changeset 78639 | ca56952b7322 | 
| parent 74101 | d804e93ae9ff | 
| permissions | -rw-r--r-- | 
| 33436 | 1  | 
(* Title: HOL/ex/Tree23.thy  | 
2  | 
Author: Tobias Nipkow, TU Muenchen  | 
|
3  | 
*)  | 
|
4  | 
||
| 61343 | 5  | 
section \<open>2-3 Trees\<close>  | 
| 33436 | 6  | 
|
7  | 
theory Tree23  | 
|
8  | 
imports Main  | 
|
9  | 
begin  | 
|
10  | 
||
| 74101 | 11  | 
hide_const (open) or  | 
12  | 
||
| 61343 | 13  | 
text\<open>This is a very direct translation of some of the functions in table.ML  | 
| 33436 | 14  | 
in the Isabelle source code. That source is due to Makarius Wenzel and Stefan  | 
15  | 
Berghofer.  | 
|
16  | 
||
17  | 
Note that because of complicated patterns and mutual recursion, these  | 
|
18  | 
function definitions take a few minutes and can also be seen as stress tests  | 
|
| 61343 | 19  | 
for the function definition facility.\<close>  | 
| 33436 | 20  | 
|
| 61933 | 21  | 
type_synonym key = int \<comment> \<open>for simplicity, should be a type class\<close>  | 
| 33436 | 22  | 
|
| 58310 | 23  | 
datatype ord = LESS | EQUAL | GREATER  | 
| 33436 | 24  | 
|
25  | 
definition "ord i j = (if i<j then LESS else if i=j then EQUAL else GREATER)"  | 
|
26  | 
||
| 58310 | 27  | 
datatype 'a tree23 =  | 
| 33436 | 28  | 
Empty |  | 
29  | 
Branch2 "'a tree23" "key * 'a" "'a tree23" |  | 
|
30  | 
Branch3 "'a tree23" "key * 'a" "'a tree23" "key * 'a" "'a tree23"  | 
|
31  | 
||
| 58310 | 32  | 
datatype 'a growth =  | 
| 33436 | 33  | 
Stay "'a tree23" |  | 
34  | 
Sprout "'a tree23" "key * 'a" "'a tree23"  | 
|
35  | 
||
36  | 
fun add :: "key \<Rightarrow> 'a \<Rightarrow> 'a tree23 \<Rightarrow> 'a growth" where  | 
|
37  | 
"add key y Empty = Sprout Empty (key,y) Empty" |  | 
|
38  | 
"add key y (Branch2 left (k,x) right) =  | 
|
39  | 
(case ord key k of  | 
|
40  | 
LESS =>  | 
|
41  | 
(case add key y left of  | 
|
42  | 
Stay left' => Stay (Branch2 left' (k,x) right)  | 
|
43  | 
| Sprout left1 q left2  | 
|
44  | 
=> Stay (Branch3 left1 q left2 (k,x) right))  | 
|
45  | 
| EQUAL => Stay (Branch2 left (k,y) right)  | 
|
46  | 
| GREATER =>  | 
|
47  | 
(case add key y right of  | 
|
48  | 
Stay right' => Stay (Branch2 left (k,x) right')  | 
|
49  | 
| Sprout right1 q right2  | 
|
50  | 
=> Stay (Branch3 left (k,x) right1 q right2)))" |  | 
|
51  | 
"add key y (Branch3 left (k1,x1) mid (k2,x2) right) =  | 
|
52  | 
(case ord key k1 of  | 
|
53  | 
LESS =>  | 
|
54  | 
(case add key y left of  | 
|
55  | 
Stay left' => Stay (Branch3 left' (k1,x1) mid (k2,x2) right)  | 
|
56  | 
| Sprout left1 q left2  | 
|
57  | 
=> Sprout (Branch2 left1 q left2) (k1,x1) (Branch2 mid (k2,x2) right))  | 
|
58  | 
| EQUAL => Stay (Branch3 left (k1,y) mid (k2,x2) right)  | 
|
59  | 
| GREATER =>  | 
|
60  | 
(case ord key k2 of  | 
|
61  | 
LESS =>  | 
|
62  | 
(case add key y mid of  | 
|
63  | 
Stay mid' => Stay (Branch3 left (k1,x1) mid' (k2,x2) right)  | 
|
64  | 
| Sprout mid1 q mid2  | 
|
65  | 
=> Sprout (Branch2 left (k1,x1) mid1) q (Branch2 mid2 (k2,x2) right))  | 
|
66  | 
| EQUAL => Stay (Branch3 left (k1,x1) mid (k2,y) right)  | 
|
67  | 
| GREATER =>  | 
|
68  | 
(case add key y right of  | 
|
69  | 
Stay right' => Stay (Branch3 left (k1,x1) mid (k2,x2) right')  | 
|
70  | 
| Sprout right1 q right2  | 
|
71  | 
=> Sprout (Branch2 left (k1,x1) mid) (k2,x2) (Branch2 right1 q right2))))"  | 
|
72  | 
||
73  | 
definition add0 :: "key \<Rightarrow> 'a \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where  | 
|
74  | 
"add0 k y t =  | 
|
75  | 
(case add k y t of Stay t' => t' | Sprout l p r => Branch2 l p r)"  | 
|
76  | 
||
77  | 
value "add0 5 e (add0 4 d (add0 3 c (add0 2 b (add0 1 a Empty))))"  | 
|
78  | 
||
79  | 
fun compare where  | 
|
80  | 
"compare None (k2, _) = LESS" |  | 
|
81  | 
"compare (Some k1) (k2, _) = ord k1 k2"  | 
|
82  | 
||
83  | 
fun if_eq where  | 
|
84  | 
"if_eq EQUAL x y = x" |  | 
|
85  | 
"if_eq _ x y = y"  | 
|
86  | 
||
87  | 
fun del :: "key option \<Rightarrow> 'a tree23 \<Rightarrow> ((key * 'a) * bool * 'a tree23)option"  | 
|
88  | 
where  | 
|
89  | 
"del (Some k) Empty = None" |  | 
|
90  | 
"del None (Branch2 Empty p Empty) = Some(p, (True, Empty))" |  | 
|
91  | 
"del None (Branch3 Empty p Empty q Empty) = Some(p, (False, Branch2 Empty q Empty))" |  | 
|
92  | 
"del k (Branch2 Empty p Empty) = (case compare k p of  | 
|
93  | 
EQUAL => Some(p, (True, Empty)) | _ => None)" |  | 
|
94  | 
"del k (Branch3 Empty p Empty q Empty) = (case compare k p of  | 
|
95  | 
EQUAL => Some(p, (False, Branch2 Empty q Empty))  | 
|
96  | 
| _ => (case compare k q of  | 
|
97  | 
EQUAL => Some(q, (False, Branch2 Empty p Empty))  | 
|
98  | 
| _ => None))" |  | 
|
99  | 
"del k (Branch2 l p r) = (case compare k p of  | 
|
100  | 
LESS => (case del k l of None \<Rightarrow> None |  | 
|
101  | 
Some(p', (False, l')) => Some(p', (False, Branch2 l' p r))  | 
|
102  | 
| Some(p', (True, l')) => Some(p', case r of  | 
|
103  | 
Branch2 rl rp rr => (True, Branch3 l' p rl rp rr)  | 
|
104  | 
| Branch3 rl rp rm rq rr => (False, Branch2  | 
|
105  | 
(Branch2 l' p rl) rp (Branch2 rm rq rr))))  | 
|
106  | 
| or => (case del (if_eq or None k) r of None \<Rightarrow> None |  | 
|
107  | 
Some(p', (False, r')) => Some(p', (False, Branch2 l (if_eq or p' p) r'))  | 
|
108  | 
| Some(p', (True, r')) => Some(p', case l of  | 
|
109  | 
Branch2 ll lp lr => (True, Branch3 ll lp lr (if_eq or p' p) r')  | 
|
110  | 
| Branch3 ll lp lm lq lr => (False, Branch2  | 
|
111  | 
(Branch2 ll lp lm) lq (Branch2 lr (if_eq or p' p) r')))))" |  | 
|
112  | 
"del k (Branch3 l p m q r) = (case compare k q of  | 
|
113  | 
LESS => (case compare k p of  | 
|
114  | 
LESS => (case del k l of None \<Rightarrow> None |  | 
|
115  | 
Some(p', (False, l')) => Some(p', (False, Branch3 l' p m q r))  | 
|
116  | 
| Some(p', (True, l')) => Some(p', (False, case (m, r) of  | 
|
117  | 
(Branch2 ml mp mr, Branch2 _ _ _) => Branch2 (Branch3 l' p ml mp mr) q r  | 
|
118  | 
| (Branch3 ml mp mm mq mr, _) => Branch3 (Branch2 l' p ml) mp (Branch2 mm mq mr) q r  | 
|
119  | 
| (Branch2 ml mp mr, Branch3 rl rp rm rq rr) =>  | 
|
120  | 
Branch3 (Branch2 l' p ml) mp (Branch2 mr q rl) rp (Branch2 rm rq rr))))  | 
|
121  | 
| or => (case del (if_eq or None k) m of None \<Rightarrow> None |  | 
|
122  | 
Some(p', (False, m')) => Some(p', (False, Branch3 l (if_eq or p' p) m' q r))  | 
|
123  | 
| Some(p', (True, m')) => Some(p', (False, case (l, r) of  | 
|
124  | 
(Branch2 ll lp lr, Branch2 _ _ _) => Branch2 (Branch3 ll lp lr (if_eq or p' p) m') q r  | 
|
125  | 
| (Branch3 ll lp lm lq lr, _) => Branch3 (Branch2 ll lp lm) lq (Branch2 lr (if_eq or p' p) m') q r  | 
|
126  | 
| (_, Branch3 rl rp rm rq rr) => Branch3 l (if_eq or p' p) (Branch2 m' q rl) rp (Branch2 rm rq rr)))))  | 
|
127  | 
| or => (case del (if_eq or None k) r of None \<Rightarrow> None |  | 
|
128  | 
Some(q', (False, r')) => Some(q', (False, Branch3 l p m (if_eq or q' q) r'))  | 
|
129  | 
| Some(q', (True, r')) => Some(q', (False, case (l, m) of  | 
|
130  | 
(Branch2 _ _ _, Branch2 ml mp mr) => Branch2 l p (Branch3 ml mp mr (if_eq or q' q) r')  | 
|
131  | 
| (_, Branch3 ml mp mm mq mr) => Branch3 l p (Branch2 ml mp mm) mq (Branch2 mr (if_eq or q' q) r')  | 
|
132  | 
| (Branch3 ll lp lm lq lr, Branch2 ml mp mr) =>  | 
|
133  | 
Branch3 (Branch2 ll lp lm) lq (Branch2 lr p ml) mp (Branch2 mr (if_eq or q' q) r')))))"  | 
|
134  | 
||
135  | 
definition del0 :: "key \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where  | 
|
136  | 
"del0 k t = (case del (Some k) t of None \<Rightarrow> t | Some(_,(_,t')) \<Rightarrow> t')"  | 
|
137  | 
||
| 61343 | 138  | 
text \<open>Ordered trees\<close>  | 
| 33436 | 139  | 
|
| 
45325
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
140  | 
definition opt_less :: "key option \<Rightarrow> key option \<Rightarrow> bool" where  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
141  | 
"opt_less x y = (case x of None \<Rightarrow> True | Some a \<Rightarrow> (case y of None \<Rightarrow> True | Some b \<Rightarrow> a < b))"  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
142  | 
|
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
143  | 
lemma opt_less_simps [simp]:  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
144  | 
"opt_less None y = True"  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
145  | 
"opt_less x None = True"  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
146  | 
"opt_less (Some a) (Some b) = (a < b)"  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
147  | 
unfolding opt_less_def by (auto simp add: ord_def split: option.split)  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
148  | 
|
| 
45333
 
04b21922ed68
ex/Tree23.thy: simpler definition of ordered-ness predicate
 
huffman 
parents: 
45332 
diff
changeset
 | 
149  | 
primrec ord' :: "key option \<Rightarrow> 'a tree23 \<Rightarrow> key option \<Rightarrow> bool" where  | 
| 
45325
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
150  | 
"ord' x Empty y = opt_less x y" |  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
151  | 
"ord' x (Branch2 l p r) y = (ord' x l (Some (fst p)) & ord' (Some (fst p)) r y)" |  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
152  | 
"ord' x (Branch3 l p m q r) y = (ord' x l (Some (fst p)) & ord' (Some (fst p)) m (Some (fst q)) & ord' (Some (fst q)) r y)"  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
153  | 
|
| 
45333
 
04b21922ed68
ex/Tree23.thy: simpler definition of ordered-ness predicate
 
huffman 
parents: 
45332 
diff
changeset
 | 
154  | 
definition ord0 :: "'a tree23 \<Rightarrow> bool" where  | 
| 
 
04b21922ed68
ex/Tree23.thy: simpler definition of ordered-ness predicate
 
huffman 
parents: 
45332 
diff
changeset
 | 
155  | 
"ord0 t = ord' None t None"  | 
| 
45325
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
156  | 
|
| 61343 | 157  | 
text \<open>Balanced trees\<close>  | 
| 45334 | 158  | 
|
159  | 
inductive full :: "nat \<Rightarrow> 'a tree23 \<Rightarrow> bool" where  | 
|
160  | 
"full 0 Empty" |  | 
|
161  | 
"\<lbrakk>full n l; full n r\<rbrakk> \<Longrightarrow> full (Suc n) (Branch2 l p r)" |  | 
|
162  | 
"\<lbrakk>full n l; full n m; full n r\<rbrakk> \<Longrightarrow> full (Suc n) (Branch3 l p m q r)"  | 
|
163  | 
||
164  | 
inductive_cases full_elims:  | 
|
165  | 
"full n Empty"  | 
|
166  | 
"full n (Branch2 l p r)"  | 
|
167  | 
"full n (Branch3 l p m q r)"  | 
|
168  | 
||
169  | 
inductive_cases full_0_elim: "full 0 t"  | 
|
170  | 
inductive_cases full_Suc_elim: "full (Suc n) t"  | 
|
171  | 
||
172  | 
lemma full_0_iff [simp]: "full 0 t \<longleftrightarrow> t = Empty"  | 
|
173  | 
by (auto elim: full_0_elim intro: full.intros)  | 
|
174  | 
||
175  | 
lemma full_Empty_iff [simp]: "full n Empty \<longleftrightarrow> n = 0"  | 
|
176  | 
by (auto elim: full_elims intro: full.intros)  | 
|
177  | 
||
178  | 
lemma full_Suc_Branch2_iff [simp]:  | 
|
179  | 
"full (Suc n) (Branch2 l p r) \<longleftrightarrow> full n l \<and> full n r"  | 
|
180  | 
by (auto elim: full_elims intro: full.intros)  | 
|
181  | 
||
182  | 
lemma full_Suc_Branch3_iff [simp]:  | 
|
183  | 
"full (Suc n) (Branch3 l p m q r) \<longleftrightarrow> full n l \<and> full n m \<and> full n r"  | 
|
184  | 
by (auto elim: full_elims intro: full.intros)  | 
|
185  | 
||
| 33436 | 186  | 
fun height :: "'a tree23 \<Rightarrow> nat" where  | 
187  | 
"height Empty = 0" |  | 
|
188  | 
"height (Branch2 l _ r) = Suc(max (height l) (height r))" |  | 
|
189  | 
"height (Branch3 l _ m _ r) = Suc(max (height l) (max (height m) (height r)))"  | 
|
190  | 
||
| 61343 | 191  | 
text\<open>Is a tree balanced?\<close>  | 
| 33436 | 192  | 
fun bal :: "'a tree23 \<Rightarrow> bool" where  | 
193  | 
"bal Empty = True" |  | 
|
194  | 
"bal (Branch2 l _ r) = (bal l & bal r & height l = height r)" |  | 
|
195  | 
"bal (Branch3 l _ m _ r) = (bal l & bal m & bal r & height l = height m & height m = height r)"  | 
|
196  | 
||
| 45334 | 197  | 
lemma full_imp_height: "full n t \<Longrightarrow> height t = n"  | 
198  | 
by (induct set: full, simp_all)  | 
|
199  | 
||
200  | 
lemma full_imp_bal: "full n t \<Longrightarrow> bal t"  | 
|
201  | 
by (induct set: full, auto dest: full_imp_height)  | 
|
202  | 
||
203  | 
lemma bal_imp_full: "bal t \<Longrightarrow> full (height t) t"  | 
|
204  | 
by (induct t, simp_all)  | 
|
205  | 
||
206  | 
lemma bal_iff_full: "bal t \<longleftrightarrow> (\<exists>n. full n t)"  | 
|
207  | 
by (auto elim!: bal_imp_full full_imp_bal)  | 
|
208  | 
||
| 69597 | 209  | 
text \<open>The \<^term>\<open>add0\<close> function either preserves the height of the  | 
210  | 
tree, or increases it by one. The constructor returned by the \<^term>\<open>add\<close> function determines which: A return value of the form \<^term>\<open>Stay t\<close> indicates that the height will be the same. A value of the  | 
|
211  | 
form \<^term>\<open>Sprout l p r\<close> indicates an increase in height.\<close>  | 
|
| 
45325
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
212  | 
|
| 45334 | 213  | 
primrec gfull :: "nat \<Rightarrow> 'a growth \<Rightarrow> bool" where  | 
214  | 
"gfull n (Stay t) \<longleftrightarrow> full n t" |  | 
|
215  | 
"gfull n (Sprout l p r) \<longleftrightarrow> full n l \<and> full n r"  | 
|
| 
45325
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
216  | 
|
| 45334 | 217  | 
lemma gfull_add: "full n t \<Longrightarrow> gfull n (add k y t)"  | 
| 45336 | 218  | 
by (induct set: full, auto split: ord.split growth.split)  | 
| 
45325
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
219  | 
|
| 69597 | 220  | 
text \<open>The \<^term>\<open>add0\<close> operation preserves balance.\<close>  | 
| 
45325
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
221  | 
|
| 45334 | 222  | 
lemma bal_add0: "bal t \<Longrightarrow> bal (add0 k y t)"  | 
223  | 
unfolding bal_iff_full add0_def  | 
|
224  | 
apply (erule exE)  | 
|
225  | 
apply (drule gfull_add [of _ _ k y])  | 
|
226  | 
apply (cases "add k y t")  | 
|
227  | 
apply (auto intro: full.intros)  | 
|
228  | 
done  | 
|
229  | 
||
| 69597 | 230  | 
text \<open>The \<^term>\<open>add0\<close> operation preserves order.\<close>  | 
| 45334 | 231  | 
|
232  | 
lemma ord_cases:  | 
|
233  | 
fixes a b :: int obtains  | 
|
234  | 
"ord a b = LESS" and "a < b" |  | 
|
235  | 
"ord a b = EQUAL" and "a = b" |  | 
|
236  | 
"ord a b = GREATER" and "a > b"  | 
|
237  | 
unfolding ord_def by (rule linorder_cases [of a b]) auto  | 
|
| 
45325
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
238  | 
|
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
239  | 
definition gtree :: "'a growth \<Rightarrow> 'a tree23" where  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
240  | 
"gtree g = (case g of Stay t \<Rightarrow> t | Sprout l p r \<Rightarrow> Branch2 l p r)"  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
241  | 
|
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
242  | 
lemma gtree_simps [simp]:  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
243  | 
"gtree (Stay t) = t"  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
244  | 
"gtree (Sprout l p r) = Branch2 l p r"  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
245  | 
unfolding gtree_def by simp_all  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
246  | 
|
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
247  | 
lemma add0: "add0 k y t = gtree (add k y t)"  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
248  | 
unfolding add0_def by (simp split: growth.split)  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
249  | 
|
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
250  | 
lemma ord'_add0:  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
251  | 
"\<lbrakk>ord' k1 t k2; opt_less k1 (Some k); opt_less (Some k) k2\<rbrakk> \<Longrightarrow> ord' k1 (add0 k y t) k2"  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
252  | 
unfolding add0  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
253  | 
apply (induct t arbitrary: k1 k2)  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
254  | 
apply simp  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
255  | 
apply clarsimp  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
256  | 
apply (rule_tac a=k and b=a in ord_cases)  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
257  | 
apply simp  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
258  | 
apply (case_tac "add k y t1", simp, simp)  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
259  | 
apply simp  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
260  | 
apply simp  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
261  | 
apply (case_tac "add k y t2", simp, simp)  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
262  | 
apply clarsimp  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
263  | 
apply (rule_tac a=k and b=a in ord_cases)  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
264  | 
apply simp  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
265  | 
apply (case_tac "add k y t1", simp, simp)  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
266  | 
apply simp  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
267  | 
apply simp  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
268  | 
apply (rule_tac a=k and b=aa in ord_cases)  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
269  | 
apply simp  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
270  | 
apply (case_tac "add k y t2", simp, simp)  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
271  | 
apply simp  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
272  | 
apply simp  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
273  | 
apply (case_tac "add k y t3", simp, simp)  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
274  | 
done  | 
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
275  | 
|
| 
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
276  | 
lemma ord0_add0: "ord0 t \<Longrightarrow> ord0 (add0 k y t)"  | 
| 
45333
 
04b21922ed68
ex/Tree23.thy: simpler definition of ordered-ness predicate
 
huffman 
parents: 
45332 
diff
changeset
 | 
277  | 
by (simp add: ord0_def ord'_add0)  | 
| 
45325
 
26b6179b5a45
ex/Tree23.thy: prove that insertion preserves tree balance and order
 
huffman 
parents: 
42463 
diff
changeset
 | 
278  | 
|
| 69597 | 279  | 
text \<open>The \<^term>\<open>del\<close> function preserves balance.\<close>  | 
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
280  | 
|
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
281  | 
lemma del_extra_simps:  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
282  | 
"l \<noteq> Empty \<or> r \<noteq> Empty \<Longrightarrow>  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
283  | 
del k (Branch2 l p r) = (case compare k p of  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
284  | 
LESS => (case del k l of None \<Rightarrow> None |  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
285  | 
Some(p', (False, l')) => Some(p', (False, Branch2 l' p r))  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
286  | 
| Some(p', (True, l')) => Some(p', case r of  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
287  | 
Branch2 rl rp rr => (True, Branch3 l' p rl rp rr)  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
288  | 
| Branch3 rl rp rm rq rr => (False, Branch2  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
289  | 
(Branch2 l' p rl) rp (Branch2 rm rq rr))))  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
290  | 
| or => (case del (if_eq or None k) r of None \<Rightarrow> None |  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
291  | 
Some(p', (False, r')) => Some(p', (False, Branch2 l (if_eq or p' p) r'))  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
292  | 
| Some(p', (True, r')) => Some(p', case l of  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
293  | 
Branch2 ll lp lr => (True, Branch3 ll lp lr (if_eq or p' p) r')  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
294  | 
| Branch3 ll lp lm lq lr => (False, Branch2  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
295  | 
(Branch2 ll lp lm) lq (Branch2 lr (if_eq or p' p) r')))))"  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
296  | 
"l \<noteq> Empty \<or> m \<noteq> Empty \<or> r \<noteq> Empty \<Longrightarrow>  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
297  | 
del k (Branch3 l p m q r) = (case compare k q of  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
298  | 
LESS => (case compare k p of  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
299  | 
LESS => (case del k l of None \<Rightarrow> None |  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
300  | 
Some(p', (False, l')) => Some(p', (False, Branch3 l' p m q r))  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
301  | 
| Some(p', (True, l')) => Some(p', (False, case (m, r) of  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
302  | 
(Branch2 ml mp mr, Branch2 _ _ _) => Branch2 (Branch3 l' p ml mp mr) q r  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
303  | 
| (Branch3 ml mp mm mq mr, _) => Branch3 (Branch2 l' p ml) mp (Branch2 mm mq mr) q r  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
304  | 
| (Branch2 ml mp mr, Branch3 rl rp rm rq rr) =>  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
305  | 
Branch3 (Branch2 l' p ml) mp (Branch2 mr q rl) rp (Branch2 rm rq rr))))  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
306  | 
| or => (case del (if_eq or None k) m of None \<Rightarrow> None |  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
307  | 
Some(p', (False, m')) => Some(p', (False, Branch3 l (if_eq or p' p) m' q r))  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
308  | 
| Some(p', (True, m')) => Some(p', (False, case (l, r) of  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
309  | 
(Branch2 ll lp lr, Branch2 _ _ _) => Branch2 (Branch3 ll lp lr (if_eq or p' p) m') q r  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
310  | 
| (Branch3 ll lp lm lq lr, _) => Branch3 (Branch2 ll lp lm) lq (Branch2 lr (if_eq or p' p) m') q r  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
311  | 
| (_, Branch3 rl rp rm rq rr) => Branch3 l (if_eq or p' p) (Branch2 m' q rl) rp (Branch2 rm rq rr)))))  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
312  | 
| or => (case del (if_eq or None k) r of None \<Rightarrow> None |  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
313  | 
Some(q', (False, r')) => Some(q', (False, Branch3 l p m (if_eq or q' q) r'))  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
314  | 
| Some(q', (True, r')) => Some(q', (False, case (l, m) of  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
315  | 
(Branch2 _ _ _, Branch2 ml mp mr) => Branch2 l p (Branch3 ml mp mr (if_eq or q' q) r')  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
316  | 
| (_, Branch3 ml mp mm mq mr) => Branch3 l p (Branch2 ml mp mm) mq (Branch2 mr (if_eq or q' q) r')  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
317  | 
| (Branch3 ll lp lm lq lr, Branch2 ml mp mr) =>  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
318  | 
Branch3 (Branch2 ll lp lm) lq (Branch2 lr p ml) mp (Branch2 mr (if_eq or q' q) r')))))"  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
319  | 
apply -  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
320  | 
apply (cases l, cases r, simp_all only: del.simps, simp)  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
321  | 
apply (cases l, cases m, cases r, simp_all only: del.simps, simp)  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
322  | 
done  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
323  | 
|
| 45335 | 324  | 
fun dfull where  | 
325  | 
"dfull n None \<longleftrightarrow> True" |  | 
|
326  | 
"dfull n (Some (p, (True, t'))) \<longleftrightarrow> full n t'" |  | 
|
327  | 
"dfull n (Some (p, (False, t'))) \<longleftrightarrow> full (Suc n) t'"  | 
|
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
328  | 
|
| 45335 | 329  | 
lemmas dfull_case_intros =  | 
| 
55417
 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 
blanchet 
parents: 
55414 
diff
changeset
 | 
330  | 
ord.exhaust [of y "dfull a (case_ord b c d y)"]  | 
| 
55413
 
a8e96847523c
adapted theories to '{case,rec}_{list,option}' names
 
blanchet 
parents: 
45605 
diff
changeset
 | 
331  | 
option.exhaust [of y "dfull a (case_option b c y)"]  | 
| 
55414
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55413 
diff
changeset
 | 
332  | 
prod.exhaust [of y "dfull a (case_prod b y)"]  | 
| 
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55413 
diff
changeset
 | 
333  | 
bool.exhaust [of y "dfull a (case_bool b c y)"]  | 
| 
55417
 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 
blanchet 
parents: 
55414 
diff
changeset
 | 
334  | 
tree23.exhaust [of y "dfull a (Some (b, case_tree23 c d e y))"]  | 
| 
 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 
blanchet 
parents: 
55414 
diff
changeset
 | 
335  | 
tree23.exhaust [of y "full a (case_tree23 b c d y)"]  | 
| 45605 | 336  | 
for a b c d e y  | 
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
337  | 
|
| 45335 | 338  | 
lemma dfull_del: "full (Suc n) t \<Longrightarrow> dfull n (del k t)"  | 
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
339  | 
proof -  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
340  | 
  { fix n :: "nat" and p :: "key \<times> 'a" and l r :: "'a tree23" and k
 | 
| 45335 | 341  | 
assume "\<And>n. \<lbrakk>compare k p = LESS; full (Suc n) l\<rbrakk> \<Longrightarrow> dfull n (del k l)"  | 
342  | 
and "\<And>n. \<lbrakk>compare k p = EQUAL; full (Suc n) r\<rbrakk> \<Longrightarrow> dfull n (del (if_eq EQUAL None k) r)"  | 
|
343  | 
and "\<And>n. \<lbrakk>compare k p = GREATER; full (Suc n) r\<rbrakk> \<Longrightarrow> dfull n (del (if_eq GREATER None k) r)"  | 
|
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
344  | 
and "full (Suc n) (Branch2 l p r)"  | 
| 45335 | 345  | 
hence "dfull n (del k (Branch2 l p r))"  | 
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
346  | 
apply clarsimp  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
347  | 
apply (cases n)  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
348  | 
apply (cases k)  | 
| 45335 | 349  | 
apply simp  | 
350  | 
apply (simp split: ord.split)  | 
|
351  | 
apply simp  | 
|
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
352  | 
apply (subst del_extra_simps, force)  | 
| 45336 | 353  | 
(* This should work, but it is way too slow!  | 
354  | 
apply (force split: ord.split option.split bool.split tree23.split) *)  | 
|
| 45335 | 355  | 
apply (simp | rule dfull_case_intros)+  | 
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
356  | 
done  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
357  | 
} note A = this  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
358  | 
  { fix n :: "nat" and p q :: "key \<times> 'a" and l m r :: "'a tree23" and k
 | 
| 45335 | 359  | 
assume "\<And>n. \<lbrakk>compare k q = LESS; compare k p = LESS; full (Suc n) l\<rbrakk> \<Longrightarrow> dfull n (del k l)"  | 
360  | 
and "\<And>n. \<lbrakk>compare k q = LESS; compare k p = EQUAL; full (Suc n) m\<rbrakk> \<Longrightarrow> dfull n (del (if_eq EQUAL None k) m)"  | 
|
361  | 
and "\<And>n. \<lbrakk>compare k q = LESS; compare k p = GREATER; full (Suc n) m\<rbrakk> \<Longrightarrow> dfull n (del (if_eq GREATER None k) m)"  | 
|
362  | 
and "\<And>n. \<lbrakk>compare k q = EQUAL; full (Suc n) r\<rbrakk> \<Longrightarrow> dfull n (del (if_eq EQUAL None k) r)"  | 
|
363  | 
and "\<And>n. \<lbrakk>compare k q = GREATER; full (Suc n) r\<rbrakk> \<Longrightarrow> dfull n (del (if_eq GREATER None k) r)"  | 
|
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
364  | 
and "full (Suc n) (Branch3 l p m q r)"  | 
| 45335 | 365  | 
hence "dfull n (del k (Branch3 l p m q r))"  | 
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
366  | 
apply clarsimp  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
367  | 
apply (cases n)  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
368  | 
apply (cases k)  | 
| 45335 | 369  | 
apply simp  | 
370  | 
apply (simp split: ord.split)  | 
|
371  | 
apply simp  | 
|
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
372  | 
apply (subst del_extra_simps, force)  | 
| 45335 | 373  | 
apply (simp | rule dfull_case_intros)+  | 
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
374  | 
done  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
375  | 
} note B = this  | 
| 45335 | 376  | 
show "full (Suc n) t \<Longrightarrow> dfull n (del k t)"  | 
| 
61166
 
5976fe402824
renamed method "goals" to "goal_cases" to emphasize its meaning;
 
wenzelm 
parents: 
60580 
diff
changeset
 | 
377  | 
proof (induct k t arbitrary: n rule: del.induct, goal_cases)  | 
| 60580 | 378  | 
case (1 k n)  | 
379  | 
thus "dfull n (del (Some k) Empty)" by simp  | 
|
380  | 
next  | 
|
381  | 
case (2 p n)  | 
|
382  | 
thus "dfull n (del None (Branch2 Empty p Empty))" by simp  | 
|
383  | 
next  | 
|
384  | 
case (3 p q n)  | 
|
385  | 
thus "dfull n (del None (Branch3 Empty p Empty q Empty))" by simp  | 
|
386  | 
next  | 
|
387  | 
case (4 v p n)  | 
|
388  | 
thus "dfull n (del (Some v) (Branch2 Empty p Empty))"  | 
|
389  | 
by (simp split: ord.split)  | 
|
390  | 
next  | 
|
391  | 
case (5 v p q n)  | 
|
392  | 
thus "dfull n (del (Some v) (Branch3 Empty p Empty q Empty))"  | 
|
393  | 
by (simp split: ord.split)  | 
|
394  | 
next  | 
|
395  | 
case (26 n)  | 
|
396  | 
thus ?case by simp  | 
|
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
397  | 
qed (fact A | fact B)+  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
398  | 
qed  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
399  | 
|
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
400  | 
lemma bal_del0: "bal t \<Longrightarrow> bal (del0 k t)"  | 
| 45335 | 401  | 
unfolding bal_iff_full del0_def  | 
402  | 
apply (erule exE)  | 
|
403  | 
apply (case_tac n, simp, simp)  | 
|
404  | 
apply (frule dfull_del [where k="Some k"])  | 
|
405  | 
apply (cases "del (Some k) t", force)  | 
|
| 
55417
 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 
blanchet 
parents: 
55414 
diff
changeset
 | 
406  | 
apply (rename_tac a, case_tac "a", rename_tac b t', case_tac "b", auto)  | 
| 45335 | 407  | 
done  | 
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
408  | 
|
| 61343 | 409  | 
text\<open>This is a little test harness and should be commented out once the  | 
410  | 
above functions have been proved correct.\<close>  | 
|
| 33436 | 411  | 
|
| 58310 | 412  | 
datatype 'a act = Add int 'a | Del int  | 
| 33436 | 413  | 
|
414  | 
fun exec where  | 
|
415  | 
"exec [] t = t" |  | 
|
416  | 
"exec (Add k x # as) t = exec as (add0 k x t)" |  | 
|
417  | 
"exec (Del k # as) t = exec as (del0 k t)"  | 
|
418  | 
||
| 61343 | 419  | 
text\<open>Some quick checks:\<close>  | 
| 33436 | 420  | 
|
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
421  | 
lemma bal_exec: "bal t \<Longrightarrow> bal (exec as t)"  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
422  | 
by (induct as t arbitrary: t rule: exec.induct,  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
423  | 
simp_all add: bal_add0 bal_del0)  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
424  | 
|
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
425  | 
lemma "bal(exec as Empty)"  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
426  | 
by (simp add: bal_exec)  | 
| 
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
427  | 
|
| 33436 | 428  | 
lemma "ord0(exec as Empty)"  | 
429  | 
quickcheck  | 
|
430  | 
oops  | 
|
431  | 
||
| 
45332
 
ede9dc025150
ex/Tree23.thy: prove that deletion preserves balance
 
huffman 
parents: 
45325 
diff
changeset
 | 
432  | 
end  |