18 "update a b (Node l (x,y) r) = |
18 "update a b (Node l (x,y) r) = |
19 (if a < x then Node (update a b l) (x,y) r |
19 (if a < x then Node (update a b l) (x,y) r |
20 else if a=x then Node l (a,b) r |
20 else if a=x then Node l (a,b) r |
21 else Node l (x,y) (update a b r))" |
21 else Node l (x,y) (update a b r))" |
22 |
22 |
23 fun del_min :: "'a tree \<Rightarrow> 'a * 'a tree" where |
|
24 "del_min (Node Leaf a r) = (a, r)" | |
|
25 "del_min (Node l a r) = (let (x,l') = del_min l in (x, Node l' a r))" |
|
26 |
|
27 fun delete :: "'a::linorder \<Rightarrow> ('a*'b) tree \<Rightarrow> ('a*'b) tree" where |
23 fun delete :: "'a::linorder \<Rightarrow> ('a*'b) tree \<Rightarrow> ('a*'b) tree" where |
28 "delete k Leaf = Leaf" | |
24 "delete k Leaf = Leaf" | |
29 "delete k (Node l (a,b) r) = (if k<a then Node (delete k l) (a,b) r else |
25 "delete k (Node l (a,b) r) = (if k<a then Node (delete k l) (a,b) r else |
30 if k > a then Node l (a,b) (delete k r) else |
26 if k > a then Node l (a,b) (delete k r) else |
31 if r = Leaf then l else let (ab',r') = del_min r in Node l ab' r')" |
27 if r = Leaf then l else let (ab',r') = del_min r in Node l ab' r')" |
33 |
29 |
34 subsection "Functional Correctness Proofs" |
30 subsection "Functional Correctness Proofs" |
35 |
31 |
36 lemma lookup_eq: |
32 lemma lookup_eq: |
37 "sorted1(inorder t) \<Longrightarrow> lookup t x = map_of (inorder t) x" |
33 "sorted1(inorder t) \<Longrightarrow> lookup t x = map_of (inorder t) x" |
38 apply (induction t) |
34 by (induction t) (auto simp: map_of_simps split: option.split) |
39 apply (auto simp: sorted_lems map_of_append map_of_sorteds split: option.split) |
|
40 done |
|
41 |
35 |
42 |
36 |
43 lemma inorder_update: |
37 lemma inorder_update: |
44 "sorted1(inorder t) \<Longrightarrow> inorder(update a b t) = upd_list a b (inorder t)" |
38 "sorted1(inorder t) \<Longrightarrow> inorder(update a b t) = upd_list a b (inorder t)" |
45 by(induction t) (auto simp: upd_list_simps) |
39 by(induction t) (auto simp: upd_list_simps) |
47 |
41 |
48 lemma del_minD: |
42 lemma del_minD: |
49 "del_min t = (x,t') \<Longrightarrow> t \<noteq> Leaf \<Longrightarrow> sorted1(inorder t) \<Longrightarrow> |
43 "del_min t = (x,t') \<Longrightarrow> t \<noteq> Leaf \<Longrightarrow> sorted1(inorder t) \<Longrightarrow> |
50 x # inorder t' = inorder t" |
44 x # inorder t' = inorder t" |
51 by(induction t arbitrary: t' rule: del_min.induct) |
45 by(induction t arbitrary: t' rule: del_min.induct) |
52 (auto simp: sorted_lems split: prod.splits) |
46 (auto simp: del_list_simps split: prod.splits) |
53 |
47 |
54 lemma inorder_delete: |
48 lemma inorder_delete: |
55 "sorted1(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)" |
49 "sorted1(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)" |
56 by(induction t) |
50 by(induction t) (auto simp: del_list_simps del_minD split: prod.splits) |
57 (auto simp: del_list_sorted sorted_lems dest!: del_minD split: prod.splits) |
|
58 |
51 |
59 |
52 |
60 interpretation Map_by_Ordered |
53 interpretation Map_by_Ordered |
61 where empty = Leaf and lookup = lookup and update = update and delete = delete |
54 where empty = Leaf and lookup = lookup and update = update and delete = delete |
62 and inorder = inorder and wf = "\<lambda>_. True" |
55 and inorder = inorder and wf = "\<lambda>_. True" |