src/HOL/Data_Structures/Tree_Set.thy
author haftmann
Mon, 06 Feb 2017 20:56:34 +0100
changeset 64990 c6a7de505796
parent 63411 e051eea34990
child 66453 cc19f7ca2ed6
permissions -rw-r--r--
more explicit errors in pathological cases
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
     1
(* Author: Tobias Nipkow *)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
     2
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 61678
diff changeset
     3
section \<open>Unbalanced Tree Implementation of Set\<close>
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
     4
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
     5
theory Tree_Set
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
     6
imports
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
     7
  "~~/src/HOL/Library/Tree"
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
     8
  Cmp
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
     9
  Set_by_Ordered
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    10
begin
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    11
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 61678
diff changeset
    12
fun isin :: "'a::linorder tree \<Rightarrow> 'a \<Rightarrow> bool" where
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    13
"isin Leaf x = False" |
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    14
"isin (Node l a r) x =
61678
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    15
  (case cmp x a of
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    16
     LT \<Rightarrow> isin l x |
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    17
     EQ \<Rightarrow> True |
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    18
     GT \<Rightarrow> isin r x)"
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    19
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    20
hide_const (open) insert
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    21
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 61678
diff changeset
    22
fun insert :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    23
"insert x Leaf = Node Leaf x Leaf" |
61678
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    24
"insert x (Node l a r) =
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    25
  (case cmp x a of
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    26
     LT \<Rightarrow> Node (insert x l) a r |
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    27
     EQ \<Rightarrow> Node l a r |
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    28
     GT \<Rightarrow> Node l a (insert x r))"
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    29
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    30
fun del_min :: "'a tree \<Rightarrow> 'a * 'a tree" where
61678
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    31
"del_min (Node l a r) =
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    32
  (if l = Leaf then (a,r) else let (x,l') = del_min l in (x, Node l' a r))"
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    33
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 61678
diff changeset
    34
fun delete :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    35
"delete x Leaf = Leaf" |
61678
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    36
"delete x (Node l a r) =
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    37
  (case cmp x a of
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    38
     LT \<Rightarrow>  Node (delete x l) a r |
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    39
     GT \<Rightarrow>  Node l a (delete x r) |
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    40
     EQ \<Rightarrow> if r = Leaf then l else let (a',r') = del_min r in Node l a' r')"
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    41
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    42
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    43
subsection "Functional Correctness Proofs"
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    44
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    45
lemma "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    46
by (induction t) (auto simp: elems_simps1)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    47
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    48
lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    49
by (induction t) (auto simp: elems_simps2)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    50
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    51
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    52
lemma inorder_insert:
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    53
  "sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)"
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    54
by(induction t) (auto simp: ins_list_simps)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    55
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    56
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    57
lemma del_minD:
61651
415e816d3f37 unnecessary precondition
nipkow
parents: 61647
diff changeset
    58
  "del_min t = (x,t') \<Longrightarrow> t \<noteq> Leaf \<Longrightarrow> x # inorder t' = inorder t"
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    59
by(induction t arbitrary: t' rule: del_min.induct)
61647
nipkow
parents: 61640
diff changeset
    60
  (auto simp: sorted_lems split: prod.splits if_splits)
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    61
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    62
lemma inorder_delete:
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    63
  "sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    64
by(induction t) (auto simp: del_list_simps del_minD split: prod.splits)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    65
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    66
interpretation Set_by_Ordered
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    67
where empty = Leaf and isin = isin and insert = insert and delete = delete
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    68
and inorder = inorder and inv = "\<lambda>_. True"
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    69
proof (standard, goal_cases)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    70
  case 1 show ?case by simp
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    71
next
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    72
  case 2 thus ?case by(simp add: isin_set)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    73
next
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    74
  case 3 thus ?case by(simp add: inorder_insert)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    75
next
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    76
  case 4 thus ?case by(simp add: inorder_delete)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    77
qed (rule TrueI)+
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    78
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    79
end