src/HOL/Data_Structures/AVL_Map.thy
author nipkow
Wed, 23 Sep 2015 09:47:04 +0200
changeset 61232 c46faf9762f7
child 61581 00d9682e8dd7
permissions -rw-r--r--
added AVL and lookup function
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61232
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
     2
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
     3
section "AVL Tree Implementation of Maps"
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
     4
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
     5
theory AVL_Map
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
     6
imports
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
     7
  AVL_Set
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
     8
  Lookup2
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
     9
begin
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    10
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    11
fun update :: "'a::order \<Rightarrow> 'b \<Rightarrow> ('a*'b) avl_tree \<Rightarrow> ('a*'b) avl_tree" where
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    12
"update x y Leaf = Node 1 Leaf (x,y) Leaf" |
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    13
"update x y (Node h l (a,b) r) = 
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    14
   (if x = a then Node h l (x,y) r else
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    15
    if x < a then node_bal_l (update x y l) (a,b) r
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    16
    else node_bal_r l (a,b) (update x y r))"
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    17
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    18
fun delete :: "'a::order \<Rightarrow> ('a*'b) avl_tree \<Rightarrow> ('a*'b) avl_tree" where
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    19
"delete _ Leaf = Leaf" |
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    20
"delete x (Node h l (a,b) r) = (
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    21
   if x = a then delete_root (Node h l (a,b) r) else
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    22
   if x < a then node_bal_r (delete x l) (a,b) r
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    23
   else node_bal_l l (a,b) (delete x r))"
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    24
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    25
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    26
subsection {* Functional Correctness Proofs *}
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    27
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    28
theorem inorder_update:
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    29
  "sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    30
by (induct t) 
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    31
   (auto simp: upd_list_simps inorder_node_bal_l inorder_node_bal_r)
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    32
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    33
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    34
theorem inorder_delete:
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    35
  "sorted1(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    36
by(induction t)
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    37
  (auto simp: del_list_simps inorder_node_bal_l inorder_node_bal_r
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    38
     inorder_delete_root inorder_delete_maxD split: prod.splits)
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    39
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    40
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    41
interpretation Map_by_Ordered
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    42
where empty = Leaf and lookup = lookup and update = update and delete = delete
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    43
and inorder = inorder and wf = "\<lambda>_. True"
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    44
proof (standard, goal_cases)
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    45
  case 1 show ?case by simp
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    46
next
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    47
  case 2 thus ?case by(simp add: lookup_eq)
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    48
next
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    49
  case 3 thus ?case by(simp add: inorder_update)
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    50
next
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    51
  case 4 thus ?case by(simp add: inorder_delete)
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    52
qed (rule TrueI)+
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    53
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    54
end