src/HOL/Data_Structures/Tree_Set.thy
author nipkow
Tue, 12 May 2020 10:24:53 +0200
changeset 71829 6f2663df8374
parent 71463 a31a9da43694
permissions -rw-r--r--
"app" -> "join" for uniformity with Join theory; tuned defs
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
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 63411
diff changeset
     7
  "HOL-Library.Tree"
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
     8
  Cmp
67965
aaa31cd0caef more name tuning
nipkow
parents: 67964
diff changeset
     9
  Set_Specs
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    10
begin
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    11
68431
b294e095f64c more abstract naming
nipkow
parents: 68020
diff changeset
    12
definition empty :: "'a tree" where
71829
6f2663df8374 "app" -> "join" for uniformity with Join theory; tuned defs
nipkow
parents: 71463
diff changeset
    13
"empty = Leaf"
68431
b294e095f64c more abstract naming
nipkow
parents: 68020
diff changeset
    14
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 61678
diff changeset
    15
fun isin :: "'a::linorder tree \<Rightarrow> 'a \<Rightarrow> bool" where
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    16
"isin Leaf x = False" |
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    17
"isin (Node l a r) x =
61678
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    18
  (case cmp x a of
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    19
     LT \<Rightarrow> isin l x |
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    20
     EQ \<Rightarrow> True |
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    21
     GT \<Rightarrow> isin r x)"
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    22
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    23
hide_const (open) insert
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    24
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 61678
diff changeset
    25
fun insert :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    26
"insert x Leaf = Node Leaf x Leaf" |
61678
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    27
"insert x (Node l a r) =
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    28
  (case cmp x a of
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    29
     LT \<Rightarrow> Node (insert x l) a r |
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    30
     EQ \<Rightarrow> Node l a r |
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    31
     GT \<Rightarrow> Node l a (insert x r))"
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    32
71463
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    33
text \<open>Deletion by replacing:\<close>
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    34
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
    35
fun split_min :: "'a tree \<Rightarrow> 'a * 'a tree" where
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
    36
"split_min (Node l a r) =
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
    37
  (if l = Leaf then (a,r) else let (x,l') = split_min l in (x, Node l' a r))"
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    38
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 61678
diff changeset
    39
fun delete :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    40
"delete x Leaf = Leaf" |
61678
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    41
"delete x (Node l a r) =
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    42
  (case cmp x a of
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    43
     LT \<Rightarrow>  Node (delete x l) a r |
b594e9277be3 tuned white space
nipkow
parents: 61651
diff changeset
    44
     GT \<Rightarrow>  Node l a (delete x r) |
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
    45
     EQ \<Rightarrow> if r = Leaf then l else let (a',r') = split_min r in Node l a' r')"
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    46
71829
6f2663df8374 "app" -> "join" for uniformity with Join theory; tuned defs
nipkow
parents: 71463
diff changeset
    47
text \<open>Deletion by joining:\<close>
71463
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    48
71829
6f2663df8374 "app" -> "join" for uniformity with Join theory; tuned defs
nipkow
parents: 71463
diff changeset
    49
fun join :: "('a::linorder)tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
6f2663df8374 "app" -> "join" for uniformity with Join theory; tuned defs
nipkow
parents: 71463
diff changeset
    50
"join t Leaf = t" |
6f2663df8374 "app" -> "join" for uniformity with Join theory; tuned defs
nipkow
parents: 71463
diff changeset
    51
"join Leaf t = t" |
6f2663df8374 "app" -> "join" for uniformity with Join theory; tuned defs
nipkow
parents: 71463
diff changeset
    52
"join (Node t1 a t2) (Node t3 b t4) =
6f2663df8374 "app" -> "join" for uniformity with Join theory; tuned defs
nipkow
parents: 71463
diff changeset
    53
  (case join t2 t3 of
71463
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    54
     Leaf \<Rightarrow> Node t1 a (Node Leaf b t4) |
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    55
     Node u2 x u3 \<Rightarrow> Node (Node t1 a u2) x (Node u3 b t4))"
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    56
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    57
fun delete2 :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    58
"delete2 x Leaf = Leaf" |
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    59
"delete2 x (Node l a r) =
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    60
  (case cmp x a of
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    61
     LT \<Rightarrow>  Node (delete2 x l) a r |
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    62
     GT \<Rightarrow>  Node l a (delete2 x r) |
71829
6f2663df8374 "app" -> "join" for uniformity with Join theory; tuned defs
nipkow
parents: 71463
diff changeset
    63
     EQ \<Rightarrow> join l r)"
