src/HOL/Data_Structures/AVL_Map.thy
author nipkow
Mon, 16 Nov 2015 13:08:52 +0100
changeset 61686 e6784939d645
parent 61648 f7662ca95f1b
child 61790 0494964bb226
permissions -rw-r--r--
tuned names
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
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61232
diff changeset
    11
fun update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) avl_tree \<Rightarrow> ('a*'b) avl_tree" where
61232
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    12
"update x y Leaf = Node 1 Leaf (x,y) Leaf" |
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61232
diff changeset
    13
"update x y (Node h l (a,b) r) = (case cmp x a of
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61232
diff changeset
    14
   EQ \<Rightarrow> Node h l (x,y) r |
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61232
diff changeset
    15
   LT \<Rightarrow> balL (update x y l) (a,b) r |
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61232
diff changeset
    16
   GT \<Rightarrow> balR l (a,b) (update x y r))"
61232
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    17
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61232
diff changeset
    18
fun delete :: "'a::cmp \<Rightarrow> ('a*'b) avl_tree \<Rightarrow> ('a*'b) avl_tree" where
61232
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    19
"delete _ Leaf = Leaf" |
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61232
diff changeset
    20
"delete x (Node h l (a,b) r) = (case cmp x a of
61648
f7662ca95f1b tuned name
nipkow
parents: 61581
diff changeset
    21
   EQ \<Rightarrow> del_root (Node h l (a,b) r) |
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61232
diff changeset
    22
   LT \<Rightarrow> balR (delete x l) (a,b) r |
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61232
diff changeset
    23
   GT \<Rightarrow> balL l (a,b) (delete x r))"
61232
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)"
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61232
diff changeset
    30
by (induct t) (auto simp: upd_list_simps inorder_balL inorder_balR)
61232
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    31
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    32
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    33
theorem inorder_delete:
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    34
  "sorted1(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    35
by(induction t)
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61232
diff changeset
    36
  (auto simp: del_list_simps inorder_balL inorder_balR
61648
f7662ca95f1b tuned name
nipkow
parents: 61581
diff changeset
    37
     inorder_del_root inorder_del_maxD split: prod.splits)
61232
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    38
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    39
interpretation Map_by_Ordered
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    40
where empty = Leaf and lookup = lookup and update = update and delete = delete
61686
e6784939d645 tuned names
nipkow
parents: 61648
diff changeset
    41
and inorder = inorder and inv = "\<lambda>_. True"
61232
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    42
proof (standard, goal_cases)
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    43
  case 1 show ?case by simp
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    44
next
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    45
  case 2 thus ?case by(simp add: lookup_eq)
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    46
next
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    47
  case 3 thus ?case by(simp add: inorder_update)
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    48
next
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    49
  case 4 thus ?case by(simp add: inorder_delete)
61686
e6784939d645 tuned names
nipkow
parents: 61648
diff changeset
    50
qed auto
61232
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    51
c46faf9762f7 added AVL and lookup function
nipkow
parents:
diff changeset
    52
end