src/HOL/Data_Structures/Tree_Set.thy
author nipkow
Thu, 05 Nov 2015 18:38:08 +0100
changeset 61588 1d2907d0ed73
parent 61581 00d9682e8dd7
child 61640 44c9198f210c
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     2
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     3
section {* Tree Implementation of Sets *}
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     4
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     5
theory Tree_Set
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     6
imports
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     7
  "~~/src/HOL/Library/Tree"
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
     8
  Cmp
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     9
  Set_by_Ordered
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    10
begin
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    11
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    12
fun isin :: "'a::cmp tree \<Rightarrow> 'a \<Rightarrow> bool" where
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    13
"isin Leaf x = False" |
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    14
"isin (Node l a r) x =
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    15
  (case cmp x a of LT \<Rightarrow> isin l x | EQ \<Rightarrow> True | GT \<Rightarrow> isin r x)"
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    16
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    17
hide_const (open) insert
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    18
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    19
fun insert :: "'a::cmp \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
61534
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    20
"insert x Leaf = Node Leaf x Leaf" |
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    21
"insert x (Node l a r) = (case cmp x a of
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    22
      LT \<Rightarrow> Node (insert x l) a r |
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    23
      EQ \<Rightarrow> Node l a r |
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    24
      GT \<Rightarrow> Node l a (insert x r))"
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    25
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    26
fun del_min :: "'a tree \<Rightarrow> 'a * 'a tree" where
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    27
"del_min (Node Leaf a r) = (a, r)" |
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    28
"del_min (Node l a r) = (let (x,l') = del_min l in (x, Node l' a r))"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    29
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    30
fun delete :: "'a::cmp \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
61534
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    31
"delete x Leaf = Leaf" |
61581
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    32
"delete x (Node l a r) = (case cmp x a of
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    33
  LT \<Rightarrow>  Node (delete x l) a r |
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    34
  GT \<Rightarrow>  Node l a (delete x r) |
00d9682e8dd7 Convertd to 3-way comparisons
nipkow
parents: 61534
diff changeset
    35
  EQ \<Rightarrow> if r = Leaf then l else let (a',r') = del_min r in Node l a' r')"
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    36
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    37
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    38
subsection "Functional Correctness Proofs"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    39
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    40
lemma "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
61229
0b9c45c4af29 unified isin-proofs
nipkow
parents: 61203
diff changeset
    41
by (induction t) (auto simp: elems_simps1)
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    42
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    43
lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
61229
0b9c45c4af29 unified isin-proofs
nipkow
parents: 61203
diff changeset
    44
by (induction t) (auto simp: elems_simps2)
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    45
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    46
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    47
lemma inorder_insert:
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    48
  "sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)"
61231
nipkow
parents: 61229
diff changeset
    49
by(induction t) (auto simp: ins_list_simps)
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    50
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    51
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    52
lemma del_minD:
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    53
  "del_min t = (x,t') \<Longrightarrow> t \<noteq> Leaf \<Longrightarrow> sorted(inorder t) \<Longrightarrow>
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    54
   x # inorder t' = inorder t"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    55
by(induction t arbitrary: t' rule: del_min.induct)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    56
  (auto simp: sorted_lems split: prod.splits)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    57
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    58
lemma inorder_delete:
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    59
  "sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
61231
nipkow
parents: 61229
diff changeset
    60
by(induction t) (auto simp: del_list_simps del_minD split: prod.splits)
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    61
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    62
interpretation Set_by_Ordered
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    63
where empty = Leaf and isin = isin and insert = insert and delete = delete
61588
nipkow
parents: 61581
diff changeset
    64
and inorder = inorder and inv = "\<lambda>_. True"
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    65
proof (standard, goal_cases)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    66
  case 1 show ?case by simp
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    67
next
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    68
  case 2 thus ?case by(simp add: isin_set)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    69
next
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    70
  case 3 thus ?case by(simp add: inorder_insert)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    71
next
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    72
  case 4 thus ?case by(simp add: inorder_delete)
61428
5e1938107371 added invar empty
nipkow
parents: 61231
diff changeset
    73
qed (rule TrueI)+
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    74
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    75
end