71463
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
    64
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    65
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    66
subsection "Functional Correctness Proofs"
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    67
67929
30486b96274d eliminated "elems"
nipkow
parents: 66453
diff changeset
    68
lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> set (inorder t))"
30486b96274d eliminated "elems"
nipkow
parents: 66453
diff changeset
    69
by (induction t) (auto simp: isin_simps)
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    70
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    71
lemma inorder_insert:
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    72
  "sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)"
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    73
by(induction t) (auto simp: ins_list_simps)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    74
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    75
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
    76
lemma split_minD:
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
    77
  "split_min t = (x,t') \<Longrightarrow> t \<noteq> Leaf \<Longrightarrow> x # inorder t' = inorder t"
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
    78
by(induction t arbitrary: t' rule: split_min.induct)
61647
nipkow
parents: 61640
diff changeset
    79
  (auto simp: sorted_lems split: prod.splits if_splits)
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    80
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    81
lemma inorder_delete:
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    82
  "sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
    83
by(induction t) (auto simp: del_list_simps split_minD split: prod.splits)
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    84
68440
6826718f732d qualify interpretations to avoid clashes
nipkow
parents: 68431
diff changeset
    85
interpretation S: Set_by_Ordered
68431
b294e095f64c more abstract naming
nipkow
parents: 68020
diff changeset
    86
where empty = empty and isin = isin and insert = insert and delete = delete
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    87
and inorder = inorder and inv = "\<lambda>_. True"
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    88
proof (standard, goal_cases)
68431
b294e095f64c more abstract naming
nipkow
parents: 68020
diff changeset
    89
  case 1 show ?case by (simp add: empty_def)
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    90
next
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    91
  case 2 thus ?case by(simp add: isin_set)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    92
next
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    93
  case 3 thus ?case by(simp add: inorder_insert)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    94
next
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    95
  case 4 thus ?case by(simp add: inorder_delete)
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    96
qed (rule TrueI)+
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
    97
71829
6f2663df8374 "app" -> "join" for uniformity with Join theory; tuned defs
nipkow
parents: 71463
diff changeset
    98
lemma inorder_join:
6f2663df8374 "app" -> "join" for uniformity with Join theory; tuned defs
nipkow
parents: 71463
diff changeset
    99
  "inorder(join l r) = inorder l @ inorder r"
6f2663df8374 "app" -> "join" for uniformity with Join theory; tuned defs
nipkow
parents: 71463
diff changeset
   100
by(induction l r rule: join.induct) (auto split: tree.split)
71463
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   101
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   102
lemma inorder_delete2:
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   103
  "sorted(inorder t) \<Longrightarrow> inorder(delete2 x t) = del_list x (inorder t)"
71829
6f2663df8374 "app" -> "join" for uniformity with Join theory; tuned defs
nipkow
parents: 71463
diff changeset
   104
by(induction t) (auto simp: inorder_join del_list_simps)
71463
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   105
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   106
interpretation S2: Set_by_Ordered
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   107
where empty = empty and isin = isin and insert = insert and delete = delete2
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   108
and inorder = inorder and inv = "\<lambda>_. True"
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   109
proof (standard, goal_cases)
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   110
  case 1 show ?case by (simp add: empty_def)
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   111
next
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   112
  case 2 thus ?case by(simp add: isin_set)
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   113
next
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   114
  case 3 thus ?case by(simp add: inorder_insert)
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   115
next
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   116
  case 4 thus ?case by(simp add: inorder_delete2)
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   117
qed (rule TrueI)+
a31a9da43694 tuned deletion
nipkow
parents: 68440
diff changeset
   118
61640
44c9198f210c no CRLF
nipkow
parents: 61588
diff changeset
   119
end