src/HOL/Data_Structures/AA_Map.thy
author wenzelm
Mon, 15 Feb 2016 14:55:44 +0100
changeset 62337 d3996d5873dd
parent 62130 90a3016a6c12
child 62496 f187aaf602c4
permissions -rw-r--r--
proper syntax;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62130
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
     2
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
     3
section "AA Implementation of Maps"
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
     4
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
     5
theory AA_Map
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
     6
imports
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
     7
  AA_Set
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
     8
  Lookup2
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
     9
begin
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    10
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    11
fun update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) aa_tree \<Rightarrow> ('a*'b) aa_tree" where
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    12
"update x y Leaf = Node 1 Leaf (x,y) Leaf" |
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    13
"update x y (Node lv t1 (a,b) t2) =
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    14
  (case cmp x a of
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    15
     LT \<Rightarrow> split (skew (Node lv (update x y t1) (a,b) t2)) |
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    16
     GT \<Rightarrow> split (skew (Node lv t1 (a,b) (update x y t2))) |
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    17
     EQ \<Rightarrow> Node lv t1 (x,y) t2)"
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    18
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    19
fun delete :: "'a::cmp \<Rightarrow> ('a*'b) aa_tree \<Rightarrow> ('a*'b) aa_tree" where
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    20
"delete _ Leaf = Leaf" |
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    21
"delete x (Node lv l (a,b) r) =
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    22
  (case cmp x a of
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    23
     LT \<Rightarrow> adjust (Node lv (delete x l) (a,b) r) |
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    24
     GT \<Rightarrow> adjust (Node lv l (a,b) (delete x r)) |
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    25
     EQ \<Rightarrow> (if l = Leaf then r
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    26
            else let (l',ab') = del_max l in adjust (Node lv l' ab' r)))"
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    27
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    28
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    29
subsection {* Functional Correctness Proofs *}
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    30
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    31
theorem inorder_update:
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    32
  "sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    33
by (induct t) (auto simp: upd_list_simps inorder_split inorder_skew)
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    34
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    35
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    36
theorem inorder_delete:
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    37
  "sorted1(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    38
by(induction t)
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    39
  (auto simp: del_list_simps inorder_adjust del_maxD split: prod.splits)
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    40
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    41
interpretation Map_by_Ordered
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    42
where empty = Leaf and lookup = lookup and update = update and delete = delete
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    43
and inorder = inorder and inv = "\<lambda>_. True"
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    44
proof (standard, goal_cases)
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    45
  case 1 show ?case by simp
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    46
next
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    47
  case 2 thus ?case by(simp add: lookup_map_of)
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    48
next
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    49
  case 3 thus ?case by(simp add: inorder_update)
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    50
next
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    51
  case 4 thus ?case by(simp add: inorder_delete)
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    52
qed auto
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    53
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents:
diff changeset
    54
end