src/HOL/Library/RBT_Impl.thy
author wenzelm
Tue, 07 Mar 2023 10:57:50 +0100
changeset 77554 4465d9dff448
parent 77061 5de3772609ea
child 80914 d97fdabd9e2b
permissions -rw-r--r--
clarified terminology of "session build database", while "build database" is the one underlying Build_Process;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47455
26315a545e26 updated headers;
wenzelm
parents: 47450
diff changeset
     1
(*  Title:      HOL/Library/RBT_Impl.thy
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
     2
    Author:     Markus Reiter, TU Muenchen
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
     3
    Author:     Alexander Krauss, TU Muenchen
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
     4
*)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
     5
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
     6
section \<open>Implementation of Red-Black Trees\<close>
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
     7
36147
b43b22f63665 theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
haftmann
parents: 35618
diff changeset
     8
theory RBT_Impl
45990
b7b905b23b2a incorporated More_Set and More_List into the Main body -- to be consolidated later
haftmann
parents: 41959
diff changeset
     9
imports Main
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    10
begin
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    11
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
    12
text \<open>
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61433
diff changeset
    13
  For applications, you should use theory \<open>RBT\<close> which defines
36147
b43b22f63665 theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
haftmann
parents: 35618
diff changeset
    14
  an abstract type of red-black tree obeying the invariant.
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
    15
\<close>
36147
b43b22f63665 theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
haftmann
parents: 35618
diff changeset
    16
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
    17
subsection \<open>Datatype of RB trees\<close>
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    18
58310
91ea607a34d8 updated news
blanchet
parents: 58257
diff changeset
    19
datatype color = R | B
91ea607a34d8 updated news
blanchet
parents: 58257
diff changeset
    20
datatype ('a, 'b) rbt = Empty | Branch color "('a, 'b) rbt" 'a 'b "('a, 'b) rbt"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    21
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    22
lemma rbt_cases:
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    23
  obtains (Empty) "t = Empty" 
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    24
  | (Red) l k v r where "t = Branch R l k v r" 
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    25
  | (Black) l k v r where "t = Branch B l k v r"
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    26
proof (cases t)
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    27
  case Empty with that show thesis by blast
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    28
next
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    29
  case (Branch c) with that show thesis by (cases c) blast+
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    30
qed
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    31
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
    32
subsection \<open>Tree properties\<close>
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    33
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
    34
subsubsection \<open>Content of a tree\<close>
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    35
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    36
primrec entries :: "('a, 'b) rbt \<Rightarrow> ('a \<times> 'b) list"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    37
where 
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    38
  "entries Empty = []"
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    39
| "entries (Branch _ l k v r) = entries l @ (k,v) # entries r"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    40
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    41
abbreviation (input) entry_in_tree :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    42
where
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    43
  "entry_in_tree k v t \<equiv> (k, v) \<in> set (entries t)"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    44
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    45
definition keys :: "('a, 'b) rbt \<Rightarrow> 'a list" where
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    46
  "keys t = map fst (entries t)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    47
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    48
lemma keys_simps [simp, code]:
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    49
  "keys Empty = []"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    50
  "keys (Branch c l k v r) = keys l @ k # keys r"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    51
  by (simp_all add: keys_def)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    52
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    53
lemma entry_in_tree_keys:
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    54
  assumes "(k, v) \<in> set (entries t)"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    55
  shows "k \<in> set (keys t)"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    56
proof -
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    57
  from assms have "fst (k, v) \<in> fst ` set (entries t)" by (rule imageI)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    58
  then show ?thesis by (simp add: keys_def)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    59
qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    60
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
    61
lemma keys_entries:
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
    62
  "k \<in> set (keys t) \<longleftrightarrow> (\<exists>v. (k, v) \<in> set (entries t))"
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
    63
  by (auto intro: entry_in_tree_keys) (auto simp add: keys_def)
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
    64
48621
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
    65
lemma non_empty_rbt_keys: 
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
    66
  "t \<noteq> rbt.Empty \<Longrightarrow> keys t \<noteq> []"
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
    67
  by (cases t) simp_all
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    68
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
    69
subsubsection \<open>Search tree properties\<close>
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    70
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    71
context ord begin
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    72
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    73
definition rbt_less :: "'a \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    74
where
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    75
  rbt_less_prop: "rbt_less k t \<longleftrightarrow> (\<forall>x\<in>set (keys t). x < k)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    76
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    77
abbreviation rbt_less_symbol (infix "|\<guillemotleft>" 50)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    78
where "t |\<guillemotleft> x \<equiv> rbt_less x t"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    79
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    80
definition rbt_greater :: "'a \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool" (infix "\<guillemotleft>|" 50) 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    81
where
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    82
  rbt_greater_prop: "rbt_greater k t = (\<forall>x\<in>set (keys t). k < x)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    83
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    84
lemma rbt_less_simps [simp]:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    85
  "Empty |\<guillemotleft> k = True"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    86
  "Branch c lt kt v rt |\<guillemotleft> k \<longleftrightarrow> kt < k \<and> lt |\<guillemotleft> k \<and> rt |\<guillemotleft> k"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    87
  by (auto simp add: rbt_less_prop)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    88
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    89
lemma rbt_greater_simps [simp]:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    90
  "k \<guillemotleft>| Empty = True"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    91
  "k \<guillemotleft>| (Branch c lt kt v rt) \<longleftrightarrow> k < kt \<and> k \<guillemotleft>| lt \<and> k \<guillemotleft>| rt"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    92
  by (auto simp add: rbt_greater_prop)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    93
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    94
lemmas rbt_ord_props = rbt_less_prop rbt_greater_prop
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    95
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    96
lemmas rbt_greater_nit = rbt_greater_prop entry_in_tree_keys
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    97
lemmas rbt_less_nit = rbt_less_prop entry_in_tree_keys
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    98
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
    99
lemma (in order)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   100
  shows rbt_less_eq_trans: "l |\<guillemotleft> u \<Longrightarrow> u \<le> v \<Longrightarrow> l |\<guillemotleft> v"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   101
  and rbt_less_trans: "t |\<guillemotleft> x \<Longrightarrow> x < y \<Longrightarrow> t |\<guillemotleft> y"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   102
  and rbt_greater_eq_trans: "u \<le> v \<Longrightarrow> v \<guillemotleft>| r \<Longrightarrow> u \<guillemotleft>| r"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   103
  and rbt_greater_trans: "x < y \<Longrightarrow> y \<guillemotleft>| t \<Longrightarrow> x \<guillemotleft>| t"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   104
  by (auto simp: rbt_ord_props)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   105
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   106
primrec rbt_sorted :: "('a, 'b) rbt \<Rightarrow> bool"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   107
where
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   108
  "rbt_sorted Empty = True"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   109
| "rbt_sorted (Branch c l k v r) = (l |\<guillemotleft> k \<and> k \<guillemotleft>| r \<and> rbt_sorted l \<and> rbt_sorted r)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   110
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   111
end
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   112
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   113
context linorder begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   114
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   115
lemma rbt_sorted_entries:
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
   116
  "rbt_sorted t \<Longrightarrow> List.sorted (map fst (entries t))"
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
   117
by (induct t)  (force simp: sorted_append rbt_ord_props dest!: entry_in_tree_keys)+
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   118
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   119
lemma distinct_entries:
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
   120
  "rbt_sorted t \<Longrightarrow> distinct (map fst (entries t))"
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
   121
by (induct t) (force simp: sorted_append rbt_ord_props dest!: entry_in_tree_keys)+
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   122
48621
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
   123
lemma distinct_keys:
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
   124
  "rbt_sorted t \<Longrightarrow> distinct (keys t)"
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
   125
  by (simp add: distinct_entries keys_def)
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
   126
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
   127
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
   128
subsubsection \<open>Tree lookup\<close>
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   129
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   130
primrec (in ord) rbt_lookup :: "('a, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   131
where
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   132
  "rbt_lookup Empty k = None"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   133
| "rbt_lookup (Branch _ l x y r) k = 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   134
   (if k < x then rbt_lookup l k else if x < k then rbt_lookup r k else Some y)"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   135
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   136
lemma rbt_lookup_keys: "rbt_sorted t \<Longrightarrow> dom (rbt_lookup t) = set (keys t)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   137
  by (induct t) (auto simp: dom_def rbt_greater_prop rbt_less_prop)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   138
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   139
lemma dom_rbt_lookup_Branch: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   140
  "rbt_sorted (Branch c t1 k v t2) \<Longrightarrow> 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   141
    dom (rbt_lookup (Branch c t1 k v t2)) 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   142
    = Set.insert k (dom (rbt_lookup t1) \<union> dom (rbt_lookup t2))"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   143
proof -
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   144
  assume "rbt_sorted (Branch c t1 k v t2)"
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
   145
  then show ?thesis by (simp add: rbt_lookup_keys)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   146
qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   147
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   148
lemma finite_dom_rbt_lookup [simp, intro!]: "finite (dom (rbt_lookup t))"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   149
proof (induct t)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   150
  case Empty then show ?case by simp
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   151
next
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   152
  case (Branch color t1 a b t2)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   153
  let ?A = "Set.insert a (dom (rbt_lookup t1) \<union> dom (rbt_lookup t2))"
62390
842917225d56 more canonical names
nipkow
parents: 62093
diff changeset
   154
  have "dom (rbt_lookup (Branch color t1 a b t2)) \<subseteq> ?A" by (auto split: if_split_asm)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   155
  moreover from Branch have "finite (insert a (dom (rbt_lookup t1) \<union> dom (rbt_lookup t2)))" by simp
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   156
  ultimately show ?case by (rule finite_subset)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   157
qed 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   158
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   159
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   160
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   161
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   162
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   163
lemma rbt_lookup_rbt_less[simp]: "t |\<guillemotleft> k \<Longrightarrow> rbt_lookup t k = None" 
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   164
by (induct t) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   165
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   166
lemma rbt_lookup_rbt_greater[simp]: "k \<guillemotleft>| t \<Longrightarrow> rbt_lookup t k = None"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   167
by (induct t) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   168
68450
41de07c7a0f3 Map.empty now qualified to avoid name clashes
nipkow
parents: 68109
diff changeset
   169
lemma rbt_lookup_Empty: "rbt_lookup Empty = Map.empty"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   170
by (rule ext) simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   171
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   172
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   173
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   174
context linorder begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   175
35618
b7bfd4cbcfc0 some lemma refinements
haftmann
parents: 35606
diff changeset
   176
lemma map_of_entries:
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   177
  "rbt_sorted t \<Longrightarrow> map_of (entries t) = rbt_lookup t"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   178
proof (induct t)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   179
  case Empty thus ?case by (simp add: rbt_lookup_Empty)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   180
next
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   181
  case (Branch c t1 k v t2)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   182
  have "rbt_lookup (Branch c t1 k v t2) = rbt_lookup t2 ++ [k\<mapsto>v] ++ rbt_lookup t1"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   183
  proof (rule ext)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   184
    fix x
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   185
    from Branch have RBT_SORTED: "rbt_sorted (Branch c t1 k v t2)" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   186
    let ?thesis = "rbt_lookup (Branch c t1 k v t2) x = (rbt_lookup t2 ++ [k \<mapsto> v] ++ rbt_lookup t1) x"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   187
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   188
    have DOM_T1: "!!k'. k'\<in>dom (rbt_lookup t1) \<Longrightarrow> k>k'"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   189
    proof -
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   190
      fix k'
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   191
      from RBT_SORTED have "t1 |\<guillemotleft> k" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   192
      with rbt_less_prop have "\<forall>k'\<in>set (keys t1). k>k'" by auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   193
      moreover assume "k'\<in>dom (rbt_lookup t1)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   194
      ultimately show "k>k'" using rbt_lookup_keys RBT_SORTED by auto
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   195
    qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   196
    
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   197
    have DOM_T2: "!!k'. k'\<in>dom (rbt_lookup t2) \<Longrightarrow> k<k'"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   198
    proof -
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   199
      fix k'
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   200
      from RBT_SORTED have "k \<guillemotleft>| t2" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   201
      with rbt_greater_prop have "\<forall>k'\<in>set (keys t2). k<k'" by auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   202
      moreover assume "k'\<in>dom (rbt_lookup t2)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   203
      ultimately show "k<k'" using rbt_lookup_keys RBT_SORTED by auto
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   204
    qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   205
    
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   206
    {
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   207
      assume C: "x<k"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   208
      hence "rbt_lookup (Branch c t1 k v t2) x = rbt_lookup t1 x" by simp
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   209
      moreover from C have "x\<notin>dom [k\<mapsto>v]" by simp
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   210
      moreover have "x \<notin> dom (rbt_lookup t2)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   211
      proof
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   212
        assume "x \<in> dom (rbt_lookup t2)"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   213
        with DOM_T2 have "k<x" by blast
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   214
        with C show False by simp
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   215
      qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   216
      ultimately have ?thesis by (simp add: map_add_upd_left map_add_dom_app_simps)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   217
    } moreover {
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   218
      assume [simp]: "x=k"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   219
      hence "rbt_lookup (Branch c t1 k v t2) x = [k \<mapsto> v] x" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   220
      moreover have "x \<notin> dom (rbt_lookup t1)" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   221
      proof
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   222
        assume "x \<in> dom (rbt_lookup t1)"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   223
        with DOM_T1 have "k>x" by blast
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   224
        thus False by simp
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   225
      qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   226
      ultimately have ?thesis by (simp add: map_add_upd_left map_add_dom_app_simps)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   227
    } moreover {
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   228
      assume C: "x>k"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   229
      hence "rbt_lookup (Branch c t1 k v t2) x = rbt_lookup t2 x" by (simp add: less_not_sym[of k x])
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   230
      moreover from C have "x\<notin>dom [k\<mapsto>v]" by simp
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   231
      moreover have "x\<notin>dom (rbt_lookup t1)" proof
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   232
        assume "x\<in>dom (rbt_lookup t1)"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   233
        with DOM_T1 have "k>x" by simp
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   234
        with C show False by simp
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   235
      qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   236
      ultimately have ?thesis by (simp add: map_add_upd_left map_add_dom_app_simps)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   237
    } ultimately show ?thesis using less_linear by blast
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   238
  qed
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   239
  also from Branch 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   240
  have "rbt_lookup t2 ++ [k \<mapsto> v] ++ rbt_lookup t1 = map_of (entries (Branch c t1 k v t2))" by simp
35618
b7bfd4cbcfc0 some lemma refinements
haftmann
parents: 35606
diff changeset
   241
  finally show ?case by simp
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   242
qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   243
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   244
lemma rbt_lookup_in_tree: "rbt_sorted t \<Longrightarrow> rbt_lookup t k = Some v \<longleftrightarrow> (k, v) \<in> set (entries t)"
35618
b7bfd4cbcfc0 some lemma refinements
haftmann
parents: 35606
diff changeset
   245
  by (simp add: map_of_entries [symmetric] distinct_entries)
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   246
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   247
lemma set_entries_inject:
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   248
  assumes rbt_sorted: "rbt_sorted t1" "rbt_sorted t2" 
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   249
  shows "set (entries t1) = set (entries t2) \<longleftrightarrow> entries t1 = entries t2"
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   250
proof -
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   251
  from rbt_sorted have "distinct (map fst (entries t1))"
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   252
    "distinct (map fst (entries t2))"
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   253
    by (auto intro: distinct_entries)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   254
  with rbt_sorted show ?thesis
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   255
    by (auto intro: map_sorted_distinct_set_unique rbt_sorted_entries simp add: distinct_map)
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   256
qed
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   257
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   258
lemma entries_eqI:
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   259
  assumes rbt_sorted: "rbt_sorted t1" "rbt_sorted t2" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   260
  assumes rbt_lookup: "rbt_lookup t1 = rbt_lookup t2"
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   261
  shows "entries t1 = entries t2"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   262
proof -
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   263
  from rbt_sorted rbt_lookup have "map_of (entries t1) = map_of (entries t2)"
35618
b7bfd4cbcfc0 some lemma refinements
haftmann
parents: 35606
diff changeset
   264
    by (simp add: map_of_entries)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   265
  with rbt_sorted have "set (entries t1) = set (entries t2)"
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   266
    by (simp add: map_of_inject_set distinct_entries)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   267
  with rbt_sorted show ?thesis by (simp add: set_entries_inject)
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   268
qed
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   269
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   270
lemma entries_rbt_lookup:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   271
  assumes "rbt_sorted t1" "rbt_sorted t2" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   272
  shows "entries t1 = entries t2 \<longleftrightarrow> rbt_lookup t1 = rbt_lookup t2"
35618
b7bfd4cbcfc0 some lemma refinements
haftmann
parents: 35606
diff changeset
   273
  using assms by (auto intro: entries_eqI simp add: map_of_entries [symmetric])
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   274
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   275
lemma rbt_lookup_from_in_tree: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   276
  assumes "rbt_sorted t1" "rbt_sorted t2" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   277
  and "\<And>v. (k, v) \<in> set (entries t1) \<longleftrightarrow> (k, v) \<in> set (entries t2)" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   278
  shows "rbt_lookup t1 k = rbt_lookup t2 k"
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   279
proof -
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   280
  from assms have "k \<in> dom (rbt_lookup t1) \<longleftrightarrow> k \<in> dom (rbt_lookup t2)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   281
    by (simp add: keys_entries rbt_lookup_keys)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   282
  with assms show ?thesis by (auto simp add: rbt_lookup_in_tree [symmetric])
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   283
qed
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   284
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   285
end
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   286
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
   287
subsubsection \<open>Red-black properties\<close>
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   288
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   289
primrec color_of :: "('a, 'b) rbt \<Rightarrow> color"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   290
where
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   291
  "color_of Empty = B"
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   292
| "color_of (Branch c _ _ _ _) = c"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   293
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   294
primrec bheight :: "('a,'b) rbt \<Rightarrow> nat"
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   295
where
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   296
  "bheight Empty = 0"
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   297
| "bheight (Branch c lt k v rt) = (if c = B then Suc (bheight lt) else bheight lt)"
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   298
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   299
primrec inv1 :: "('a, 'b) rbt \<Rightarrow> bool"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   300
where
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   301
  "inv1 Empty = True"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   302
| "inv1 (Branch c lt k v rt) \<longleftrightarrow> inv1 lt \<and> inv1 rt \<and> (c = B \<or> color_of lt = B \<and> color_of rt = B)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   303
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61433
diff changeset
   304
primrec inv1l :: "('a, 'b) rbt \<Rightarrow> bool" \<comment> \<open>Weaker version\<close>
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   305
where
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   306
  "inv1l Empty = True"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   307
| "inv1l (Branch c l k v r) = (inv1 l \<and> inv1 r)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   308
lemma [simp]: "inv1 t \<Longrightarrow> inv1l t" by (cases t) simp+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   309
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   310
primrec inv2 :: "('a, 'b) rbt \<Rightarrow> bool"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   311
where
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   312
  "inv2 Empty = True"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   313
| "inv2 (Branch c lt k v rt) = (inv2 lt \<and> inv2 rt \<and> bheight lt = bheight rt)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   314
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   315
context ord begin
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   316
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   317
definition is_rbt :: "('a, 'b) rbt \<Rightarrow> bool" where
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   318
  "is_rbt t \<longleftrightarrow> inv1 t \<and> inv2 t \<and> color_of t = B \<and> rbt_sorted t"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   319
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   320
lemma is_rbt_rbt_sorted [simp]:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   321
  "is_rbt t \<Longrightarrow> rbt_sorted t" by (simp add: is_rbt_def)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   322
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   323
theorem Empty_is_rbt [simp]:
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   324
  "is_rbt Empty" by (simp add: is_rbt_def)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   325
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   326
end
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   327
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
   328
subsection \<open>Insertion\<close>
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   329
61225
1a690dce8cfc tuned references
nipkow
parents: 61121
diff changeset
   330
text \<open>The function definitions are based on the book by Okasaki.\<close>
1a690dce8cfc tuned references
nipkow
parents: 61121
diff changeset
   331
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   332
fun (* slow, due to massive case splitting *)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   333
  balance :: "('a,'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   334
where
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   335
  "balance (Branch R a w x b) s t (Branch R c y z d) = Branch R (Branch B a w x b) s t (Branch B c y z d)" |
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   336
  "balance (Branch R (Branch R a w x b) s t c) y z d = Branch R (Branch B a w x b) s t (Branch B c y z d)" |
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   337
  "balance (Branch R a w x (Branch R b s t c)) y z d = Branch R (Branch B a w x b) s t (Branch B c y z d)" |
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   338
  "balance a w x (Branch R b s t (Branch R c y z d)) = Branch R (Branch B a w x b) s t (Branch B c y z d)" |
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   339
  "balance a w x (Branch R (Branch R b s t c) y z d) = Branch R (Branch B a w x b) s t (Branch B c y z d)" |
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   340
  "balance a s t b = Branch B a s t b"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   341
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   342
lemma balance_inv1: "\<lbrakk>inv1l l; inv1l r\<rbrakk> \<Longrightarrow> inv1 (balance l k v r)" 
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   343
  by (induct l k v r rule: balance.induct) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   344
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   345
lemma balance_bheight: "bheight l = bheight r \<Longrightarrow> bheight (balance l k v r) = Suc (bheight l)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   346
  by (induct l k v r rule: balance.induct) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   347
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   348
lemma balance_inv2: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   349
  assumes "inv2 l" "inv2 r" "bheight l = bheight r"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   350
  shows "inv2 (balance l k v r)"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   351
  using assms
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   352
  by (induct l k v r rule: balance.induct) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   353
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   354
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   355
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   356
lemma balance_rbt_greater[simp]: "(v \<guillemotleft>| balance a k x b) = (v \<guillemotleft>| a \<and> v \<guillemotleft>| b \<and> v < k)" 
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   357
  by (induct a k x b rule: balance.induct) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   358
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   359
lemma balance_rbt_less[simp]: "(balance a k x b |\<guillemotleft> v) = (a |\<guillemotleft> v \<and> b |\<guillemotleft> v \<and> k < v)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   360
  by (induct a k x b rule: balance.induct) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   361
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   362
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   363
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   364
lemma (in linorder) balance_rbt_sorted: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   365
  fixes k :: "'a"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   366
  assumes "rbt_sorted l" "rbt_sorted r" "l |\<guillemotleft> k" "k \<guillemotleft>| r"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   367
  shows "rbt_sorted (balance l k v r)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   368
using assms proof (induct l k v r rule: balance.induct)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   369
  case ("2_2" a x w b y t c z s va vb vd vc)
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   370
  hence "y < z \<and> z \<guillemotleft>| Branch B va vb vd vc" 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   371
    by (auto simp add: rbt_ord_props)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   372
  hence "y \<guillemotleft>| (Branch B va vb vd vc)" by (blast dest: rbt_greater_trans)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   373
  with "2_2" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   374
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   375
  case ("3_2" va vb vd vc x w b y s c z)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   376
  from "3_2" have "x < y \<and> Branch B va vb vd vc |\<guillemotleft> x" 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   377
    by simp
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   378
  hence "Branch B va vb vd vc |\<guillemotleft> y" by (blast dest: rbt_less_trans)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   379
  with "3_2" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   380
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   381
  case ("3_3" x w b y s c z t va vb vd vc)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   382
  from "3_3" have "y < z \<and> z \<guillemotleft>| Branch B va vb vd vc" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   383
  hence "y \<guillemotleft>| Branch B va vb vd vc" by (blast dest: rbt_greater_trans)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   384
  with "3_3" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   385
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   386
  case ("3_4" vd ve vg vf x w b y s c z t va vb vii vc)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   387
  hence "x < y \<and> Branch B vd ve vg vf |\<guillemotleft> x" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   388
  hence 1: "Branch B vd ve vg vf |\<guillemotleft> y" by (blast dest: rbt_less_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   389
  from "3_4" have "y < z \<and> z \<guillemotleft>| Branch B va vb vii vc" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   390
  hence "y \<guillemotleft>| Branch B va vb vii vc" by (blast dest: rbt_greater_trans)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   391
  with 1 "3_4" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   392
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   393
  case ("4_2" va vb vd vc x w b y s c z t dd)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   394
  hence "x < y \<and> Branch B va vb vd vc |\<guillemotleft> x" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   395
  hence "Branch B va vb vd vc |\<guillemotleft> y" by (blast dest: rbt_less_trans)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   396
  with "4_2" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   397
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   398
  case ("5_2" x w b y s c z t va vb vd vc)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   399
  hence "y < z \<and> z \<guillemotleft>| Branch B va vb vd vc" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   400
  hence "y \<guillemotleft>| Branch B va vb vd vc" by (blast dest: rbt_greater_trans)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   401
  with "5_2" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   402
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   403
  case ("5_3" va vb vd vc x w b y s c z t)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   404
  hence "x < y \<and> Branch B va vb vd vc |\<guillemotleft> x" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   405
  hence "Branch B va vb vd vc |\<guillemotleft> y" by (blast dest: rbt_less_trans)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   406
  with "5_3" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   407
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   408
  case ("5_4" va vb vg vc x w b y s c z t vd ve vii vf)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   409
  hence "x < y \<and> Branch B va vb vg vc |\<guillemotleft> x" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   410
  hence 1: "Branch B va vb vg vc |\<guillemotleft> y" by (blast dest: rbt_less_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   411
  from "5_4" have "y < z \<and> z \<guillemotleft>| Branch B vd ve vii vf" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   412
  hence "y \<guillemotleft>| Branch B vd ve vii vf" by (blast dest: rbt_greater_trans)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   413
  with 1 "5_4" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   414
qed simp+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   415
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   416
lemma entries_balance [simp]:
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   417
  "entries (balance l k v r) = entries l @ (k, v) # entries r"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   418
  by (induct l k v r rule: balance.induct) auto
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   419
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   420
lemma keys_balance [simp]: 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   421
  "keys (balance l k v r) = keys l @ k # keys r"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   422
  by (simp add: keys_def)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   423
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   424
lemma balance_in_tree:  
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   425
  "entry_in_tree k x (balance l v y r) \<longleftrightarrow> entry_in_tree k x l \<or> k = v \<and> x = y \<or> entry_in_tree k x r"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   426
  by (auto simp add: keys_def)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   427
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   428
lemma (in linorder) rbt_lookup_balance[simp]: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   429
fixes k :: "'a"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   430
assumes "rbt_sorted l" "rbt_sorted r" "l |\<guillemotleft> k" "k \<guillemotleft>| r"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   431
shows "rbt_lookup (balance l k v r) x = rbt_lookup (Branch B l k v r) x"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   432
by (rule rbt_lookup_from_in_tree) (auto simp:assms balance_in_tree balance_rbt_sorted)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   433
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   434
primrec paint :: "color \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   435
where
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   436
  "paint c Empty = Empty"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   437
| "paint c (Branch _ l k v r) = Branch c l k v r"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   438
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   439
lemma paint_inv1l[simp]: "inv1l t \<Longrightarrow> inv1l (paint c t)" by (cases t) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   440
lemma paint_inv1[simp]: "inv1l t \<Longrightarrow> inv1 (paint B t)" by (cases t) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   441
lemma paint_inv2[simp]: "inv2 t \<Longrightarrow> inv2 (paint c t)" by (cases t) auto
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   442
lemma paint_color_of[simp]: "color_of (paint B t) = B" by (cases t) auto
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   443
lemma paint_in_tree[simp]: "entry_in_tree k x (paint c t) = entry_in_tree k x t" by (cases t) auto
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   444
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   445
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   446
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   447
lemma paint_rbt_sorted[simp]: "rbt_sorted t \<Longrightarrow> rbt_sorted (paint c t)" by (cases t) auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   448
lemma paint_rbt_lookup[simp]: "rbt_lookup (paint c t) = rbt_lookup t" by (rule ext) (cases t, auto)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   449
lemma paint_rbt_greater[simp]: "(v \<guillemotleft>| paint c t) = (v \<guillemotleft>| t)" by (cases t) auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   450
lemma paint_rbt_less[simp]: "(paint c t |\<guillemotleft> v) = (t |\<guillemotleft> v)" by (cases t) auto
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   451
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   452
fun
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   453
  rbt_ins :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   454
where
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   455
  "rbt_ins f k v Empty = Branch R Empty k v Empty" |
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   456
  "rbt_ins f k v (Branch B l x y r) = (if k < x then balance (rbt_ins f k v l) x y r
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   457
                                       else if k > x then balance l x y (rbt_ins f k v r)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   458
                                       else Branch B l x (f k y v) r)" |
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   459
  "rbt_ins f k v (Branch R l x y r) = (if k < x then Branch R (rbt_ins f k v l) x y r
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   460
                                       else if k > x then Branch R l x y (rbt_ins f k v r)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   461
                                       else Branch R l x (f k y v) r)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   462
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   463
lemma ins_inv1_inv2: 
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   464
  assumes "inv1 t" "inv2 t"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   465
  shows "inv2 (rbt_ins f k x t)" "bheight (rbt_ins f k x t) = bheight t" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   466
  "color_of t = B \<Longrightarrow> inv1 (rbt_ins f k x t)" "inv1l (rbt_ins f k x t)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   467
  using assms
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   468
  by (induct f k x t rule: rbt_ins.induct) (auto simp: balance_inv1 balance_inv2 balance_bheight)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   469
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   470
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   471
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   472
context linorder begin
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   473
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   474
lemma ins_rbt_greater[simp]: "(v \<guillemotleft>| rbt_ins f (k :: 'a) x t) = (v \<guillemotleft>| t \<and> k > v)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   475
  by (induct f k x t rule: rbt_ins.induct) auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   476
lemma ins_rbt_less[simp]: "(rbt_ins f k x t |\<guillemotleft> v) = (t |\<guillemotleft> v \<and> k < v)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   477
  by (induct f k x t rule: rbt_ins.induct) auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   478
lemma ins_rbt_sorted[simp]: "rbt_sorted t \<Longrightarrow> rbt_sorted (rbt_ins f k x t)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   479
  by (induct f k x t rule: rbt_ins.induct) (auto simp: balance_rbt_sorted)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   480
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   481
lemma keys_ins: "set (keys (rbt_ins f k v t)) = { k } \<union> set (keys t)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   482
  by (induct f k v t rule: rbt_ins.induct) auto
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   483
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   484
lemma rbt_lookup_ins: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   485
  fixes k :: "'a"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   486
  assumes "rbt_sorted t"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   487
  shows "rbt_lookup (rbt_ins f k v t) x = ((rbt_lookup t)(k |-> case rbt_lookup t k of None \<Rightarrow> v 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   488
                                                                | Some w \<Rightarrow> f k w v)) x"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   489
using assms by (induct f k v t rule: rbt_ins.induct) auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   490
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   491
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   492
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   493
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   494
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   495
definition rbt_insert_with_key :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   496
where "rbt_insert_with_key f k v t = paint B (rbt_ins f k v t)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   497
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   498
definition rbt_insertw_def: "rbt_insert_with f = rbt_insert_with_key (\<lambda>_. f)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   499
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   500
definition rbt_insert :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   501
  "rbt_insert = rbt_insert_with_key (\<lambda>_ _ nv. nv)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   502
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   503
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   504
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   505
context linorder begin
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   506
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   507
lemma rbt_insertwk_rbt_sorted: "rbt_sorted t \<Longrightarrow> rbt_sorted (rbt_insert_with_key f (k :: 'a) x t)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   508
  by (auto simp: rbt_insert_with_key_def)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   509
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   510
theorem rbt_insertwk_is_rbt: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   511
  assumes inv: "is_rbt t" 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   512
  shows "is_rbt (rbt_insert_with_key f k x t)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   513
using assms
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   514
unfolding rbt_insert_with_key_def is_rbt_def
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   515
by (auto simp: ins_inv1_inv2)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   516
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   517
lemma rbt_lookup_rbt_insertwk: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   518
  assumes "rbt_sorted t"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   519
  shows "rbt_lookup (rbt_insert_with_key f k v t) x = ((rbt_lookup t)(k |-> case rbt_lookup t k of None \<Rightarrow> v 
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   520
                                                       | Some w \<Rightarrow> f k w v)) x"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   521
unfolding rbt_insert_with_key_def using assms
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   522
by (simp add:rbt_lookup_ins)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   523
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   524
lemma rbt_insertw_rbt_sorted: "rbt_sorted t \<Longrightarrow> rbt_sorted (rbt_insert_with f k v t)" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   525
  by (simp add: rbt_insertwk_rbt_sorted rbt_insertw_def)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   526
theorem rbt_insertw_is_rbt: "is_rbt t \<Longrightarrow> is_rbt (rbt_insert_with f k v t)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   527
  by (simp add: rbt_insertwk_is_rbt rbt_insertw_def)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   528
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   529
lemma rbt_lookup_rbt_insertw:
63649
e690d6f2185b tuned proofs;
wenzelm
parents: 63040
diff changeset
   530
  "is_rbt t \<Longrightarrow>
e690d6f2185b tuned proofs;
wenzelm
parents: 63040
diff changeset
   531
    rbt_lookup (rbt_insert_with f k v t) =
e690d6f2185b tuned proofs;
wenzelm
parents: 63040
diff changeset
   532
      (rbt_lookup t)(k \<mapsto> (if k \<in> dom (rbt_lookup t) then f (the (rbt_lookup t k)) v else v))"
e690d6f2185b tuned proofs;
wenzelm
parents: 63040
diff changeset
   533
  by (rule ext, cases "rbt_lookup t k") (auto simp: rbt_lookup_rbt_insertwk dom_def rbt_insertw_def)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   534
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   535
lemma rbt_insert_rbt_sorted: "rbt_sorted t \<Longrightarrow> rbt_sorted (rbt_insert k v t)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   536
  by (simp add: rbt_insertwk_rbt_sorted rbt_insert_def)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   537
theorem rbt_insert_is_rbt [simp]: "is_rbt t \<Longrightarrow> is_rbt (rbt_insert k v t)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   538
  by (simp add: rbt_insertwk_is_rbt rbt_insert_def)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   539
63649
e690d6f2185b tuned proofs;
wenzelm
parents: 63040
diff changeset
   540
lemma rbt_lookup_rbt_insert: "is_rbt t \<Longrightarrow> rbt_lookup (rbt_insert k v t) = (rbt_lookup t)(k\<mapsto>v)"
e690d6f2185b tuned proofs;
wenzelm
parents: 63040
diff changeset
   541
  by (rule ext) (simp add: rbt_insert_def rbt_lookup_rbt_insertwk split: option.split)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   542
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   543
end
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   544
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
   545
subsection \<open>Deletion\<close>
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   546
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   547
lemma bheight_paintR'[simp]: "color_of t = B \<Longrightarrow> bheight (paint R t) = bheight t - 1"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   548
by (cases t rule: rbt_cases) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   549
63680
6e1e8b5abbfa more symbols;
wenzelm
parents: 63649
diff changeset
   550
text \<open>
6e1e8b5abbfa more symbols;
wenzelm
parents: 63649
diff changeset
   551
  The function definitions are based on the Haskell code by Stefan Kahrs
6e1e8b5abbfa more symbols;
wenzelm
parents: 63649
diff changeset
   552
  at \<^url>\<open>http://www.cs.ukc.ac.uk/people/staff/smk/redblack/rb.html\<close>.
6e1e8b5abbfa more symbols;
wenzelm
parents: 63649
diff changeset
   553
\<close>
61225
1a690dce8cfc tuned references
nipkow
parents: 61121
diff changeset
   554
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   555
fun
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   556
  balance_left :: "('a,'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   557
where
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   558
  "balance_left (Branch R a k x b) s y c = Branch R (Branch B a k x b) s y c" |
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   559
  "balance_left bl k x (Branch B a s y b) = balance bl k x (Branch R a s y b)" |
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   560
  "balance_left bl k x (Branch R (Branch B a s y b) t z c) = Branch R (Branch B bl k x a) s y (balance b t z (paint R c))" |
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   561
  "balance_left t k x s = Empty"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   562
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   563
lemma balance_left_inv2_with_inv1:
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   564
  assumes "inv2 lt" "inv2 rt" "bheight lt + 1 = bheight rt" "inv1 rt"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   565
  shows "bheight (balance_left lt k v rt) = bheight lt + 1"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   566
  and   "inv2 (balance_left lt k v rt)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   567
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   568
by (induct lt k v rt rule: balance_left.induct) (auto simp: balance_inv2 balance_bheight)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   569
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   570
lemma balance_left_inv2_app: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   571
  assumes "inv2 lt" "inv2 rt" "bheight lt + 1 = bheight rt" "color_of rt = B"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   572
  shows "inv2 (balance_left lt k v rt)" 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   573
        "bheight (balance_left lt k v rt) = bheight rt"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   574
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   575
by (induct lt k v rt rule: balance_left.induct) (auto simp add: balance_inv2 balance_bheight)+ 
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   576
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   577
lemma balance_left_inv1: "\<lbrakk>inv1l a; inv1 b; color_of b = B\<rbrakk> \<Longrightarrow> inv1 (balance_left a k x b)"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   578
  by (induct a k x b rule: balance_left.induct) (simp add: balance_inv1)+
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   579
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   580
lemma balance_left_inv1l: "\<lbrakk> inv1l lt; inv1 rt \<rbrakk> \<Longrightarrow> inv1l (balance_left lt k x rt)"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   581
by (induct lt k x rt rule: balance_left.induct) (auto simp: balance_inv1)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   582
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   583
lemma (in linorder) balance_left_rbt_sorted: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   584
  "\<lbrakk> rbt_sorted l; rbt_sorted r; rbt_less k l; k \<guillemotleft>| r \<rbrakk> \<Longrightarrow> rbt_sorted (balance_left l k v r)"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   585
apply (induct l k v r rule: balance_left.induct)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   586
apply (auto simp: balance_rbt_sorted)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   587
apply (unfold rbt_greater_prop rbt_less_prop)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   588
by force+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   589
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   590
context order begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   591
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   592
lemma balance_left_rbt_greater: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   593
  fixes k :: "'a"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   594
  assumes "k \<guillemotleft>| a" "k \<guillemotleft>| b" "k < x" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   595
  shows "k \<guillemotleft>| balance_left a x t b"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   596
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   597
by (induct a x t b rule: balance_left.induct) auto
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   598
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   599
lemma balance_left_rbt_less: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   600
  fixes k :: "'a"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   601
  assumes "a |\<guillemotleft> k" "b |\<guillemotleft> k" "x < k" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   602
  shows "balance_left a x t b |\<guillemotleft> k"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   603
using assms
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   604
by (induct a x t b rule: balance_left.induct) auto
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   605
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   606
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   607
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   608
lemma balance_left_in_tree: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   609
  assumes "inv1l l" "inv1 r" "bheight l + 1 = bheight r"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   610
  shows "entry_in_tree k v (balance_left l a b r) = (entry_in_tree k v l \<or> k = a \<and> v = b \<or> entry_in_tree k v r)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   611
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   612
by (induct l k v r rule: balance_left.induct) (auto simp: balance_in_tree)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   613
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   614
fun
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   615
  balance_right :: "('a,'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   616
where
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   617
  "balance_right a k x (Branch R b s y c) = Branch R a k x (Branch B b s y c)" |
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   618
  "balance_right (Branch B a k x b) s y bl = balance (Branch R a k x b) s y bl" |
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   619
  "balance_right (Branch R a k x (Branch B b s y c)) t z bl = Branch R (balance (paint R a) k x b) s y (Branch B c t z bl)" |
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   620
  "balance_right t k x s = Empty"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   621
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   622
lemma balance_right_inv2_with_inv1:
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   623
  assumes "inv2 lt" "inv2 rt" "bheight lt = bheight rt + 1" "inv1 lt"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   624
  shows "inv2 (balance_right lt k v rt) \<and> bheight (balance_right lt k v rt) = bheight lt"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   625
using assms
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   626
by (induct lt k v rt rule: balance_right.induct) (auto simp: balance_inv2 balance_bheight)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   627
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   628
lemma balance_right_inv1: "\<lbrakk>inv1 a; inv1l b; color_of a = B\<rbrakk> \<Longrightarrow> inv1 (balance_right a k x b)"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   629
by (induct a k x b rule: balance_right.induct) (simp add: balance_inv1)+
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   630
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   631
lemma balance_right_inv1l: "\<lbrakk> inv1 lt; inv1l rt \<rbrakk> \<Longrightarrow>inv1l (balance_right lt k x rt)"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   632
by (induct lt k x rt rule: balance_right.induct) (auto simp: balance_inv1)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   633
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   634
lemma (in linorder) balance_right_rbt_sorted:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   635
  "\<lbrakk> rbt_sorted l; rbt_sorted r; rbt_less k l; k \<guillemotleft>| r \<rbrakk> \<Longrightarrow> rbt_sorted (balance_right l k v r)"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   636
apply (induct l k v r rule: balance_right.induct)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   637
apply (auto simp:balance_rbt_sorted)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   638
apply (unfold rbt_less_prop rbt_greater_prop)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   639
by force+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   640
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   641
context order begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   642
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   643
lemma balance_right_rbt_greater: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   644
  fixes k :: "'a"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   645
  assumes "k \<guillemotleft>| a" "k \<guillemotleft>| b" "k < x" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   646
  shows "k \<guillemotleft>| balance_right a x t b"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   647
using assms by (induct a x t b rule: balance_right.induct) auto
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   648
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   649
lemma balance_right_rbt_less: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   650
  fixes k :: "'a"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   651
  assumes "a |\<guillemotleft> k" "b |\<guillemotleft> k" "x < k" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   652
  shows "balance_right a x t b |\<guillemotleft> k"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   653
using assms by (induct a x t b rule: balance_right.induct) auto
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   654
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   655
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   656
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   657
lemma balance_right_in_tree:
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   658
  assumes "inv1 l" "inv1l r" "bheight l = bheight r + 1" "inv2 l" "inv2 r"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   659
  shows "entry_in_tree x y (balance_right l k v r) = (entry_in_tree x y l \<or> x = k \<and> y = v \<or> entry_in_tree x y r)"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   660
using assms by (induct l k v r rule: balance_right.induct) (auto simp: balance_in_tree)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   661
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   662
fun
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   663
  combine :: "('a,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   664
where
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   665
  "combine Empty x = x" 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   666
| "combine x Empty = x" 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   667
| "combine (Branch R a k x b) (Branch R c s y d) = (case (combine b c) of
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   668
                                    Branch R b2 t z c2 \<Rightarrow> (Branch R (Branch R a k x b2) t z (Branch R c2 s y d)) |
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   669
                                    bc \<Rightarrow> Branch R a k x (Branch R bc s y d))" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   670
| "combine (Branch B a k x b) (Branch B c s y d) = (case (combine b c) of
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   671
                                    Branch R b2 t z c2 \<Rightarrow> Branch R (Branch B a k x b2) t z (Branch B c2 s y d) |
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   672
                                    bc \<Rightarrow> balance_left a k x (Branch B bc s y d))" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   673
| "combine a (Branch R b k x c) = Branch R (combine a b) k x c" 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   674
| "combine (Branch R a k x b) c = Branch R a k x (combine b c)" 
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   675
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   676
lemma combine_inv2:
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   677
  assumes "inv2 lt" "inv2 rt" "bheight lt = bheight rt"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   678
  shows "bheight (combine lt rt) = bheight lt" "inv2 (combine lt rt)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   679
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   680
by (induct lt rt rule: combine.induct) 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   681
   (auto simp: balance_left_inv2_app split: rbt.splits color.splits)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   682
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   683
lemma combine_inv1: 
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   684
  assumes "inv1 lt" "inv1 rt"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   685
  shows "color_of lt = B \<Longrightarrow> color_of rt = B \<Longrightarrow> inv1 (combine lt rt)"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   686
         "inv1l (combine lt rt)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   687
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   688
by (induct lt rt rule: combine.induct)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   689
   (auto simp: balance_left_inv1 split: rbt.splits color.splits)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   690
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   691
context linorder begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   692
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   693
lemma combine_rbt_greater[simp]: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   694
  fixes k :: "'a"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   695
  assumes "k \<guillemotleft>| l" "k \<guillemotleft>| r" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   696
  shows "k \<guillemotleft>| combine l r"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   697
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   698
by (induct l r rule: combine.induct)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   699
   (auto simp: balance_left_rbt_greater split:rbt.splits color.splits)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   700
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   701
lemma combine_rbt_less[simp]: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   702
  fixes k :: "'a"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   703
  assumes "l |\<guillemotleft> k" "r |\<guillemotleft> k" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   704
  shows "combine l r |\<guillemotleft> k"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   705
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   706
by (induct l r rule: combine.induct)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   707
   (auto simp: balance_left_rbt_less split:rbt.splits color.splits)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   708
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   709
lemma combine_rbt_sorted: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   710
  fixes k :: "'a"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   711
  assumes "rbt_sorted l" "rbt_sorted r" "l |\<guillemotleft> k" "k \<guillemotleft>| r"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   712
  shows "rbt_sorted (combine l r)"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   713
using assms proof (induct l r rule: combine.induct)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   714
  case (3 a x v b c y w d)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   715
  hence ineqs: "a |\<guillemotleft> x" "x \<guillemotleft>| b" "b |\<guillemotleft> k" "k \<guillemotleft>| c" "c |\<guillemotleft> y" "y \<guillemotleft>| d"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   716
    by auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   717
  with 3
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   718
  show ?case
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   719
    by (cases "combine b c" rule: rbt_cases)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   720
      (auto, (metis combine_rbt_greater combine_rbt_less ineqs ineqs rbt_less_simps(2) rbt_greater_simps(2) rbt_greater_trans rbt_less_trans)+)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   721
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   722
  case (4 a x v b c y w d)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   723
  hence "x < k \<and> rbt_greater k c" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   724
  hence "rbt_greater x c" by (blast dest: rbt_greater_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   725
  with 4 have 2: "rbt_greater x (combine b c)" by (simp add: combine_rbt_greater)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   726
  from 4 have "k < y \<and> rbt_less k b" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   727
  hence "rbt_less y b" by (blast dest: rbt_less_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   728
  with 4 have 3: "rbt_less y (combine b c)" by (simp add: combine_rbt_less)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   729
  show ?case
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   730
  proof (cases "combine b c" rule: rbt_cases)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   731
    case Empty
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   732
    from 4 have "x < y \<and> rbt_greater y d" by auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   733
    hence "rbt_greater x d" by (blast dest: rbt_greater_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   734
    with 4 Empty have "rbt_sorted a" and "rbt_sorted (Branch B Empty y w d)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   735
      and "rbt_less x a" and "rbt_greater x (Branch B Empty y w d)" by auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   736
    with Empty show ?thesis by (simp add: balance_left_rbt_sorted)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   737
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   738
    case (Red lta va ka rta)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   739
    with 2 4 have "x < va \<and> rbt_less x a" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   740
    hence 5: "rbt_less va a" by (blast dest: rbt_less_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   741
    from Red 3 4 have "va < y \<and> rbt_greater y d" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   742
    hence "rbt_greater va d" by (blast dest: rbt_greater_trans)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   743
    with Red 2 3 4 5 show ?thesis by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   744
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   745
    case (Black lta va ka rta)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   746
    from 4 have "x < y \<and> rbt_greater y d" by auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   747
    hence "rbt_greater x d" by (blast dest: rbt_greater_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   748
    with Black 2 3 4 have "rbt_sorted a" and "rbt_sorted (Branch B (combine b c) y w d)" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   749
      and "rbt_less x a" and "rbt_greater x (Branch B (combine b c) y w d)" by auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   750
    with Black show ?thesis by (simp add: balance_left_rbt_sorted)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   751
  qed
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   752
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   753
  case (5 va vb vd vc b x w c)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   754
  hence "k < x \<and> rbt_less k (Branch B va vb vd vc)" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   755
  hence "rbt_less x (Branch B va vb vd vc)" by (blast dest: rbt_less_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   756
  with 5 show ?case by (simp add: combine_rbt_less)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   757
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   758
  case (6 a x v b va vb vd vc)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   759
  hence "x < k \<and> rbt_greater k (Branch B va vb vd vc)" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   760
  hence "rbt_greater x (Branch B va vb vd vc)" by (blast dest: rbt_greater_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   761
  with 6 show ?case by (simp add: combine_rbt_greater)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   762
qed simp+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   763
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   764
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   765
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   766
lemma combine_in_tree: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   767
  assumes "inv2 l" "inv2 r" "bheight l = bheight r" "inv1 l" "inv1 r"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   768
  shows "entry_in_tree k v (combine l r) = (entry_in_tree k v l \<or> entry_in_tree k v r)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   769
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   770
proof (induct l r rule: combine.induct)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   771
  case (4 _ _ _ b c)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   772
  hence a: "bheight (combine b c) = bheight b" by (simp add: combine_inv2)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   773
  from 4 have b: "inv1l (combine b c)" by (simp add: combine_inv1)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   774
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   775
  show ?case
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   776
  proof (cases "combine b c" rule: rbt_cases)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   777
    case Empty
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   778
    with 4 a show ?thesis by (auto simp: balance_left_in_tree)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   779
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   780
    case (Red lta ka va rta)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   781
    with 4 show ?thesis by auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   782
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   783
    case (Black lta ka va rta)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   784
    with a b 4  show ?thesis by (auto simp: balance_left_in_tree)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   785
  qed 
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   786
qed (auto split: rbt.splits color.splits)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   787
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   788
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   789
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   790
fun
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   791
  rbt_del_from_left :: "'a \<Rightarrow> ('a,'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt" and
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   792
  rbt_del_from_right :: "'a \<Rightarrow> ('a,'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt" and
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   793
  rbt_del :: "'a\<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   794
where
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   795
  "rbt_del x Empty = Empty" |
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   796
  "rbt_del x (Branch c a y s b) = 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   797
   (if x < y then rbt_del_from_left x a y s b 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   798
    else (if x > y then rbt_del_from_right x a y s b else combine a b))" |
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   799
  "rbt_del_from_left x (Branch B lt z v rt) y s b = balance_left (rbt_del x (Branch B lt z v rt)) y s b" |
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   800
  "rbt_del_from_left x a y s b = Branch R (rbt_del x a) y s b" |
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   801
  "rbt_del_from_right x a y s (Branch B lt z v rt) = balance_right a y s (rbt_del x (Branch B lt z v rt))" | 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   802
  "rbt_del_from_right x a y s b = Branch R a y s (rbt_del x b)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   803
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   804
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   805
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   806
context linorder begin
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   807
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   808
lemma 
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   809
  assumes "inv2 lt" "inv1 lt"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   810
  shows
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   811
  "\<lbrakk>inv2 rt; bheight lt = bheight rt; inv1 rt\<rbrakk> \<Longrightarrow>
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   812
   inv2 (rbt_del_from_left x lt k v rt) \<and> 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   813
   bheight (rbt_del_from_left x lt k v rt) = bheight lt \<and> 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   814
   (color_of lt = B \<and> color_of rt = B \<and> inv1 (rbt_del_from_left x lt k v rt) \<or> 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   815
    (color_of lt \<noteq> B \<or> color_of rt \<noteq> B) \<and> inv1l (rbt_del_from_left x lt k v rt))"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   816
  and "\<lbrakk>inv2 rt; bheight lt = bheight rt; inv1 rt\<rbrakk> \<Longrightarrow>
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   817
  inv2 (rbt_del_from_right x lt k v rt) \<and> 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   818
  bheight (rbt_del_from_right x lt k v rt) = bheight lt \<and> 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   819
  (color_of lt = B \<and> color_of rt = B \<and> inv1 (rbt_del_from_right x lt k v rt) \<or> 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   820
   (color_of lt \<noteq> B \<or> color_of rt \<noteq> B) \<and> inv1l (rbt_del_from_right x lt k v rt))"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   821
  and rbt_del_inv1_inv2: "inv2 (rbt_del x lt) \<and> (color_of lt = R \<and> bheight (rbt_del x lt) = bheight lt \<and> inv1 (rbt_del x lt) 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   822
  \<or> color_of lt = B \<and> bheight (rbt_del x lt) = bheight lt - 1 \<and> inv1l (rbt_del x lt))"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   823
using assms
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   824
proof (induct x lt k v rt and x lt k v rt and x lt rule: rbt_del_from_left_rbt_del_from_right_rbt_del.induct)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   825
case (2 y c _ y')
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   826
  have "y = y' \<or> y < y' \<or> y > y'" by auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   827
  thus ?case proof (elim disjE)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   828
    assume "y = y'"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   829
    with 2 show ?thesis by (cases c) (simp add: combine_inv2 combine_inv1)+
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   830
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   831
    assume "y < y'"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   832
    with 2 show ?thesis by (cases c) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   833
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   834
    assume "y' < y"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   835
    with 2 show ?thesis by (cases c) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   836
  qed
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   837
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   838
  case (3 y lt z v rta y' ss bb) 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   839
  thus ?case by (cases "color_of (Branch B lt z v rta) = B \<and> color_of bb = B") (simp add: balance_left_inv2_with_inv1 balance_left_inv1 balance_left_inv1l)+
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   840
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   841
  case (5 y a y' ss lt z v rta)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   842
  thus ?case by (cases "color_of a = B \<and> color_of (Branch B lt z v rta) = B") (simp add: balance_right_inv2_with_inv1 balance_right_inv1 balance_right_inv1l)+
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   843
next
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   844
  case ("6_1" y a y' ss) thus ?case by (cases "color_of a = B \<and> color_of Empty = B") simp+
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   845
qed auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   846
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   847
lemma 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   848
  rbt_del_from_left_rbt_less: "\<lbrakk> lt |\<guillemotleft> v; rt |\<guillemotleft> v; k < v\<rbrakk> \<Longrightarrow> rbt_del_from_left x lt k y rt |\<guillemotleft> v"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   849
  and rbt_del_from_right_rbt_less: "\<lbrakk>lt |\<guillemotleft> v; rt |\<guillemotleft> v; k < v\<rbrakk> \<Longrightarrow> rbt_del_from_right x lt k y rt |\<guillemotleft> v"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   850
  and rbt_del_rbt_less: "lt |\<guillemotleft> v \<Longrightarrow> rbt_del x lt |\<guillemotleft> v"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   851
by (induct x lt k y rt and x lt k y rt and x lt rule: rbt_del_from_left_rbt_del_from_right_rbt_del.induct) 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   852
   (auto simp: balance_left_rbt_less balance_right_rbt_less)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   853
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   854
lemma rbt_del_from_left_rbt_greater: "\<lbrakk>v \<guillemotleft>| lt; v \<guillemotleft>| rt; k > v\<rbrakk> \<Longrightarrow> v \<guillemotleft>| rbt_del_from_left x lt k y rt"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   855
  and rbt_del_from_right_rbt_greater: "\<lbrakk>v \<guillemotleft>| lt; v \<guillemotleft>| rt; k > v\<rbrakk> \<Longrightarrow> v \<guillemotleft>| rbt_del_from_right x lt k y rt"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   856
  and rbt_del_rbt_greater: "v \<guillemotleft>| lt \<Longrightarrow> v \<guillemotleft>| rbt_del x lt"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   857
by (induct x lt k y rt and x lt k y rt and x lt rule: rbt_del_from_left_rbt_del_from_right_rbt_del.induct)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   858
   (auto simp: balance_left_rbt_greater balance_right_rbt_greater)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   859
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   860
lemma "\<lbrakk>rbt_sorted lt; rbt_sorted rt; lt |\<guillemotleft> k; k \<guillemotleft>| rt\<rbrakk> \<Longrightarrow> rbt_sorted (rbt_del_from_left x lt k y rt)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   861
  and "\<lbrakk>rbt_sorted lt; rbt_sorted rt; lt |\<guillemotleft> k; k \<guillemotleft>| rt\<rbrakk> \<Longrightarrow> rbt_sorted (rbt_del_from_right x lt k y rt)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   862
  and rbt_del_rbt_sorted: "rbt_sorted lt \<Longrightarrow> rbt_sorted (rbt_del x lt)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   863
proof (induct x lt k y rt and x lt k y rt and x lt rule: rbt_del_from_left_rbt_del_from_right_rbt_del.induct)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   864
  case (3 x lta zz v rta yy ss bb)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   865
  from 3 have "Branch B lta zz v rta |\<guillemotleft> yy" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   866
  hence "rbt_del x (Branch B lta zz v rta) |\<guillemotleft> yy" by (rule rbt_del_rbt_less)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   867
  with 3 show ?case by (simp add: balance_left_rbt_sorted)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   868
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   869
  case ("4_2" x vaa vbb vdd vc yy ss bb)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   870
  hence "Branch R vaa vbb vdd vc |\<guillemotleft> yy" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   871
  hence "rbt_del x (Branch R vaa vbb vdd vc) |\<guillemotleft> yy" by (rule rbt_del_rbt_less)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   872
  with "4_2" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   873
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   874
  case (5 x aa yy ss lta zz v rta) 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   875
  hence "yy \<guillemotleft>| Branch B lta zz v rta" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   876
  hence "yy \<guillemotleft>| rbt_del x (Branch B lta zz v rta)" by (rule rbt_del_rbt_greater)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   877
  with 5 show ?case by (simp add: balance_right_rbt_sorted)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   878
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   879
  case ("6_2" x aa yy ss vaa vbb vdd vc)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   880
  hence "yy \<guillemotleft>| Branch R vaa vbb vdd vc" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   881
  hence "yy \<guillemotleft>| rbt_del x (Branch R vaa vbb vdd vc)" by (rule rbt_del_rbt_greater)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   882
  with "6_2" show ?case by simp
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   883
qed (auto simp: combine_rbt_sorted)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   884
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   885
lemma "\<lbrakk>rbt_sorted lt; rbt_sorted rt; lt |\<guillemotleft> kt; kt \<guillemotleft>| rt; inv1 lt; inv1 rt; inv2 lt; inv2 rt; bheight lt = bheight rt; x < kt\<rbrakk> \<Longrightarrow> entry_in_tree k v (rbt_del_from_left x lt kt y rt) = (False \<or> (x \<noteq> k \<and> entry_in_tree k v (Branch c lt kt y rt)))"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   886
  and "\<lbrakk>rbt_sorted lt; rbt_sorted rt; lt |\<guillemotleft> kt; kt \<guillemotleft>| rt; inv1 lt; inv1 rt; inv2 lt; inv2 rt; bheight lt = bheight rt; x > kt\<rbrakk> \<Longrightarrow> entry_in_tree k v (rbt_del_from_right x lt kt y rt) = (False \<or> (x \<noteq> k \<and> entry_in_tree k v (Branch c lt kt y rt)))"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   887
  and rbt_del_in_tree: "\<lbrakk>rbt_sorted t; inv1 t; inv2 t\<rbrakk> \<Longrightarrow> entry_in_tree k v (rbt_del x t) = (False \<or> (x \<noteq> k \<and> entry_in_tree k v t))"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   888
proof (induct x lt kt y rt and x lt kt y rt and x t rule: rbt_del_from_left_rbt_del_from_right_rbt_del.induct)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   889
  case (2 xx c aa yy ss bb)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   890
  have "xx = yy \<or> xx < yy \<or> xx > yy" by auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   891
  from this 2 show ?case proof (elim disjE)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   892
    assume "xx = yy"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   893
    with 2 show ?thesis proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   894
      case True
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
   895
      from 2 \<open>xx = yy\<close> \<open>xx = k\<close> have "rbt_sorted (Branch c aa yy ss bb) \<and> k = yy" by simp
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   896
      hence "\<not> entry_in_tree k v aa" "\<not> entry_in_tree k v bb" by (auto simp: rbt_less_nit rbt_greater_prop)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
   897
      with \<open>xx = yy\<close> 2 \<open>xx = k\<close> show ?thesis by (simp add: combine_in_tree)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   898
    qed (simp add: combine_in_tree)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   899
  qed simp+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   900
next    
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   901
  case (3 xx lta zz vv rta yy ss bb)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
   902
  define mt where [simp]: "mt = Branch B lta zz vv rta"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   903
  from 3 have "inv2 mt \<and> inv1 mt" by simp
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   904
  hence "inv2 (rbt_del xx mt) \<and> (color_of mt = R \<and> bheight (rbt_del xx mt) = bheight mt \<and> inv1 (rbt_del xx mt) \<or> color_of mt = B \<and> bheight (rbt_del xx mt) = bheight mt - 1 \<and> inv1l (rbt_del xx mt))" by (blast dest: rbt_del_inv1_inv2)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   905
  with 3 have 4: "entry_in_tree k v (rbt_del_from_left xx mt yy ss bb) = (False \<or> xx \<noteq> k \<and> entry_in_tree k v mt \<or> (k = yy \<and> v = ss) \<or> entry_in_tree k v bb)" by (simp add: balance_left_in_tree)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   906
  thus ?case proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   907
    case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   908
    from 3 True have "yy \<guillemotleft>| bb \<and> yy > k" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   909
    hence "k \<guillemotleft>| bb" by (blast dest: rbt_greater_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   910
    with 3 4 True show ?thesis by (auto simp: rbt_greater_nit)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   911
  qed auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   912
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   913
  case ("4_1" xx yy ss bb)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   914
  show ?case proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   915
    case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   916
    with "4_1" have "yy \<guillemotleft>| bb \<and> k < yy" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   917
    hence "k \<guillemotleft>| bb" by (blast dest: rbt_greater_trans)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
   918
    with "4_1" \<open>xx = k\<close> 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   919
   have "entry_in_tree k v (Branch R Empty yy ss bb) = entry_in_tree k v Empty" by (auto simp: rbt_greater_nit)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   920
    thus ?thesis by auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   921
  qed simp+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   922
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   923
  case ("4_2" xx vaa vbb vdd vc yy ss bb)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   924
  thus ?case proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   925
    case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   926
    with "4_2" have "k < yy \<and> yy \<guillemotleft>| bb" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   927
    hence "k \<guillemotleft>| bb" by (blast dest: rbt_greater_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   928
    with True "4_2" show ?thesis by (auto simp: rbt_greater_nit)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   929
  qed auto
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   930
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   931
  case (5 xx aa yy ss lta zz vv rta)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
   932
  define mt where [simp]: "mt = Branch B lta zz vv rta"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   933
  from 5 have "inv2 mt \<and> inv1 mt" by simp
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   934
  hence "inv2 (rbt_del xx mt) \<and> (color_of mt = R \<and> bheight (rbt_del xx mt) = bheight mt \<and> inv1 (rbt_del xx mt) \<or> color_of mt = B \<and> bheight (rbt_del xx mt) = bheight mt - 1 \<and> inv1l (rbt_del xx mt))" by (blast dest: rbt_del_inv1_inv2)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   935
  with 5 have 3: "entry_in_tree k v (rbt_del_from_right xx aa yy ss mt) = (entry_in_tree k v aa \<or> (k = yy \<and> v = ss) \<or> False \<or> xx \<noteq> k \<and> entry_in_tree k v mt)" by (simp add: balance_right_in_tree)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   936
  thus ?case proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   937
    case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   938
    from 5 True have "aa |\<guillemotleft> yy \<and> yy < k" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   939
    hence "aa |\<guillemotleft> k" by (blast dest: rbt_less_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   940
    with 3 5 True show ?thesis by (auto simp: rbt_less_nit)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   941
  qed auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   942
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   943
  case ("6_1" xx aa yy ss)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   944
  show ?case proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   945
    case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   946
    with "6_1" have "aa |\<guillemotleft> yy \<and> k > yy" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   947
    hence "aa |\<guillemotleft> k" by (blast dest: rbt_less_trans)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
   948
    with "6_1" \<open>xx = k\<close> show ?thesis by (auto simp: rbt_less_nit)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   949
  qed simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   950
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   951
  case ("6_2" xx aa yy ss vaa vbb vdd vc)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   952
  thus ?case proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   953
    case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   954
    with "6_2" have "k > yy \<and> aa |\<guillemotleft> yy" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   955
    hence "aa |\<guillemotleft> k" by (blast dest: rbt_less_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   956
    with True "6_2" show ?thesis by (auto simp: rbt_less_nit)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   957
  qed auto
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   958
qed simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   959
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   960
definition (in ord) rbt_delete where
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   961
  "rbt_delete k t = paint B (rbt_del k t)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   962
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   963
theorem rbt_delete_is_rbt [simp]: assumes "is_rbt t" shows "is_rbt (rbt_delete k t)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   964
proof -
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   965
  from assms have "inv2 t" and "inv1 t" unfolding is_rbt_def by auto 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   966
  hence "inv2 (rbt_del k t) \<and> (color_of t = R \<and> bheight (rbt_del k t) = bheight t \<and> inv1 (rbt_del k t) \<or> color_of t = B \<and> bheight (rbt_del k t) = bheight t - 1 \<and> inv1l (rbt_del k t))" by (rule rbt_del_inv1_inv2)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   967
  hence "inv2 (rbt_del k t) \<and> inv1l (rbt_del k t)" by (cases "color_of t") auto
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   968
  with assms show ?thesis
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   969
    unfolding is_rbt_def rbt_delete_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   970
    by (auto intro: paint_rbt_sorted rbt_del_rbt_sorted)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   971
qed
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   972
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   973
lemma rbt_delete_in_tree: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   974
  assumes "is_rbt t" 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   975
  shows "entry_in_tree k v (rbt_delete x t) = (x \<noteq> k \<and> entry_in_tree k v t)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   976
  using assms unfolding is_rbt_def rbt_delete_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   977
  by (auto simp: rbt_del_in_tree)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   978
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   979
lemma rbt_lookup_rbt_delete:
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   980
  assumes is_rbt: "is_rbt t"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   981
  shows "rbt_lookup (rbt_delete k t) = (rbt_lookup t)|`(-{k})"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   982
proof
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   983
  fix x
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   984
  show "rbt_lookup (rbt_delete k t) x = (rbt_lookup t |` (-{k})) x" 
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   985
  proof (cases "x = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   986
    assume "x = k" 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   987
    with is_rbt show ?thesis
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   988
      by (cases "rbt_lookup (rbt_delete k t) k") (auto simp: rbt_lookup_in_tree rbt_delete_in_tree)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   989
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   990
    assume "x \<noteq> k"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   991
    thus ?thesis
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   992
      by auto (metis is_rbt rbt_delete_is_rbt rbt_delete_in_tree is_rbt_rbt_sorted rbt_lookup_from_in_tree)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   993
  qed
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   994
qed
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   995
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   996
end
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   997
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
   998
subsection \<open>Modifying existing entries\<close>
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   999
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1000
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1001
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1002
primrec
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1003
  rbt_map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1004
where
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1005
  "rbt_map_entry k f Empty = Empty"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1006
| "rbt_map_entry k f (Branch c lt x v rt) =
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1007
    (if k < x then Branch c (rbt_map_entry k f lt) x v rt
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1008
    else if k > x then (Branch c lt x v (rbt_map_entry k f rt))
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
  1009
    else Branch c lt x (f v) rt)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1010
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1011
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1012
lemma rbt_map_entry_color_of: "color_of (rbt_map_entry k f t) = color_of t" by (induct t) simp+
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1013
lemma rbt_map_entry_inv1: "inv1 (rbt_map_entry k f t) = inv1 t" by (induct t) (simp add: rbt_map_entry_color_of)+
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1014
lemma rbt_map_entry_inv2: "inv2 (rbt_map_entry k f t) = inv2 t" "bheight (rbt_map_entry k f t) = bheight t" by (induct t) simp+
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1015
lemma rbt_map_entry_rbt_greater: "rbt_greater a (rbt_map_entry k f t) = rbt_greater a t" by (induct t) simp+
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1016
lemma rbt_map_entry_rbt_less: "rbt_less a (rbt_map_entry k f t) = rbt_less a t" by (induct t) simp+
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1017
lemma rbt_map_entry_rbt_sorted: "rbt_sorted (rbt_map_entry k f t) = rbt_sorted t"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1018
  by (induct t) (simp_all add: rbt_map_entry_rbt_less rbt_map_entry_rbt_greater)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1019
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1020
theorem rbt_map_entry_is_rbt [simp]: "is_rbt (rbt_map_entry k f t) = is_rbt t" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1021
unfolding is_rbt_def by (simp add: rbt_map_entry_inv2 rbt_map_entry_color_of rbt_map_entry_rbt_sorted rbt_map_entry_inv1 )
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1022
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1023
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1024
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1025
theorem (in linorder) rbt_lookup_rbt_map_entry:
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55417
diff changeset
  1026
  "rbt_lookup (rbt_map_entry k f t) = (rbt_lookup t)(k := map_option f (rbt_lookup t k))"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1027
  by (induct t) (auto split: option.splits simp add: fun_eq_iff)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1028
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1029
subsection \<open>Mapping all entries\<close>
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1030
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1031
primrec
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
  1032
  map :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'c) rbt"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1033
where
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1034
  "map f Empty = Empty"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1035
| "map f (Branch c lt k v rt) = Branch c (map f lt) k (f k v) (map f rt)"
32237
cdc76a42fed4 added missing proof of RBT.map_of_alist_of (contributed by Peter Lammich)
krauss
parents: 30738
diff changeset
  1036
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1037
lemma map_entries [simp]: "entries (map f t) = List.map (\<lambda>(k, v). (k, f k v)) (entries t)"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1038
  by (induct t) auto
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1039
lemma map_keys [simp]: "keys (map f t) = keys t" by (simp add: keys_def split_def)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1040
lemma map_color_of: "color_of (map f t) = color_of t" by (induct t) simp+
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1041
lemma map_inv1: "inv1 (map f t) = inv1 t" by (induct t) (simp add: map_color_of)+
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1042
lemma map_inv2: "inv2 (map f t) = inv2 t" "bheight (map f t) = bheight t" by (induct t) simp+
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1043
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1044
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1045
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1046
lemma map_rbt_greater: "rbt_greater k (map f t) = rbt_greater k t" by (induct t) simp+
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1047
lemma map_rbt_less: "rbt_less k (map f t) = rbt_less k t" by (induct t) simp+
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1048
lemma map_rbt_sorted: "rbt_sorted (map f t) = rbt_sorted t"  by (induct t) (simp add: map_rbt_less map_rbt_greater)+
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1049
theorem map_is_rbt [simp]: "is_rbt (map f t) = is_rbt t" 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1050
unfolding is_rbt_def by (simp add: map_inv1 map_inv2 map_rbt_sorted map_color_of)
32237
cdc76a42fed4 added missing proof of RBT.map_of_alist_of (contributed by Peter Lammich)
krauss
parents: 30738
diff changeset
  1051
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1052
end
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1053
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55417
diff changeset
  1054
theorem (in linorder) rbt_lookup_map: "rbt_lookup (map f t) x = map_option (f x) (rbt_lookup t x)"
73526
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 73212
diff changeset
  1055
  by (induct t) (auto simp: antisym_conv3)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1056
 (* FIXME: simproc "antisym less" does not work for linorder context, only for linorder type class
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1057
    by (induct t) auto *)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1058
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1059
hide_const (open) map
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1060
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1061
subsection \<open>Folding over entries\<close>
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1062
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1063
definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c" where
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55412
diff changeset
  1064
  "fold f t = List.fold (case_prod f) (entries t)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1065
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1066
lemma fold_simps [simp]:
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1067
  "fold f Empty = id"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1068
  "fold f (Branch c lt k v rt) = fold f rt \<circ> f k v \<circ> fold f lt"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1069
  by (simp_all add: fold_def fun_eq_iff)
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
  1070
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1071
lemma fold_code [code]:
49810
53f14f62cca2 fix code equation for RBT_Impl.fold
Andreas Lochbihler
parents: 49807
diff changeset
  1072
  "fold f Empty x = x"
53f14f62cca2 fix code equation for RBT_Impl.fold
Andreas Lochbihler
parents: 49807
diff changeset
  1073
  "fold f (Branch c lt k v rt) x = fold f rt (f k v (fold f lt x))"
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1074
by(simp_all)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1075
67408
4a4c14b24800 prefer formal comments;
wenzelm
parents: 66808
diff changeset
  1076
\<comment> \<open>fold with continuation predicate\<close>
48621
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1077
fun foldi :: "('c \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a :: linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c" 
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1078
  where
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1079
  "foldi c f Empty s = s" |
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1080
  "foldi c f (Branch col l k v r) s = (
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1081
    if (c s) then
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1082
      let s' = foldi c f l s in
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1083
        if (c s') then
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1084
          foldi c f r (f k v s')
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1085
        else s'
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1086
    else 
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1087
      s
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1088
  )"
35606
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1089
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1090
subsection \<open>Bulkloading a tree\<close>
35606
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1091
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1092
definition (in ord) rbt_bulkload :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) rbt" where
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1093
  "rbt_bulkload xs = foldr (\<lambda>(k, v). rbt_insert k v) xs Empty"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1094
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1095
context linorder begin
35606
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1096
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1097
lemma rbt_bulkload_is_rbt [simp, intro]:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1098
  "is_rbt (rbt_bulkload xs)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1099
  unfolding rbt_bulkload_def by (induct xs) auto
35606
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1100
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1101
lemma rbt_lookup_rbt_bulkload:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1102
  "rbt_lookup (rbt_bulkload xs) = map_of xs"
35606
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1103
proof -
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1104
  obtain ys where "ys = rev xs" by simp
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1105
  have "\<And>t. is_rbt t \<Longrightarrow>
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55412
diff changeset
  1106
    rbt_lookup (List.fold (case_prod rbt_insert) ys t) = rbt_lookup t ++ map_of (rev ys)"
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55412
diff changeset
  1107
      by (induct ys) (simp_all add: rbt_bulkload_def rbt_lookup_rbt_insert case_prod_beta)
35606
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1108
  from this Empty_is_rbt have
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55412
diff changeset
  1109
    "rbt_lookup (List.fold (case_prod rbt_insert) (rev xs) Empty) = rbt_lookup Empty ++ map_of xs"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1110
     by (simp add: \<open>ys = rev xs\<close>)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1111
  then show ?thesis by (simp add: rbt_bulkload_def rbt_lookup_Empty foldr_conv_fold)
35606
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1112
qed
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1113
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1114
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1115
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1116
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1117
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1118
subsection \<open>Building a RBT from a sorted list\<close>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1119
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1120
text \<open>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1121
  These functions have been adapted from 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1122
  Andrew W. Appel, Efficient Verified Red-Black Trees (September 2011) 
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1123
\<close>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1124
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1125
fun rbtreeify_f :: "nat \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a, 'b) rbt \<times> ('a \<times> 'b) list"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1126
  and rbtreeify_g :: "nat \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a, 'b) rbt \<times> ('a \<times> 'b) list"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1127
where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1128
  "rbtreeify_f n kvs =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1129
   (if n = 0 then (Empty, kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1130
    else if n = 1 then
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1131
      case kvs of (k, v) # kvs' \<Rightarrow> (Branch R Empty k v Empty, kvs')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1132
    else if (n mod 2 = 0) then
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1133
      case rbtreeify_f (n div 2) kvs of (t1, (k, v) # kvs') \<Rightarrow>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1134
        apfst (Branch B t1 k v) (rbtreeify_g (n div 2) kvs')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1135
    else case rbtreeify_f (n div 2) kvs of (t1, (k, v) # kvs') \<Rightarrow>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1136
        apfst (Branch B t1 k v) (rbtreeify_f (n div 2) kvs'))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1137
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1138
| "rbtreeify_g n kvs =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1139
   (if n = 0 \<or> n = 1 then (Empty, kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1140
    else if n mod 2 = 0 then
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1141
      case rbtreeify_g (n div 2) kvs of (t1, (k, v) # kvs') \<Rightarrow>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1142
        apfst (Branch B t1 k v) (rbtreeify_g (n div 2) kvs')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1143
    else case rbtreeify_f (n div 2) kvs of (t1, (k, v) # kvs') \<Rightarrow>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1144
        apfst (Branch B t1 k v) (rbtreeify_g (n div 2) kvs'))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1145
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1146
definition rbtreeify :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) rbt"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1147
where "rbtreeify kvs = fst (rbtreeify_g (Suc (length kvs)) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1148
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1149
declare rbtreeify_f.simps [simp del] rbtreeify_g.simps [simp del]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1150
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1151
lemma rbtreeify_f_code [code]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1152
  "rbtreeify_f n kvs =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1153
   (if n = 0 then (Empty, kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1154
    else if n = 1 then
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1155
      case kvs of (k, v) # kvs' \<Rightarrow> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1156
        (Branch R Empty k v Empty, kvs')
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
  1157
    else let (n', r) = Euclidean_Rings.divmod_nat n 2 in
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1158
      if r = 0 then
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1159
        case rbtreeify_f n' kvs of (t1, (k, v) # kvs') \<Rightarrow>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1160
          apfst (Branch B t1 k v) (rbtreeify_g n' kvs')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1161
      else case rbtreeify_f n' kvs of (t1, (k, v) # kvs') \<Rightarrow>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1162
          apfst (Branch B t1 k v) (rbtreeify_f n' kvs'))"
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
  1163
by (subst rbtreeify_f.simps) (simp only: Let_def Euclidean_Rings.divmod_nat_def prod.case)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1164
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1165
lemma rbtreeify_g_code [code]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1166
  "rbtreeify_g n kvs =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1167
   (if n = 0 \<or> n = 1 then (Empty, kvs)
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
  1168
    else let (n', r) = Euclidean_Rings.divmod_nat n 2 in
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1169
      if r = 0 then
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1170
        case rbtreeify_g n' kvs of (t1, (k, v) # kvs') \<Rightarrow>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1171
          apfst (Branch B t1 k v) (rbtreeify_g n' kvs')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1172
      else case rbtreeify_f n' kvs of (t1, (k, v) # kvs') \<Rightarrow>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1173
          apfst (Branch B t1 k v) (rbtreeify_g n' kvs'))"
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
  1174
by(subst rbtreeify_g.simps)(simp only: Let_def Euclidean_Rings.divmod_nat_def prod.case)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1175
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1176
lemma Suc_double_half: "Suc (2 * n) div 2 = n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1177
by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1178
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1179
lemma div2_plus_div2: "n div 2 + n div 2 = (n :: nat) - n mod 2"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1180
by arith
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1181
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1182
lemma rbtreeify_f_rec_aux_lemma:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1183
  "\<lbrakk>k - n div 2 = Suc k'; n \<le> k; n mod 2 = Suc 0\<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1184
  \<Longrightarrow> k' - n div 2 = k - n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1185
apply(rule add_right_imp_eq[where a = "n - n div 2"])
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1186
apply(subst add_diff_assoc2, arith)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1187
apply(simp add: div2_plus_div2)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1188
done
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1189
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1190
lemma rbtreeify_f_simps:
59575
55f5e1cbf2a7 removed needless (and inconsistent) qualifier that messes up with Mirabelle
blanchet
parents: 59554
diff changeset
  1191
  "rbtreeify_f 0 kvs = (Empty, kvs)"
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1192
  "rbtreeify_f (Suc 0) ((k, v) # kvs) = 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1193
  (Branch R Empty k v Empty, kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1194
  "0 < n \<Longrightarrow> rbtreeify_f (2 * n) kvs =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1195
   (case rbtreeify_f n kvs of (t1, (k, v) # kvs') \<Rightarrow>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1196
     apfst (Branch B t1 k v) (rbtreeify_g n kvs'))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1197
  "0 < n \<Longrightarrow> rbtreeify_f (Suc (2 * n)) kvs =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1198
   (case rbtreeify_f n kvs of (t1, (k, v) # kvs') \<Rightarrow> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1199
     apfst (Branch B t1 k v) (rbtreeify_f n kvs'))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1200
by(subst (1) rbtreeify_f.simps, simp add: Suc_double_half)+
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1201
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1202
lemma rbtreeify_g_simps:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1203
  "rbtreeify_g 0 kvs = (Empty, kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1204
  "rbtreeify_g (Suc 0) kvs = (Empty, kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1205
  "0 < n \<Longrightarrow> rbtreeify_g (2 * n) kvs =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1206
   (case rbtreeify_g n kvs of (t1, (k, v) # kvs') \<Rightarrow> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1207
     apfst (Branch B t1 k v) (rbtreeify_g n kvs'))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1208
  "0 < n \<Longrightarrow> rbtreeify_g (Suc (2 * n)) kvs =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1209
   (case rbtreeify_f n kvs of (t1, (k, v) # kvs') \<Rightarrow> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1210
     apfst (Branch B t1 k v) (rbtreeify_g n kvs'))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1211
by(subst (1) rbtreeify_g.simps, simp add: Suc_double_half)+
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1212
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1213
declare rbtreeify_f_simps[simp] rbtreeify_g_simps[simp]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1214
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1215
lemma length_rbtreeify_f: "n \<le> length kvs
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1216
  \<Longrightarrow> length (snd (rbtreeify_f n kvs)) = length kvs - n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1217
  and length_rbtreeify_g:"\<lbrakk> 0 < n; n \<le> Suc (length kvs) \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1218
  \<Longrightarrow> length (snd (rbtreeify_g n kvs)) = Suc (length kvs) - n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1219
proof(induction n kvs and n kvs rule: rbtreeify_f_rbtreeify_g.induct)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1220
  case (1 n kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1221
  show ?case
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1222
  proof(cases "n \<le> 1")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1223
    case True thus ?thesis using "1.prems"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1224
      by(cases n kvs rule: nat.exhaust[case_product list.exhaust]) auto
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1225
  next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1226
    case False
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1227
    hence "n \<noteq> 0" "n \<noteq> 1" by simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1228
    note IH = "1.IH"[OF this]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1229
    show ?thesis
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1230
    proof(cases "n mod 2 = 0")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1231
      case True
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1232
      hence "length (snd (rbtreeify_f n kvs)) = 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1233
        length (snd (rbtreeify_f (2 * (n div 2)) kvs))"
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64243
diff changeset
  1234
        by(metis minus_nat.diff_0 minus_mod_eq_mult_div [symmetric])
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1235
      also from "1.prems" False obtain k v kvs' 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1236
        where kvs: "kvs = (k, v) # kvs'" by(cases kvs) auto
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1237
      also have "0 < n div 2" using False by(simp) 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1238
      note rbtreeify_f_simps(3)[OF this]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1239
      also note kvs[symmetric] 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1240
      also let ?rest1 = "snd (rbtreeify_f (n div 2) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1241
      from "1.prems" have "n div 2 \<le> length kvs" by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1242
      with True have len: "length ?rest1 = length kvs - n div 2" by(rule IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1243
      with "1.prems" False obtain t1 k' v' kvs''
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1244
        where kvs'': "rbtreeify_f (n div 2) kvs = (t1, (k', v') # kvs'')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1245
         by(cases ?rest1)(auto simp add: snd_def split: prod.split_asm)
55412
eb2caacf3ba4 avoid old 'prod.simps' -- better be more specific
blanchet
parents: 53374
diff changeset
  1246
      note this also note prod.case also note list.simps(5) 
eb2caacf3ba4 avoid old 'prod.simps' -- better be more specific
blanchet
parents: 53374
diff changeset
  1247
      also note prod.case also note snd_apfst
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1248
      also have "0 < n div 2" "n div 2 \<le> Suc (length kvs'')" 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1249
        using len "1.prems" False unfolding kvs'' by simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1250
      with True kvs''[symmetric] refl refl
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1251
      have "length (snd (rbtreeify_g (n div 2) kvs'')) = 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1252
        Suc (length kvs'') - n div 2" by(rule IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1253
      finally show ?thesis using len[unfolded kvs''] "1.prems" True
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64243
diff changeset
  1254
        by(simp add: Suc_diff_le[symmetric] mult_2[symmetric] minus_mod_eq_mult_div [symmetric])
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1255
    next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1256
      case False
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1257
      hence "length (snd (rbtreeify_f n kvs)) = 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1258
        length (snd (rbtreeify_f (Suc (2 * (n div 2))) kvs))"
59554
4044f53326c9 inlined rules to free user-space from technical names
haftmann
parents: 58881
diff changeset
  1259
        by (simp add: mod_eq_0_iff_dvd)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1260
      also from "1.prems" \<open>\<not> n \<le> 1\<close> obtain k v kvs' 
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1261
        where kvs: "kvs = (k, v) # kvs'" by(cases kvs) auto
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1262
      also have "0 < n div 2" using \<open>\<not> n \<le> 1\<close> by(simp) 
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1263
      note rbtreeify_f_simps(4)[OF this]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1264
      also note kvs[symmetric] 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1265
      also let ?rest1 = "snd (rbtreeify_f (n div 2) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1266
      from "1.prems" have "n div 2 \<le> length kvs" by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1267
      with False have len: "length ?rest1 = length kvs - n div 2" by(rule IH)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1268
      with "1.prems" \<open>\<not> n \<le> 1\<close> obtain t1 k' v' kvs''
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1269
        where kvs'': "rbtreeify_f (n div 2) kvs = (t1, (k', v') # kvs'')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1270
        by(cases ?rest1)(auto simp add: snd_def split: prod.split_asm)
55412
eb2caacf3ba4 avoid old 'prod.simps' -- better be more specific
blanchet
parents: 53374
diff changeset
  1271
      note this also note prod.case also note list.simps(5)
eb2caacf3ba4 avoid old 'prod.simps' -- better be more specific
blanchet
parents: 53374
diff changeset
  1272
      also note prod.case also note snd_apfst
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1273
      also have "n div 2 \<le> length kvs''" 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1274
        using len "1.prems" False unfolding kvs'' by simp arith
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1275
      with False kvs''[symmetric] refl refl
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1276
      have "length (snd (rbtreeify_f (n div 2) kvs'')) = length kvs'' - n div 2"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1277
        by(rule IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1278
      finally show ?thesis using len[unfolded kvs''] "1.prems" False
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1279
        by simp(rule rbtreeify_f_rec_aux_lemma[OF sym])
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1280
    qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1281
  qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1282
next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1283
  case (2 n kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1284
  show ?case
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1285
  proof(cases "n > 1")
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1286
    case False with \<open>0 < n\<close> show ?thesis
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1287
      by(cases n kvs rule: nat.exhaust[case_product list.exhaust]) simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1288
  next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1289
    case True
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1290
    hence "\<not> (n = 0 \<or> n = 1)" by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1291
    note IH = "2.IH"[OF this]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1292
    show ?thesis
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1293
    proof(cases "n mod 2 = 0")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1294
      case True
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1295
      hence "length (snd (rbtreeify_g n kvs)) =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1296
        length (snd (rbtreeify_g (2 * (n div 2)) kvs))"
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64243
diff changeset
  1297
        by(metis minus_nat.diff_0 minus_mod_eq_mult_div [symmetric])
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1298
      also from "2.prems" True obtain k v kvs' 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1299
        where kvs: "kvs = (k, v) # kvs'" by(cases kvs) auto
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1300
      also have "0 < n div 2" using \<open>1 < n\<close> by(simp) 
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1301
      note rbtreeify_g_simps(3)[OF this]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1302
      also note kvs[symmetric] 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1303
      also let ?rest1 = "snd (rbtreeify_g (n div 2) kvs)"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1304
      from "2.prems" \<open>1 < n\<close>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1305
      have "0 < n div 2" "n div 2 \<le> Suc (length kvs)" by simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1306
      with True have len: "length ?rest1 = Suc (length kvs) - n div 2" by(rule IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1307
      with "2.prems" obtain t1 k' v' kvs''
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1308
        where kvs'': "rbtreeify_g (n div 2) kvs = (t1, (k', v') # kvs'')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1309
        by(cases ?rest1)(auto simp add: snd_def split: prod.split_asm)
55412
eb2caacf3ba4 avoid old 'prod.simps' -- better be more specific
blanchet
parents: 53374
diff changeset
  1310
      note this also note prod.case also note list.simps(5) 
eb2caacf3ba4 avoid old 'prod.simps' -- better be more specific
blanchet
parents: 53374
diff changeset
  1311
      also note prod.case also note snd_apfst
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1312
      also have "n div 2 \<le> Suc (length kvs'')" 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1313
        using len "2.prems" unfolding kvs'' by simp
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1314
      with True kvs''[symmetric] refl refl \<open>0 < n div 2\<close>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1315
      have "length (snd (rbtreeify_g (n div 2) kvs'')) = Suc (length kvs'') - n div 2"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1316
        by(rule IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1317
      finally show ?thesis using len[unfolded kvs''] "2.prems" True
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64243
diff changeset
  1318
        by(simp add: Suc_diff_le[symmetric] mult_2[symmetric] minus_mod_eq_mult_div [symmetric])
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1319
    next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1320
      case False
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1321
      hence "length (snd (rbtreeify_g n kvs)) = 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1322
        length (snd (rbtreeify_g (Suc (2 * (n div 2))) kvs))"
59554
4044f53326c9 inlined rules to free user-space from technical names
haftmann
parents: 58881
diff changeset
  1323
        by (simp add: mod_eq_0_iff_dvd)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1324
      also from "2.prems" \<open>1 < n\<close> obtain k v kvs'
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1325
        where kvs: "kvs = (k, v) # kvs'" by(cases kvs) auto
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1326
      also have "0 < n div 2" using \<open>1 < n\<close> by(simp)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1327
      note rbtreeify_g_simps(4)[OF this]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1328
      also note kvs[symmetric] 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1329
      also let ?rest1 = "snd (rbtreeify_f (n div 2) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1330
      from "2.prems" have "n div 2 \<le> length kvs" by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1331
      with False have len: "length ?rest1 = length kvs - n div 2" by(rule IH)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1332
      with "2.prems" \<open>1 < n\<close> False obtain t1 k' v' kvs'' 
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1333
        where kvs'': "rbtreeify_f (n div 2) kvs = (t1, (k', v') # kvs'')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1334
        by(cases ?rest1)(auto simp add: snd_def split: prod.split_asm, arith)
55412
eb2caacf3ba4 avoid old 'prod.simps' -- better be more specific
blanchet
parents: 53374
diff changeset
  1335
      note this also note prod.case also note list.simps(5) 
eb2caacf3ba4 avoid old 'prod.simps' -- better be more specific
blanchet
parents: 53374
diff changeset
  1336
      also note prod.case also note snd_apfst
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1337
      also have "n div 2 \<le> Suc (length kvs'')" 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1338
        using len "2.prems" False unfolding kvs'' by simp arith
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1339
      with False kvs''[symmetric] refl refl \<open>0 < n div 2\<close>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1340
      have "length (snd (rbtreeify_g (n div 2) kvs'')) = Suc (length kvs'') - n div 2"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1341
        by(rule IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1342
      finally show ?thesis using len[unfolded kvs''] "2.prems" False
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1343
        by(simp add: div2_plus_div2)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1344
    qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1345
  qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1346
qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1347
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1348
lemma rbtreeify_induct [consumes 1, case_names f_0 f_1 f_even f_odd g_0 g_1 g_even g_odd]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1349
  fixes P Q
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1350
  defines "f0 == (\<And>kvs. P 0 kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1351
  and "f1 == (\<And>k v kvs. P (Suc 0) ((k, v) # kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1352
  and "feven ==
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1353
    (\<And>n kvs t k v kvs'. \<lbrakk> n > 0; n \<le> length kvs; P n kvs; 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1354
       rbtreeify_f n kvs = (t, (k, v) # kvs'); n \<le> Suc (length kvs'); Q n kvs' \<rbrakk> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1355
     \<Longrightarrow> P (2 * n) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1356
  and "fodd == 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1357
    (\<And>n kvs t k v kvs'. \<lbrakk> n > 0; n \<le> length kvs; P n kvs;
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1358
       rbtreeify_f n kvs = (t, (k, v) # kvs'); n \<le> length kvs'; P n kvs' \<rbrakk> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1359
    \<Longrightarrow> P (Suc (2 * n)) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1360
  and "g0 == (\<And>kvs. Q 0 kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1361
  and "g1 == (\<And>kvs. Q (Suc 0) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1362
  and "geven == 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1363
    (\<And>n kvs t k v kvs'. \<lbrakk> n > 0; n \<le> Suc (length kvs); Q n kvs; 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1364
       rbtreeify_g n kvs = (t, (k, v) # kvs'); n \<le> Suc (length kvs'); Q n kvs' \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1365
    \<Longrightarrow> Q (2 * n) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1366
  and "godd == 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1367
    (\<And>n kvs t k v kvs'. \<lbrakk> n > 0; n \<le> length kvs; P n kvs;
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1368
       rbtreeify_f n kvs = (t, (k, v) # kvs'); n \<le> Suc (length kvs'); Q n kvs' \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1369
    \<Longrightarrow> Q (Suc (2 * n)) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1370
  shows "\<lbrakk> n \<le> length kvs; 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1371
           PROP f0; PROP f1; PROP feven; PROP fodd; 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1372
           PROP g0; PROP g1; PROP geven; PROP godd \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1373
         \<Longrightarrow> P n kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1374
  and "\<lbrakk> n \<le> Suc (length kvs);
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1375
          PROP f0; PROP f1; PROP feven; PROP fodd; 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1376
          PROP g0; PROP g1; PROP geven; PROP godd \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1377
       \<Longrightarrow> Q n kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1378
proof -
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1379
  assume f0: "PROP f0" and f1: "PROP f1" and feven: "PROP feven" and fodd: "PROP fodd"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1380
    and g0: "PROP g0" and g1: "PROP g1" and geven: "PROP geven" and godd: "PROP godd"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1381
  show "n \<le> length kvs \<Longrightarrow> P n kvs" and "n \<le> Suc (length kvs) \<Longrightarrow> Q n kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1382
  proof(induction rule: rbtreeify_f_rbtreeify_g.induct)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1383
    case (1 n kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1384
    show ?case
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1385
    proof(cases "n \<le> 1")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1386
      case True thus ?thesis using "1.prems"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1387
        by(cases n kvs rule: nat.exhaust[case_product list.exhaust])
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1388
          (auto simp add: f0[unfolded f0_def] f1[unfolded f1_def])
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1389
    next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1390
      case False 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1391
      hence ns: "n \<noteq> 0" "n \<noteq> 1" by simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1392
      hence ge0: "n div 2 > 0" by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1393
      note IH = "1.IH"[OF ns]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1394
      show ?thesis
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1395
      proof(cases "n mod 2 = 0")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1396
        case True note ge0 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1397
        moreover from "1.prems" have n2: "n div 2 \<le> length kvs" by simp
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
  1398
        moreover from True n2 have "P (n div 2) kvs" by(rule IH)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1399
        moreover from length_rbtreeify_f[OF n2] ge0 "1.prems" obtain t k v kvs' 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1400
          where kvs': "rbtreeify_f (n div 2) kvs = (t, (k, v) # kvs')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1401
          by(cases "snd (rbtreeify_f (n div 2) kvs)")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1402
            (auto simp add: snd_def split: prod.split_asm)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1403
        moreover from "1.prems" length_rbtreeify_f[OF n2] ge0
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
  1404
        have n2': "n div 2 \<le> Suc (length kvs')" by(simp add: kvs')
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
  1405
        moreover from True kvs'[symmetric] refl refl n2'
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1406
        have "Q (n div 2) kvs'" by(rule IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1407
        moreover note feven[unfolded feven_def]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1408
          (* FIXME: why does by(rule feven[unfolded feven_def]) not work? *)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1409
        ultimately have "P (2 * (n div 2)) kvs" by -
64243
aee949f6642d eliminated irregular aliasses
haftmann
parents: 63680
diff changeset
  1410
        thus ?thesis using True by (metis minus_mod_eq_div_mult [symmetric] minus_nat.diff_0 mult.commute)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1411
      next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1412
        case False note ge0
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1413
        moreover from "1.prems" have n2: "n div 2 \<le> length kvs" by simp
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
  1414
        moreover from False n2 have "P (n div 2) kvs" by(rule IH)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1415
        moreover from length_rbtreeify_f[OF n2] ge0 "1.prems" obtain t k v kvs' 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1416
          where kvs': "rbtreeify_f (n div 2) kvs = (t, (k, v) # kvs')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1417
          by(cases "snd (rbtreeify_f (n div 2) kvs)")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1418
            (auto simp add: snd_def split: prod.split_asm)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1419
        moreover from "1.prems" length_rbtreeify_f[OF n2] ge0 False
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
  1420
        have n2': "n div 2 \<le> length kvs'" by(simp add: kvs') arith
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
  1421
        moreover from False kvs'[symmetric] refl refl n2' have "P (n div 2) kvs'" by(rule IH)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1422
        moreover note fodd[unfolded fodd_def]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1423
        ultimately have "P (Suc (2 * (n div 2))) kvs" by -
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1424
        thus ?thesis using False 
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64243
diff changeset
  1425
          by simp (metis One_nat_def Suc_eq_plus1_left le_add_diff_inverse mod_less_eq_dividend minus_mod_eq_mult_div [symmetric])
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1426
      qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1427
    qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1428
  next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1429
    case (2 n kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1430
    show ?case
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1431
    proof(cases "n \<le> 1")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1432
      case True thus ?thesis using "2.prems"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1433
        by(cases n kvs rule: nat.exhaust[case_product list.exhaust])
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1434
          (auto simp add: g0[unfolded g0_def] g1[unfolded g1_def])
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1435
    next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1436
      case False 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1437
      hence ns: "\<not> (n = 0 \<or> n = 1)" by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1438
      hence ge0: "n div 2 > 0" by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1439
      note IH = "2.IH"[OF ns]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1440
      show ?thesis
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1441
      proof(cases "n mod 2 = 0")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1442
        case True note ge0
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1443
        moreover from "2.prems" have n2: "n div 2 \<le> Suc (length kvs)" by simp
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
  1444
        moreover from True n2 have "Q (n div 2) kvs" by(rule IH)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1445
        moreover from length_rbtreeify_g[OF ge0 n2] ge0 "2.prems" obtain t k v kvs' 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1446
          where kvs': "rbtreeify_g (n div 2) kvs = (t, (k, v) # kvs')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1447
          by(cases "snd (rbtreeify_g (n div 2) kvs)")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1448
            (auto simp add: snd_def split: prod.split_asm)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1449
        moreover from "2.prems" length_rbtreeify_g[OF ge0 n2] ge0
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
  1450
        have n2': "n div 2 \<le> Suc (length kvs')" by(simp add: kvs')
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
  1451
        moreover from True kvs'[symmetric] refl refl  n2'
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1452
        have "Q (n div 2) kvs'" by(rule IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1453
        moreover note geven[unfolded geven_def]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1454
        ultimately have "Q (2 * (n div 2)) kvs" by -
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1455
        thus ?thesis using True 
64243
aee949f6642d eliminated irregular aliasses
haftmann
parents: 63680
diff changeset
  1456
          by(metis minus_mod_eq_div_mult [symmetric] minus_nat.diff_0 mult.commute)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1457
      next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1458
        case False note ge0
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1459
        moreover from "2.prems" have n2: "n div 2 \<le> length kvs" by simp
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
  1460
        moreover from False n2 have "P (n div 2) kvs" by(rule IH)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1461
        moreover from length_rbtreeify_f[OF n2] ge0 "2.prems" False obtain t k v kvs' 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1462
          where kvs': "rbtreeify_f (n div 2) kvs = (t, (k, v) # kvs')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1463
          by(cases "snd (rbtreeify_f (n div 2) kvs)")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1464
            (auto simp add: snd_def split: prod.split_asm, arith)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1465
        moreover from "2.prems" length_rbtreeify_f[OF n2] ge0 False
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
  1466
        have n2': "n div 2 \<le> Suc (length kvs')" by(simp add: kvs') arith
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
  1467
        moreover from False kvs'[symmetric] refl refl n2'
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1468
        have "Q (n div 2) kvs'" by(rule IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1469
        moreover note godd[unfolded godd_def]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1470
        ultimately have "Q (Suc (2 * (n div 2))) kvs" by -
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1471
        thus ?thesis using False 
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64243
diff changeset
  1472
          by simp (metis One_nat_def Suc_eq_plus1_left le_add_diff_inverse mod_less_eq_dividend minus_mod_eq_mult_div [symmetric])
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1473
      qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1474
    qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1475
  qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1476
qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1477
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1478
lemma inv1_rbtreeify_f: "n \<le> length kvs 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1479
  \<Longrightarrow> inv1 (fst (rbtreeify_f n kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1480
  and inv1_rbtreeify_g: "n \<le> Suc (length kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1481
  \<Longrightarrow> inv1 (fst (rbtreeify_g n kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1482
by(induct n kvs and n kvs rule: rbtreeify_induct) simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1483
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1484
fun plog2 :: "nat \<Rightarrow> nat" 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1485
where "plog2 n = (if n \<le> 1 then 0 else plog2 (n div 2) + 1)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1486
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1487
declare plog2.simps [simp del]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1488
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1489
lemma plog2_simps [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1490
  "plog2 0 = 0" "plog2 (Suc 0) = 0"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1491
  "0 < n \<Longrightarrow> plog2 (2 * n) = 1 + plog2 n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1492
  "0 < n \<Longrightarrow> plog2 (Suc (2 * n)) = 1 + plog2 n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1493
by(subst plog2.simps, simp add: Suc_double_half)+
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1494
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1495
lemma bheight_rbtreeify_f: "n \<le> length kvs
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1496
  \<Longrightarrow> bheight (fst (rbtreeify_f n kvs)) = plog2 n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1497
  and bheight_rbtreeify_g: "n \<le> Suc (length kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1498
  \<Longrightarrow> bheight (fst (rbtreeify_g n kvs)) = plog2 n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1499
by(induct n kvs and n kvs rule: rbtreeify_induct) simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1500
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1501
lemma bheight_rbtreeify_f_eq_plog2I:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1502
  "\<lbrakk> rbtreeify_f n kvs = (t, kvs'); n \<le> length kvs \<rbrakk> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1503
  \<Longrightarrow> bheight t = plog2 n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1504
using bheight_rbtreeify_f[of n kvs] by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1505
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1506
lemma bheight_rbtreeify_g_eq_plog2I: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1507
  "\<lbrakk> rbtreeify_g n kvs = (t, kvs'); n \<le> Suc (length kvs) \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1508
  \<Longrightarrow> bheight t = plog2 n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1509
using bheight_rbtreeify_g[of n kvs] by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1510
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1511
hide_const (open) plog2
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1512
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1513
lemma inv2_rbtreeify_f: "n \<le> length kvs
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1514
  \<Longrightarrow> inv2 (fst (rbtreeify_f n kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1515
  and inv2_rbtreeify_g: "n \<le> Suc (length kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1516
  \<Longrightarrow> inv2 (fst (rbtreeify_g n kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1517
by(induct n kvs and n kvs rule: rbtreeify_induct)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1518
  (auto simp add: bheight_rbtreeify_f bheight_rbtreeify_g 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1519
        intro: bheight_rbtreeify_f_eq_plog2I bheight_rbtreeify_g_eq_plog2I)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1520
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1521
lemma "n \<le> length kvs \<Longrightarrow> True"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1522
  and color_of_rbtreeify_g:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1523
  "\<lbrakk> n \<le> Suc (length kvs); 0 < n \<rbrakk> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1524
  \<Longrightarrow> color_of (fst (rbtreeify_g n kvs)) = B"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1525
by(induct n kvs and n kvs rule: rbtreeify_induct) simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1526
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1527
lemma entries_rbtreeify_f_append:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1528
  "n \<le> length kvs 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1529
  \<Longrightarrow> entries (fst (rbtreeify_f n kvs)) @ snd (rbtreeify_f n kvs) = kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1530
  and entries_rbtreeify_g_append: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1531
  "n \<le> Suc (length kvs) 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1532
  \<Longrightarrow> entries (fst (rbtreeify_g n kvs)) @ snd (rbtreeify_g n kvs) = kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1533
by(induction rule: rbtreeify_induct) simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1534
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1535
lemma length_entries_rbtreeify_f:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1536
  "n \<le> length kvs \<Longrightarrow> length (entries (fst (rbtreeify_f n kvs))) = n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1537
  and length_entries_rbtreeify_g: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1538
  "n \<le> Suc (length kvs) \<Longrightarrow> length (entries (fst (rbtreeify_g n kvs))) = n - 1"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1539
by(induct rule: rbtreeify_induct) simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1540
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1541
lemma rbtreeify_f_conv_drop: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1542
  "n \<le> length kvs \<Longrightarrow> snd (rbtreeify_f n kvs) = drop n kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1543
using entries_rbtreeify_f_append[of n kvs]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1544
by(simp add: append_eq_conv_conj length_entries_rbtreeify_f)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1545
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1546
lemma rbtreeify_g_conv_drop: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1547
  "n \<le> Suc (length kvs) \<Longrightarrow> snd (rbtreeify_g n kvs) = drop (n - 1) kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1548
using entries_rbtreeify_g_append[of n kvs]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1549
by(simp add: append_eq_conv_conj length_entries_rbtreeify_g)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1550
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1551
lemma entries_rbtreeify_f [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1552
  "n \<le> length kvs \<Longrightarrow> entries (fst (rbtreeify_f n kvs)) = take n kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1553
using entries_rbtreeify_f_append[of n kvs]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1554
by(simp add: append_eq_conv_conj length_entries_rbtreeify_f)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1555
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1556
lemma entries_rbtreeify_g [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1557
  "n \<le> Suc (length kvs) \<Longrightarrow> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1558
  entries (fst (rbtreeify_g n kvs)) = take (n - 1) kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1559
using entries_rbtreeify_g_append[of n kvs]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1560
by(simp add: append_eq_conv_conj length_entries_rbtreeify_g)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1561
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1562
lemma keys_rbtreeify_f [simp]: "n \<le> length kvs
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1563
  \<Longrightarrow> keys (fst (rbtreeify_f n kvs)) = take n (map fst kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1564
by(simp add: keys_def take_map)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1565
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1566
lemma keys_rbtreeify_g [simp]: "n \<le> Suc (length kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1567
  \<Longrightarrow> keys (fst (rbtreeify_g n kvs)) = take (n - 1) (map fst kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1568
by(simp add: keys_def take_map)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1569
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1570
lemma rbtreeify_fD: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1571
  "\<lbrakk> rbtreeify_f n kvs = (t, kvs'); n \<le> length kvs \<rbrakk> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1572
  \<Longrightarrow> entries t = take n kvs \<and> kvs' = drop n kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1573
using rbtreeify_f_conv_drop[of n kvs] entries_rbtreeify_f[of n kvs] by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1574
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1575
lemma rbtreeify_gD: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1576
  "\<lbrakk> rbtreeify_g n kvs = (t, kvs'); n \<le> Suc (length kvs) \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1577
  \<Longrightarrow> entries t = take (n - 1) kvs \<and> kvs' = drop (n - 1) kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1578
using rbtreeify_g_conv_drop[of n kvs] entries_rbtreeify_g[of n kvs] by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1579
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1580
lemma entries_rbtreeify [simp]: "entries (rbtreeify kvs) = kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1581
by(simp add: rbtreeify_def entries_rbtreeify_g)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1582
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1583
context linorder begin
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1584
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1585
lemma rbt_sorted_rbtreeify_f: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1586
  "\<lbrakk> n \<le> length kvs; sorted (map fst kvs); distinct (map fst kvs) \<rbrakk> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1587
  \<Longrightarrow> rbt_sorted (fst (rbtreeify_f n kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1588
  and rbt_sorted_rbtreeify_g: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1589
  "\<lbrakk> n \<le> Suc (length kvs); sorted (map fst kvs); distinct (map fst kvs) \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1590
  \<Longrightarrow> rbt_sorted (fst (rbtreeify_g n kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1591
proof(induction n kvs and n kvs rule: rbtreeify_induct)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1592
  case (f_even n kvs t k v kvs')
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1593
  from rbtreeify_fD[OF \<open>rbtreeify_f n kvs = (t, (k, v) # kvs')\<close> \<open>n \<le> length kvs\<close>]
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1594
  have "entries t = take n kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1595
    and kvs': "drop n kvs = (k, v) # kvs'" by simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1596
  hence unfold: "kvs = take n kvs @ (k, v) # kvs'" by(metis append_take_drop_id)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1597
  from \<open>sorted (map fst kvs)\<close> kvs'
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1598
  have "(\<forall>(x, y) \<in> set (take n kvs). x \<le> k) \<and> (\<forall>(x, y) \<in> set kvs'. k \<le> x)"
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1599
    by(subst (asm) unfold)(auto simp add: sorted_append)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1600
  moreover from \<open>distinct (map fst kvs)\<close> kvs'
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1601
  have "(\<forall>(x, y) \<in> set (take n kvs). x \<noteq> k) \<and> (\<forall>(x, y) \<in> set kvs'. x \<noteq> k)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1602
    by(subst (asm) unfold)(auto intro: rev_image_eqI)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1603
  ultimately have "(\<forall>(x, y) \<in> set (take n kvs). x < k) \<and> (\<forall>(x, y) \<in> set kvs'. k < x)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1604
    by fastforce
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1605
  hence "fst (rbtreeify_f n kvs) |\<guillemotleft> k" "k \<guillemotleft>| fst (rbtreeify_g n kvs')"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1606
    using \<open>n \<le> Suc (length kvs')\<close> \<open>n \<le> length kvs\<close> set_take_subset[of "n - 1" kvs']
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1607
    by(auto simp add: ord.rbt_greater_prop ord.rbt_less_prop take_map split_def)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1608
  moreover from \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1609
  have "rbt_sorted (fst (rbtreeify_f n kvs))" by(rule f_even.IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1610
  moreover have "sorted (map fst kvs')" "distinct (map fst kvs')"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1611
    using \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1612
    by(subst (asm) (1 2) unfold, simp add: sorted_append)+
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1613
  hence "rbt_sorted (fst (rbtreeify_g n kvs'))" by(rule f_even.IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1614
  ultimately show ?case
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1615
    using \<open>0 < n\<close> \<open>rbtreeify_f n kvs = (t, (k, v) # kvs')\<close> by simp
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1616
next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1617
  case (f_odd n kvs t k v kvs')
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1618
  from rbtreeify_fD[OF \<open>rbtreeify_f n kvs = (t, (k, v) # kvs')\<close> \<open>n \<le> length kvs\<close>]
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1619
  have "entries t = take n kvs" 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1620
    and kvs': "drop n kvs = (k, v) # kvs'" by simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1621
  hence unfold: "kvs = take n kvs @ (k, v) # kvs'" by(metis append_take_drop_id)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1622
  from \<open>sorted (map fst kvs)\<close> kvs'
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1623
  have "(\<forall>(x, y) \<in> set (take n kvs). x \<le> k) \<and> (\<forall>(x, y) \<in> set kvs'. k \<le> x)"
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1624
    by(subst (asm) unfold)(auto simp add: sorted_append)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1625
  moreover from \<open>distinct (map fst kvs)\<close> kvs'
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1626
  have "(\<forall>(x, y) \<in> set (take n kvs). x \<noteq> k) \<and> (\<forall>(x, y) \<in> set kvs'. x \<noteq> k)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1627
    by(subst (asm) unfold)(auto intro: rev_image_eqI)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1628
  ultimately have "(\<forall>(x, y) \<in> set (take n kvs). x < k) \<and> (\<forall>(x, y) \<in> set kvs'. k < x)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1629
    by fastforce
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1630
  hence "fst (rbtreeify_f n kvs) |\<guillemotleft> k" "k \<guillemotleft>| fst (rbtreeify_f n kvs')"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1631
    using \<open>n \<le> length kvs'\<close> \<open>n \<le> length kvs\<close> set_take_subset[of n kvs']
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1632
    by(auto simp add: rbt_greater_prop rbt_less_prop take_map split_def)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1633
  moreover from \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1634
  have "rbt_sorted (fst (rbtreeify_f n kvs))" by(rule f_odd.IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1635
  moreover have "sorted (map fst kvs')" "distinct (map fst kvs')"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1636
    using \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1637
    by(subst (asm) (1 2) unfold, simp add: sorted_append)+
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1638
  hence "rbt_sorted (fst (rbtreeify_f n kvs'))" by(rule f_odd.IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1639
  ultimately show ?case 
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1640
    using \<open>0 < n\<close> \<open>rbtreeify_f n kvs = (t, (k, v) # kvs')\<close> by simp
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1641
next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1642
  case (g_even n kvs t k v kvs')
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1643
  from rbtreeify_gD[OF \<open>rbtreeify_g n kvs = (t, (k, v) # kvs')\<close> \<open>n \<le> Suc (length kvs)\<close>]
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1644
  have t: "entries t = take (n - 1) kvs" 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1645
    and kvs': "drop (n - 1) kvs = (k, v) # kvs'" by simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1646
  hence unfold: "kvs = take (n - 1) kvs @ (k, v) # kvs'" by(metis append_take_drop_id)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1647
  from \<open>sorted (map fst kvs)\<close> kvs'
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1648
  have "(\<forall>(x, y) \<in> set (take (n - 1) kvs). x \<le> k) \<and> (\<forall>(x, y) \<in> set kvs'. k \<le> x)"
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1649
    by(subst (asm) unfold)(auto simp add: sorted_append)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1650
  moreover from \<open>distinct (map fst kvs)\<close> kvs'
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1651
  have "(\<forall>(x, y) \<in> set (take (n - 1) kvs). x \<noteq> k) \<and> (\<forall>(x, y) \<in> set kvs'. x \<noteq> k)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1652
    by(subst (asm) unfold)(auto intro: rev_image_eqI)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1653
  ultimately have "(\<forall>(x, y) \<in> set (take (n - 1) kvs). x < k) \<and> (\<forall>(x, y) \<in> set kvs'. k < x)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1654
    by fastforce
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1655
  hence "fst (rbtreeify_g n kvs) |\<guillemotleft> k" "k \<guillemotleft>| fst (rbtreeify_g n kvs')"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1656
    using \<open>n \<le> Suc (length kvs')\<close> \<open>n \<le> Suc (length kvs)\<close> set_take_subset[of "n - 1" kvs']
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1657
    by(auto simp add: rbt_greater_prop rbt_less_prop take_map split_def)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1658
  moreover from \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1659
  have "rbt_sorted (fst (rbtreeify_g n kvs))" by(rule g_even.IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1660
  moreover have "sorted (map fst kvs')" "distinct (map fst kvs')"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1661
    using \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1662
    by(subst (asm) (1 2) unfold, simp add: sorted_append)+
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1663
  hence "rbt_sorted (fst (rbtreeify_g n kvs'))" by(rule g_even.IH)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1664
  ultimately show ?case using \<open>0 < n\<close> \<open>rbtreeify_g n kvs = (t, (k, v) # kvs')\<close> by simp
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1665
next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1666
  case (g_odd n kvs t k v kvs')
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1667
  from rbtreeify_fD[OF \<open>rbtreeify_f n kvs = (t, (k, v) # kvs')\<close> \<open>n \<le> length kvs\<close>]
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1668
  have "entries t = take n kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1669
    and kvs': "drop n kvs = (k, v) # kvs'" by simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1670
  hence unfold: "kvs = take n kvs @ (k, v) # kvs'" by(metis append_take_drop_id)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1671
  from \<open>sorted (map fst kvs)\<close> kvs'
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1672
  have "(\<forall>(x, y) \<in> set (take n kvs). x \<le> k) \<and> (\<forall>(x, y) \<in> set kvs'. k \<le> x)"
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1673
    by(subst (asm) unfold)(auto simp add: sorted_append)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1674
  moreover from \<open>distinct (map fst kvs)\<close> kvs'
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1675
  have "(\<forall>(x, y) \<in> set (take n kvs). x \<noteq> k) \<and> (\<forall>(x, y) \<in> set kvs'. x \<noteq> k)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1676
    by(subst (asm) unfold)(auto intro: rev_image_eqI)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1677
  ultimately have "(\<forall>(x, y) \<in> set (take n kvs). x < k) \<and> (\<forall>(x, y) \<in> set kvs'. k < x)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1678
    by fastforce
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1679
  hence "fst (rbtreeify_f n kvs) |\<guillemotleft> k" "k \<guillemotleft>| fst (rbtreeify_g n kvs')"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1680
    using \<open>n \<le> Suc (length kvs')\<close> \<open>n \<le> length kvs\<close> set_take_subset[of "n - 1" kvs']
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1681
    by(auto simp add: rbt_greater_prop rbt_less_prop take_map split_def)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1682
  moreover from \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1683
  have "rbt_sorted (fst (rbtreeify_f n kvs))" by(rule g_odd.IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1684
  moreover have "sorted (map fst kvs')" "distinct (map fst kvs')"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1685
    using \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1686
    by(subst (asm) (1 2) unfold, simp add: sorted_append)+
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1687
  hence "rbt_sorted (fst (rbtreeify_g n kvs'))" by(rule g_odd.IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1688
  ultimately show ?case
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1689
    using \<open>0 < n\<close> \<open>rbtreeify_f n kvs = (t, (k, v) # kvs')\<close> by simp
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1690
qed simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1691
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1692
lemma rbt_sorted_rbtreeify: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1693
  "\<lbrakk> sorted (map fst kvs); distinct (map fst kvs) \<rbrakk> \<Longrightarrow> rbt_sorted (rbtreeify kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1694
by(simp add: rbtreeify_def rbt_sorted_rbtreeify_g)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1695
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1696
lemma is_rbt_rbtreeify: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1697
  "\<lbrakk> sorted (map fst kvs); distinct (map fst kvs) \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1698
  \<Longrightarrow> is_rbt (rbtreeify kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1699
by(simp add: is_rbt_def rbtreeify_def inv1_rbtreeify_g inv2_rbtreeify_g rbt_sorted_rbtreeify_g color_of_rbtreeify_g)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1700
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1701
lemma rbt_lookup_rbtreeify:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1702
  "\<lbrakk> sorted (map fst kvs); distinct (map fst kvs) \<rbrakk> \<Longrightarrow> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1703
  rbt_lookup (rbtreeify kvs) = map_of kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1704
by(simp add: map_of_entries[symmetric] rbt_sorted_rbtreeify)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1705
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1706
end
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1707
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1708
text \<open>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1709
  Functions to compare the height of two rbt trees, taken from 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1710
  Andrew W. Appel, Efficient Verified Red-Black Trees (September 2011)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1711
\<close>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1712
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1713
fun skip_red :: "('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1714
where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1715
  "skip_red (Branch color.R l k v r) = l"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1716
| "skip_red t = t"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1717
49807
9a0843995fd3 correct definition for skip_black
Andreas Lochbihler
parents: 49770
diff changeset
  1718
definition skip_black :: "('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt"
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1719
where
49807
9a0843995fd3 correct definition for skip_black
Andreas Lochbihler
parents: 49770
diff changeset
  1720
  "skip_black t = (let t' = skip_red t in case t' of Branch color.B l k v r \<Rightarrow> l | _ \<Rightarrow> t')"
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1721
58310
91ea607a34d8 updated news
blanchet
parents: 58257
diff changeset
  1722
datatype compare = LT | GT | EQ
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1723
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1724
partial_function (tailrec) compare_height :: "('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> compare"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1725
where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1726
  "compare_height sx s t tx =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1727
  (case (skip_red sx, skip_red s, skip_red t, skip_red tx) of
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1728
     (Branch _ sx' _ _ _, Branch _ s' _ _ _, Branch _ t' _ _ _, Branch _ tx' _ _ _) \<Rightarrow> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1729
       compare_height (skip_black sx') s' t' (skip_black tx')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1730
   | (_, rbt.Empty, _, Branch _ _ _ _ _) \<Rightarrow> LT
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1731
   | (Branch _ _ _ _ _, _, rbt.Empty, _) \<Rightarrow> GT
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1732
   | (Branch _ sx' _ _ _, Branch _ s' _ _ _, Branch _ t' _ _ _, rbt.Empty) \<Rightarrow>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1733
       compare_height (skip_black sx') s' t' rbt.Empty
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1734
   | (rbt.Empty, Branch _ s' _ _ _, Branch _ t' _ _ _, Branch _ tx' _ _ _) \<Rightarrow>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1735
       compare_height rbt.Empty s' t' (skip_black tx')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1736
   | _ \<Rightarrow> EQ)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1737
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1738
declare compare_height.simps [code]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1739
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1740
hide_type (open) compare
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1741
hide_const (open)
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 55414
diff changeset
  1742
  compare_height skip_black skip_red LT GT EQ case_compare rec_compare
58257
0662f35534fe half-ported Imperative HOL to new datatypes
blanchet
parents: 58249
diff changeset
  1743
  Abs_compare Rep_compare
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1744
hide_fact (open)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1745
  Abs_compare_cases Abs_compare_induct Abs_compare_inject Abs_compare_inverse
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1746
  Rep_compare Rep_compare_cases Rep_compare_induct Rep_compare_inject Rep_compare_inverse
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55466
diff changeset
  1747
  compare.simps compare.exhaust compare.induct compare.rec compare.simps
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57512
diff changeset
  1748
  compare.size compare.case_cong compare.case_cong_weak compare.case
62093
bd73a2279fcd more uniform treatment of package internals;
wenzelm
parents: 61585
diff changeset
  1749
  compare.nchotomy compare.split compare.split_asm compare.eq.refl compare.eq.simps
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1750
  equal_compare_def
61121
efe8b18306b7 do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
wenzelm
parents: 61076
diff changeset
  1751
  skip_red.simps skip_red.cases skip_red.induct 
49807
9a0843995fd3 correct definition for skip_black
Andreas Lochbihler
parents: 49770
diff changeset
  1752
  skip_black_def
61121
efe8b18306b7 do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
wenzelm
parents: 61076
diff changeset
  1753
  compare_height.simps
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1754
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  1755
subsection \<open>union and intersection of sorted associative lists\<close>
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1756
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1757
context ord begin
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1758
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1759
function sunion_with :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a \<times> 'b) list" 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1760
where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1761
  "sunion_with f ((k, v) # as) ((k', v') # bs) =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1762
   (if k > k' then (k', v') # sunion_with f ((k, v) # as) bs
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1763
    else if k < k' then (k, v) # sunion_with f as ((k', v') # bs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1764
    else (k, f k v v') # sunion_with f as bs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1765
| "sunion_with f [] bs = bs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1766
| "sunion_with f as [] = as"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1767
by pat_completeness auto
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1768
termination by lexicographic_order
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1769
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1770
function sinter_with :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a \<times> 'b) list"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1771
where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1772
  "sinter_with f ((k, v) # as) ((k', v') # bs) =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1773
  (if k > k' then sinter_with f ((k, v) # as) bs
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1774
   else if k < k' then sinter_with f as ((k', v') # bs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1775
   else (k, f k v v') # sinter_with f as bs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1776
| "sinter_with f [] _ = []"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1777
| "sinter_with f _ [] = []"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1778
by pat_completeness auto
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1779
termination by lexicographic_order
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1780
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1781
end
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1782
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1783
declare ord.sunion_with.simps [code] ord.sinter_with.simps[code]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1784
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1785
context linorder begin
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1786
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1787
lemma set_fst_sunion_with: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1788
  "set (map fst (sunion_with f xs ys)) = set (map fst xs) \<union> set (map fst ys)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1789
by(induct f xs ys rule: sunion_with.induct) auto
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1790
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1791
lemma sorted_sunion_with [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1792
  "\<lbrakk> sorted (map fst xs); sorted (map fst ys) \<rbrakk> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1793
  \<Longrightarrow> sorted (map fst (sunion_with f xs ys))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1794
by(induct f xs ys rule: sunion_with.induct)
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1795
  (auto simp add: set_fst_sunion_with simp del: set_map)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1796
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1797
lemma distinct_sunion_with [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1798
  "\<lbrakk> distinct (map fst xs); distinct (map fst ys); sorted (map fst xs); sorted (map fst ys) \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1799
  \<Longrightarrow> distinct (map fst (sunion_with f xs ys))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1800
proof(induct f xs ys rule: sunion_with.induct)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1801
  case (1 f k v xs k' v' ys)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1802
  have "\<lbrakk> \<not> k < k'; \<not> k' < k \<rbrakk> \<Longrightarrow> k = k'" by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1803
  thus ?case using "1"
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1804
    by(auto simp add: set_fst_sunion_with simp del: set_map)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1805
qed simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1806
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1807
lemma map_of_sunion_with: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1808
  "\<lbrakk> sorted (map fst xs); sorted (map fst ys) \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1809
  \<Longrightarrow> map_of (sunion_with f xs ys) k = 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1810
  (case map_of xs k of None \<Rightarrow> map_of ys k 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1811
  | Some v \<Rightarrow> case map_of ys k of None \<Rightarrow> Some v 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1812
              | Some w \<Rightarrow> Some (f k v w))"
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1813
by(induct f xs ys rule: sunion_with.induct)(auto split: option.split dest: map_of_SomeD bspec)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1814
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1815
lemma set_fst_sinter_with [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1816
  "\<lbrakk> sorted (map fst xs); sorted (map fst ys) \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1817
  \<Longrightarrow> set (map fst (sinter_with f xs ys)) = set (map fst xs) \<inter> set (map fst ys)"
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1818
by(induct f xs ys rule: sinter_with.induct)(auto simp del: set_map)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1819
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1820
lemma set_fst_sinter_with_subset1:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1821
  "set (map fst (sinter_with f xs ys)) \<subseteq> set (map fst xs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1822
by(induct f xs ys rule: sinter_with.induct) auto
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1823
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1824
lemma set_fst_sinter_with_subset2:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1825
  "set (map fst (sinter_with f xs ys)) \<subseteq> set (map fst ys)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1826
by(induct f xs ys rule: sinter_with.induct)(auto simp del: set_map)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1827
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1828
lemma sorted_sinter_with [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1829
  "\<lbrakk> sorted (map fst xs); sorted (map fst ys) \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1830
  \<Longrightarrow> sorted (map fst (sinter_with f xs ys))"
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1831
by(induct f xs ys rule: sinter_with.induct)(auto simp del: set_map)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1832
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1833
lemma distinct_sinter_with [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1834
  "\<lbrakk> distinct (map fst xs); distinct (map fst ys) \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1835
  \<Longrightarrow> distinct (map fst (sinter_with f xs ys))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1836
proof(induct f xs ys rule: sinter_with.induct)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1837
  case (1 f k v as k' v' bs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1838
  have "\<lbrakk> \<not> k < k'; \<not> k' < k \<rbrakk> \<Longrightarrow> k = k'" by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1839
  thus ?case using "1" set_fst_sinter_with_subset1[of f as bs]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1840
    set_fst_sinter_with_subset2[of f as bs]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1841
    by(auto simp del: set_map)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1842
qed simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1843
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1844
lemma map_of_sinter_with:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1845
  "\<lbrakk> sorted (map fst xs); sorted (map fst ys) \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1846
  \<Longrightarrow> map_of (sinter_with f xs ys) k = 
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55417
diff changeset
  1847
  (case map_of xs k of None \<Rightarrow> None | Some v \<Rightarrow> map_option (f k v) (map_of ys k))"
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1848
apply(induct f xs ys rule: sinter_with.induct)
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67408
diff changeset
  1849
apply(auto simp add: map_option_case split: option.splits dest: map_of_SomeD bspec)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1850
done
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1851
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1852
end
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1853
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1854
lemma distinct_map_of_rev: "distinct (map fst xs) \<Longrightarrow> map_of (rev xs) = map_of xs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1855
by(induct xs)(auto 4 3 simp add: map_add_def intro!: ext split: option.split intro: rev_image_eqI)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1856
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1857
lemma map_map_filter: 
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55417
diff changeset
  1858
  "map f (List.map_filter g xs) = List.map_filter (map_option f \<circ> g) xs"
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1859
by(auto simp add: List.map_filter_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1860
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55417
diff changeset
  1861
lemma map_filter_map_option_const: 
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55417
diff changeset
  1862
  "List.map_filter (\<lambda>x. map_option (\<lambda>y. f x) (g (f x))) xs = filter (\<lambda>x. g x \<noteq> None) (map f xs)"
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1863
by(auto simp add: map_filter_def filter_map o_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1864
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1865
lemma set_map_filter: "set (List.map_filter P xs) = the ` (P ` set xs - {None})"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1866
by(auto simp add: List.map_filter_def intro: rev_image_eqI)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1867
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1868
(* Split and Join *)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1869
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1870
definition is_rbt_empty :: "('a, 'b) rbt \<Rightarrow> bool" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1871
  "is_rbt_empty t \<longleftrightarrow> (case t of RBT_Impl.Empty \<Rightarrow> True | _ \<Rightarrow> False)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1872
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1873
lemma is_rbt_empty_prop[simp]: "is_rbt_empty t \<longleftrightarrow> t = RBT_Impl.Empty"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1874
  by (auto simp: is_rbt_empty_def split: RBT_Impl.rbt.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1875
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1876
definition small_rbt :: "('a, 'b) rbt \<Rightarrow> bool" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1877
  "small_rbt t \<longleftrightarrow> bheight t < 4"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1878
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1879
definition flip_rbt :: "('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1880
  "flip_rbt t1 t2 \<longleftrightarrow> bheight t2 < bheight t1"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1881
73212
87e3c180044a hide the internal abbreviations MR and MB
Andreas Lochbihler <mail@andreas-lochbihler.de>
parents: 73211
diff changeset
  1882
abbreviation (input) MR where "MR l a b r \<equiv> Branch RBT_Impl.R l a b r"
87e3c180044a hide the internal abbreviations MR and MB
Andreas Lochbihler <mail@andreas-lochbihler.de>
parents: 73211
diff changeset
  1883
abbreviation (input) MB where "MB l a b r \<equiv> Branch RBT_Impl.B l a b r"
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1884
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1885
fun rbt_baliL :: "('a, 'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1886
  "rbt_baliL (MR (MR t1 a b t2) a' b' t3) a'' b'' t4 = MR (MB t1 a b t2) a' b' (MB t3 a'' b'' t4)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1887
| "rbt_baliL (MR t1 a b (MR t2 a' b' t3)) a'' b'' t4 = MR (MB t1 a b t2) a' b' (MB t3 a'' b'' t4)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1888
| "rbt_baliL t1 a b t2 = MB t1 a b t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1889
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1890
fun rbt_baliR :: "('a, 'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1891
  "rbt_baliR t1 a b (MR t2 a' b' (MR t3 a'' b'' t4)) = MR (MB t1 a b t2) a' b' (MB t3 a'' b'' t4)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1892
| "rbt_baliR t1 a b (MR (MR t2 a' b' t3) a'' b'' t4) = MR (MB t1 a b t2) a' b' (MB t3 a'' b'' t4)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1893
| "rbt_baliR t1 a b t2 = MB t1 a b t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1894
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1895
fun rbt_baldL :: "('a, 'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1896
  "rbt_baldL (MR t1 a b t2) a' b' t3 = MR (MB t1 a b t2) a' b' t3"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1897
| "rbt_baldL t1 a b (MB t2 a' b' t3) = rbt_baliR t1 a b (MR t2 a' b' t3)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1898
| "rbt_baldL t1 a b (MR (MB t2 a' b' t3) a'' b'' t4) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1899
  MR (MB t1 a b t2) a' b' (rbt_baliR t3 a'' b'' (paint RBT_Impl.R t4))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1900
| "rbt_baldL t1 a b t2 = MR t1 a b t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1901
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1902
fun rbt_baldR :: "('a, 'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1903
  "rbt_baldR t1 a b (MR t2 a' b' t3) = MR t1 a b (MB t2 a' b' t3)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1904
| "rbt_baldR (MB t1 a b t2) a' b' t3 = rbt_baliL (MR t1 a b t2) a' b' t3"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1905
| "rbt_baldR (MR t1 a b (MB t2 a' b' t3)) a'' b'' t4 =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1906
  MR (rbt_baliL (paint RBT_Impl.R t1) a b t2) a' b' (MB t3 a'' b'' t4)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1907
| "rbt_baldR t1 a b t2 = MR t1 a b t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1908
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1909
fun rbt_app :: "('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1910
  "rbt_app RBT_Impl.Empty t = t"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1911
| "rbt_app t RBT_Impl.Empty = t"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1912
| "rbt_app (MR t1 a b t2) (MR t3 a'' b'' t4) = (case rbt_app t2 t3 of
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1913
    MR u2 a' b' u3 \<Rightarrow> (MR (MR t1 a b u2) a' b' (MR u3 a'' b'' t4))
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1914
  | t23 \<Rightarrow> MR t1 a b (MR t23 a'' b'' t4))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1915
| "rbt_app (MB t1 a b t2) (MB t3 a'' b'' t4) = (case rbt_app t2 t3 of
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1916
    MR u2 a' b' u3 \<Rightarrow> MR (MB t1 a b u2) a' b' (MB u3 a'' b'' t4)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1917
  | t23 \<Rightarrow> rbt_baldL t1 a b (MB t23 a'' b'' t4))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1918
| "rbt_app t1 (MR t2 a b t3) = MR (rbt_app t1 t2) a b t3"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1919
| "rbt_app (MR t1 a b t2) t3 = MR t1 a b (rbt_app t2 t3)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1920
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1921
fun rbt_joinL :: "('a, 'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1922
  "rbt_joinL l a b r = (if bheight l \<ge> bheight r then MR l a b r
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1923
    else case r of MB l' a' b' r' \<Rightarrow> rbt_baliL (rbt_joinL l a b l') a' b' r'
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1924
    | MR l' a' b' r' \<Rightarrow> MR (rbt_joinL l a b l') a' b' r')"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1925
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1926
declare rbt_joinL.simps[simp del]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1927
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1928
fun rbt_joinR :: "('a, 'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1929
  "rbt_joinR l a b r = (if bheight l \<le> bheight r then MR l a b r
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1930
    else case l of MB l' a' b' r' \<Rightarrow> rbt_baliR l' a' b' (rbt_joinR r' a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1931
    | MR l' a' b' r' \<Rightarrow> MR l' a' b' (rbt_joinR r' a b r))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1932
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1933
declare rbt_joinR.simps[simp del]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1934
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1935
definition rbt_join :: "('a, 'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1936
  "rbt_join l a b r =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1937
    (let bhl = bheight l; bhr = bheight r
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1938
    in if bhl > bhr
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1939
    then paint RBT_Impl.B (rbt_joinR l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1940
    else if bhl < bhr
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1941
    then paint RBT_Impl.B (rbt_joinL l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1942
    else MB l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1943
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1944
lemma size_paint[simp]: "size (paint c t) = size t"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1945
  by (cases t) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1946
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1947
lemma size_baliL[simp]: "size (rbt_baliL t1 a b t2) = Suc (size t1 + size t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1948
  by (induction t1 a b t2 rule: rbt_baliL.induct) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1949
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1950
lemma size_baliR[simp]: "size (rbt_baliR t1 a b t2) = Suc (size t1 + size t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1951
  by (induction t1 a b t2 rule: rbt_baliR.induct) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1952
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1953
lemma size_baldL[simp]: "size (rbt_baldL t1 a b t2) = Suc (size t1 + size t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1954
  by (induction t1 a b t2 rule: rbt_baldL.induct) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1955
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1956
lemma size_baldR[simp]: "size (rbt_baldR t1 a b t2) = Suc (size t1 + size t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1957
  by (induction t1 a b t2 rule: rbt_baldR.induct) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1958
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1959
lemma size_rbt_app[simp]: "size (rbt_app t1 t2) = size t1 + size t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1960
  by (induction t1 t2 rule: rbt_app.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1961
     (auto split: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1962
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1963
lemma size_rbt_joinL[simp]: "size (rbt_joinL t1 a b t2) = Suc (size t1 + size t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1964
  by (induction t1 a b t2 rule: rbt_joinL.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1965
     (auto simp: rbt_joinL.simps split: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1966
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1967
lemma size_rbt_joinR[simp]: "size (rbt_joinR t1 a b t2) = Suc (size t1 + size t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1968
  by (induction t1 a b t2 rule: rbt_joinR.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1969
     (auto simp: rbt_joinR.simps split: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1970
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1971
lemma size_rbt_join[simp]: "size (rbt_join t1 a b t2) = Suc (size t1 + size t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1972
  by (auto simp: rbt_join_def Let_def)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1973
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1974
definition "inv_12 t \<longleftrightarrow> inv1 t \<and> inv2 t"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1975
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1976
lemma rbt_Node: "inv_12 (RBT_Impl.Branch c l a b r) \<Longrightarrow> inv_12 l \<and> inv_12 r"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1977
  by (auto simp: inv_12_def)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1978
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1979
lemma paint2: "paint c2 (paint c1 t) = paint c2 t"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1980
  by (cases t) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1981
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1982
lemma inv1_rbt_baliL: "inv1l l \<Longrightarrow> inv1 r \<Longrightarrow> inv1 (rbt_baliL l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1983
  by (induct l a b r rule: rbt_baliL.induct) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1984
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1985
lemma inv1_rbt_baliR: "inv1 l \<Longrightarrow> inv1l r \<Longrightarrow> inv1 (rbt_baliR l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1986
  by (induct l a b r rule: rbt_baliR.induct) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1987
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1988
lemma rbt_bheight_rbt_baliL: "bheight l = bheight r \<Longrightarrow> bheight (rbt_baliL l a b r) = Suc (bheight l)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1989
  by (induct l a b r rule: rbt_baliL.induct) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1990
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1991
lemma rbt_bheight_rbt_baliR: "bheight l = bheight r \<Longrightarrow> bheight (rbt_baliR l a b r) = Suc (bheight l)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1992
  by (induct l a b r rule: rbt_baliR.induct) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1993
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1994
lemma inv2_rbt_baliL: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l = bheight r \<Longrightarrow> inv2 (rbt_baliL l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1995
  by (induct l a b r rule: rbt_baliL.induct) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1996
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1997
lemma inv2_rbt_baliR: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l = bheight r \<Longrightarrow> inv2 (rbt_baliR l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1998
  by (induct l a b r rule: rbt_baliR.induct) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  1999
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2000
lemma inv_rbt_baliR: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> inv1 l \<Longrightarrow> inv1l r \<Longrightarrow> bheight l = bheight r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2001
  inv1 (rbt_baliR l a b r) \<and> inv2 (rbt_baliR l a b r) \<and> bheight (rbt_baliR l a b r) = Suc (bheight l)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2002
  by (induct l a b r rule: rbt_baliR.induct) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2003
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2004
lemma inv_rbt_baliL: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> inv1l l \<Longrightarrow> inv1 r \<Longrightarrow> bheight l = bheight r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2005
  inv1 (rbt_baliL l a b r) \<and> inv2 (rbt_baliL l a b r) \<and> bheight (rbt_baliL l a b r) = Suc (bheight l)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2006
  by (induct l a b r rule: rbt_baliL.induct) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2007
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2008
lemma inv2_rbt_baldL_inv1: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l + 1 = bheight r \<Longrightarrow> inv1 r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2009
  inv2 (rbt_baldL l a b r) \<and> bheight (rbt_baldL l a b r) = bheight r"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2010
  by (induct l a b r rule: rbt_baldL.induct) (auto simp: inv2_rbt_baliR rbt_bheight_rbt_baliR)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2011
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2012
lemma inv2_rbt_baldL_B: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l + 1 = bheight r \<Longrightarrow> color_of r = RBT_Impl.B \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2013
  inv2 (rbt_baldL l a b r) \<and> bheight (rbt_baldL l a b r) = bheight r"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2014
  by (induct l a b r rule: rbt_baldL.induct) (auto simp add: inv2_rbt_baliR rbt_bheight_rbt_baliR)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2015
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2016
lemma inv1_rbt_baldL: "inv1l l \<Longrightarrow> inv1 r \<Longrightarrow> color_of r = RBT_Impl.B \<Longrightarrow> inv1 (rbt_baldL l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2017
  by (induct l a b r rule: rbt_baldL.induct) (simp_all add: inv1_rbt_baliR)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2018
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2019
lemma inv1lI: "inv1 t \<Longrightarrow> inv1l t"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2020
  by (cases t) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2021
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2022
lemma neq_Black[simp]: "(c \<noteq> RBT_Impl.B) = (c = RBT_Impl.R)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2023
  by (cases c) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2024
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2025
lemma inv1l_rbt_baldL: "inv1l l \<Longrightarrow> inv1 r \<Longrightarrow> inv1l (rbt_baldL l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2026
  by (induct l a b r rule: rbt_baldL.induct) (auto simp: inv1_rbt_baliR paint2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2027
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2028
lemma inv2_rbt_baldR_inv1: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l = bheight r + 1 \<Longrightarrow> inv1 l \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2029
  inv2 (rbt_baldR l a b r) \<and> bheight (rbt_baldR l a b r) = bheight l"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2030
  by (induct l a b r rule: rbt_baldR.induct) (auto simp: inv2_rbt_baliL rbt_bheight_rbt_baliL)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2031
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2032
lemma inv1_rbt_baldR: "inv1 l \<Longrightarrow> inv1l r \<Longrightarrow> color_of l = RBT_Impl.B \<Longrightarrow> inv1 (rbt_baldR l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2033
  by (induct l a b r rule: rbt_baldR.induct) (simp_all add: inv1_rbt_baliL)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2034
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2035
lemma inv1l_rbt_baldR: "inv1 l \<Longrightarrow> inv1l r \<Longrightarrow>inv1l (rbt_baldR l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2036
  by (induct l a b r rule: rbt_baldR.induct) (auto simp: inv1_rbt_baliL paint2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2037
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2038
lemma inv2_rbt_app: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l = bheight r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2039
  inv2 (rbt_app l r) \<and> bheight (rbt_app l r) = bheight l"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2040
  by (induct l r rule: rbt_app.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2041
     (auto simp: inv2_rbt_baldL_B split: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2042
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2043
lemma inv1_rbt_app: "inv1 l \<Longrightarrow> inv1 r \<Longrightarrow> (color_of l = RBT_Impl.B \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2044
  color_of r = RBT_Impl.B \<longrightarrow> inv1 (rbt_app l r)) \<and> inv1l (rbt_app l r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2045
  by (induct l r rule: rbt_app.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2046
     (auto simp: inv1_rbt_baldL split: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2047
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2048
lemma inv_rbt_baldL: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l + 1 = bheight r \<Longrightarrow> inv1l l \<Longrightarrow> inv1 r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2049
  inv2 (rbt_baldL l a b r) \<and> bheight (rbt_baldL l a b r) = bheight r \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2050
  inv1l (rbt_baldL l a b r) \<and> (color_of r = RBT_Impl.B \<longrightarrow> inv1 (rbt_baldL l a b r))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2051
  by (induct l a b r rule: rbt_baldL.induct) (auto simp: inv_rbt_baliR rbt_bheight_rbt_baliR paint2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2052
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2053
lemma inv_rbt_baldR: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l = bheight r + 1 \<Longrightarrow> inv1 l \<Longrightarrow> inv1l r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2054
  inv2 (rbt_baldR l a b r) \<and> bheight (rbt_baldR l a b r) = bheight l \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2055
  inv1l (rbt_baldR l a b r) \<and> (color_of l = RBT_Impl.B \<longrightarrow> inv1 (rbt_baldR l a b r))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2056
  by (induct l a b r rule: rbt_baldR.induct) (auto simp: inv_rbt_baliL rbt_bheight_rbt_baliL paint2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2057
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2058
lemma inv_rbt_app: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l = bheight r \<Longrightarrow> inv1 l \<Longrightarrow> inv1 r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2059
  inv2 (rbt_app l r) \<and> bheight (rbt_app l r) = bheight l \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2060
  inv1l (rbt_app l r) \<and> (color_of l = RBT_Impl.B \<and> color_of r = RBT_Impl.B \<longrightarrow> inv1 (rbt_app l r))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2061
  by (induct l r rule: rbt_app.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2062
     (auto simp: inv2_rbt_baldL_B inv_rbt_baldL split: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2063
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2064
lemma inv1l_rbt_joinL: "inv1 l \<Longrightarrow> inv1 r \<Longrightarrow> bheight l \<le> bheight r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2065
  inv1l (rbt_joinL l a b r) \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2066
  (bheight l \<noteq> bheight r \<and> color_of r = RBT_Impl.B \<longrightarrow> inv1 (rbt_joinL l a b r))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2067
proof (induct l a b r rule: rbt_joinL.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2068
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2069
  then show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2070
    by (auto simp: inv1_rbt_baliL rbt_joinL.simps[of l a b r]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2071
        split!: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2072
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2073
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2074
lemma inv1l_rbt_joinR: "inv1 l \<Longrightarrow> inv2 l \<Longrightarrow> inv1 r \<Longrightarrow> inv2 r \<Longrightarrow> bheight l \<ge> bheight r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2075
  inv1l (rbt_joinR l a b r) \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2076
  (bheight l \<noteq> bheight r \<and> color_of l = RBT_Impl.B \<longrightarrow> inv1 (rbt_joinR l a b r))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2077
proof (induct l a b r rule: rbt_joinR.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2078
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2079
  then show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2080
    by (fastforce simp: inv1_rbt_baliR rbt_joinR.simps[of l a b r]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2081
        split!: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2082
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2083
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2084
lemma bheight_rbt_joinL: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l \<le> bheight r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2085
  bheight (rbt_joinL l a b r) = bheight r"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2086
proof (induct l a b r rule: rbt_joinL.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2087
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2088
  then show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2089
    by (auto simp: rbt_bheight_rbt_baliL rbt_joinL.simps[of l a b r]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2090
        split!: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2091
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2092
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2093
lemma inv2_rbt_joinL: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l \<le> bheight r \<Longrightarrow> inv2 (rbt_joinL l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2094
proof (induct l a b r rule: rbt_joinL.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2095
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2096
  then show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2097
    by (auto simp: inv2_rbt_baliL bheight_rbt_joinL rbt_joinL.simps[of l a b r]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2098
        split!: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2099
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2100
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2101
lemma bheight_rbt_joinR: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l \<ge> bheight r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2102
  bheight (rbt_joinR l a b r) = bheight l"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2103
proof (induct l a b r rule: rbt_joinR.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2104
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2105
  then show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2106
    by (fastforce simp: rbt_bheight_rbt_baliR rbt_joinR.simps[of l a b r]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2107
        split!: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2108
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2109
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2110
lemma inv2_rbt_joinR: "inv2 l \<Longrightarrow> inv2 r \<Longrightarrow> bheight l \<ge> bheight r \<Longrightarrow> inv2 (rbt_joinR l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2111
proof (induct l a b r rule: rbt_joinR.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2112
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2113
  then show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2114
    by (fastforce simp: inv2_rbt_baliR bheight_rbt_joinR rbt_joinR.simps[of l a b r]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2115
        split!: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2116
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2117
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2118
lemma keys_paint[simp]: "RBT_Impl.keys (paint c t) = RBT_Impl.keys t"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2119
  by (cases t) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2120
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2121
lemma keys_rbt_baliL: "RBT_Impl.keys (rbt_baliL l a b r) = RBT_Impl.keys l @ a # RBT_Impl.keys r"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2122
  by (cases "(l,a,b,r)" rule: rbt_baliL.cases) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2123
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2124
lemma keys_rbt_baliR: "RBT_Impl.keys (rbt_baliR l a b r) = RBT_Impl.keys l @ a # RBT_Impl.keys r"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2125
  by (cases "(l,a,b,r)" rule: rbt_baliR.cases) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2126
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2127
lemma keys_rbt_baldL: "RBT_Impl.keys (rbt_baldL l a b r) = RBT_Impl.keys l @ a # RBT_Impl.keys r"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2128
  by (cases "(l,a,b,r)" rule: rbt_baldL.cases) (auto simp: keys_rbt_baliL keys_rbt_baliR)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2129
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2130
lemma keys_rbt_baldR: "RBT_Impl.keys (rbt_baldR l a b r) = RBT_Impl.keys l @ a # RBT_Impl.keys r"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2131
  by (cases "(l,a,b,r)" rule: rbt_baldR.cases) (auto simp: keys_rbt_baliL keys_rbt_baliR)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2132
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2133
lemma keys_rbt_app: "RBT_Impl.keys (rbt_app l r) = RBT_Impl.keys l @ RBT_Impl.keys r"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2134
  by (induction l r rule: rbt_app.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2135
     (auto simp: keys_rbt_baldL keys_rbt_baldR split: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2136
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2137
lemma keys_rbt_joinL: "bheight l \<le> bheight r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2138
  RBT_Impl.keys (rbt_joinL l a b r) = RBT_Impl.keys l @ a # RBT_Impl.keys r"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2139
proof (induction l a b r rule: rbt_joinL.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2140
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2141
  thus ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2142
    by (auto simp: keys_rbt_baliL rbt_joinL.simps[of l a b r]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2143
        split!: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2144
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2145
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2146
lemma keys_rbt_joinR: "RBT_Impl.keys (rbt_joinR l a b r) = RBT_Impl.keys l @ a # RBT_Impl.keys r"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2147
proof (induction l a b r rule: rbt_joinR.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2148
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2149
  thus ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2150
    by (force simp: keys_rbt_baliR rbt_joinR.simps[of l a b r]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2151
        split!: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2152
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2153
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2154
lemma rbt_set_rbt_baliL: "set (RBT_Impl.keys (rbt_baliL l a b r)) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2155
  set (RBT_Impl.keys l) \<union> {a} \<union> set (RBT_Impl.keys r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2156
  by (cases "(l,a,b,r)" rule: rbt_baliL.cases) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2157
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2158
lemma set_rbt_joinL: "set (RBT_Impl.keys (rbt_joinL l a b r)) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2159
  set (RBT_Impl.keys l) \<union> {a} \<union> set (RBT_Impl.keys r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2160
proof (induction l a b r rule: rbt_joinL.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2161
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2162
  thus ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2163
    by (auto simp: rbt_set_rbt_baliL rbt_joinL.simps[of l a b r]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2164
        split!: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2165
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2166
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2167
lemma rbt_set_rbt_baliR: "set (RBT_Impl.keys (rbt_baliR l a b r)) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2168
  set (RBT_Impl.keys l) \<union> {a} \<union> set (RBT_Impl.keys r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2169
  by (cases "(l,a,b,r)" rule: rbt_baliR.cases) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2170
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2171
lemma set_rbt_joinR: "set (RBT_Impl.keys (rbt_joinR l a b r)) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2172
  set (RBT_Impl.keys l) \<union> {a} \<union> set (RBT_Impl.keys r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2173
proof (induction l a b r rule: rbt_joinR.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2174
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2175
  thus ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2176
    by (force simp: rbt_set_rbt_baliR rbt_joinR.simps[of l a b r]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2177
        split!: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2178
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2179
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2180
lemma set_keys_paint: "set (RBT_Impl.keys (paint c t)) = set (RBT_Impl.keys t)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2181
  by (cases t) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2182
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2183
lemma set_rbt_join: "set (RBT_Impl.keys (rbt_join l a b r)) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2184
  set (RBT_Impl.keys l) \<union> {a} \<union> set (RBT_Impl.keys r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2185
  by (simp add: set_rbt_joinL set_rbt_joinR set_keys_paint rbt_join_def Let_def)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2186
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2187
lemma inv_rbt_join: "inv_12 l \<Longrightarrow> inv_12 r \<Longrightarrow> inv_12 (rbt_join l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2188
  by (auto simp: rbt_join_def Let_def inv1l_rbt_joinL inv1l_rbt_joinR
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2189
      inv2_rbt_joinL inv2_rbt_joinR inv_12_def)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2190
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2191
fun rbt_recolor :: "('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2192
  "rbt_recolor (Branch RBT_Impl.R t1 k v t2) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2193
    (if color_of t1 = RBT_Impl.B \<and> color_of t2 = RBT_Impl.B then Branch RBT_Impl.B t1 k v t2
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2194
    else Branch RBT_Impl.R t1 k v t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2195
| "rbt_recolor t = t"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2196
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2197
lemma rbt_recolor: "inv_12 t \<Longrightarrow> inv_12 (rbt_recolor t)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2198
  by (induction t rule: rbt_recolor.induct) (auto simp: inv_12_def)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2199
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2200
fun rbt_split_min :: "('a, 'b) rbt \<Rightarrow> 'a \<times> 'b \<times> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2201
  "rbt_split_min RBT_Impl.Empty = undefined"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2202
| "rbt_split_min (RBT_Impl.Branch _ l a b r) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2203
    (if is_rbt_empty l then (a,b,r) else let (a',b',l') = rbt_split_min l in (a',b',rbt_join l' a b r))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2204
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2205
lemma rbt_split_min_set:
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2206
  "rbt_split_min t = (a,b,t') \<Longrightarrow> t \<noteq> RBT_Impl.Empty \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2207
  a \<in> set (RBT_Impl.keys t) \<and> set (RBT_Impl.keys t) = {a} \<union> set (RBT_Impl.keys t')"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2208
  by (induction t arbitrary: t') (auto simp: set_rbt_join split: prod.splits if_splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2209
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2210
lemma rbt_split_min_inv: "rbt_split_min t = (a,b,t') \<Longrightarrow> inv_12 t \<Longrightarrow> t \<noteq> RBT_Impl.Empty \<Longrightarrow> inv_12 t'"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2211
  by (induction t arbitrary: t')
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2212
     (auto simp: inv_rbt_join split: if_splits prod.splits dest: rbt_Node)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2213
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2214
definition rbt_join2 :: "('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2215
  "rbt_join2 l r = (if is_rbt_empty r then l else let (a,b,r') = rbt_split_min r in rbt_join l a b r')"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2216
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2217
lemma set_rbt_join2[simp]: "set (RBT_Impl.keys (rbt_join2 l r)) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2218
  set (RBT_Impl.keys l) \<union> set (RBT_Impl.keys r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2219
  by (simp add: rbt_join2_def rbt_split_min_set set_rbt_join split: prod.split)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2220
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2221
lemma inv_rbt_join2: "inv_12 l \<Longrightarrow> inv_12 r \<Longrightarrow> inv_12 (rbt_join2 l r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2222
  by (simp add: rbt_join2_def inv_rbt_join rbt_split_min_set rbt_split_min_inv split: prod.split)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2223
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2224
context ord begin
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2225
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2226
fun rbt_split :: "('a, 'b) rbt \<Rightarrow> 'a \<Rightarrow> ('a, 'b) rbt \<times> 'b option \<times> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2227
  "rbt_split RBT_Impl.Empty k = (RBT_Impl.Empty, None, RBT_Impl.Empty)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2228
| "rbt_split (RBT_Impl.Branch _ l a b r) x =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2229
  (if x < a then (case rbt_split l x of (l1, \<beta>, l2) \<Rightarrow> (l1, \<beta>, rbt_join l2 a b r))
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2230
  else if a < x then (case rbt_split r x of (r1, \<beta>, r2) \<Rightarrow> (rbt_join l a b r1, \<beta>, r2))
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2231
  else (l, Some b, r))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2232
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2233
lemma rbt_split: "rbt_split t x = (l,\<beta>,r) \<Longrightarrow> inv_12 t \<Longrightarrow> inv_12 l \<and> inv_12 r"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2234
  by (induction t arbitrary: l r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2235
     (auto simp: set_rbt_join inv_rbt_join rbt_greater_prop rbt_less_prop
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2236
      split: if_splits prod.splits dest!: rbt_Node)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2237
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2238
lemma rbt_split_size: "(l2,\<beta>,r2) = rbt_split t2 a \<Longrightarrow> size l2 + size r2 \<le> size t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2239
  by (induction t2 a arbitrary: l2 r2 rule: rbt_split.induct) (auto split: if_splits prod.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2240
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2241
function rbt_union_rec :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2242
  "rbt_union_rec f t1 t2 = (let (f, t2, t1) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2243
    if flip_rbt t2 t1 then (\<lambda>k v v'. f k v' v, t1, t2) else (f, t2, t1) in
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2244
    if small_rbt t2 then RBT_Impl.fold (rbt_insert_with_key f) t2 t1
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2245
    else (case t1 of RBT_Impl.Empty \<Rightarrow> t2
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2246
      | RBT_Impl.Branch _ l1 a b r1 \<Rightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2247
        case rbt_split t2 a of (l2, \<beta>, r2) \<Rightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2248
          rbt_join (rbt_union_rec f l1 l2) a (case \<beta> of None \<Rightarrow> b | Some b' \<Rightarrow> f a b b') (rbt_union_rec f r1 r2)))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2249
  by pat_completeness auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2250
termination
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2251
  using rbt_split_size
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2252
  by (relation "measure (\<lambda>(f,t1,t2). size t1 + size t2)") (fastforce split: if_splits)+
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2253
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2254
declare rbt_union_rec.simps[simp del]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2255
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2256
function rbt_union_swap_rec :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2257
  "rbt_union_swap_rec f \<gamma> t1 t2 = (let (\<gamma>, t2, t1) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2258
    if flip_rbt t2 t1 then (\<not>\<gamma>, t1, t2) else (\<gamma>, t2, t1);
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2259
    f' = (if \<gamma> then (\<lambda>k v v'. f k v' v) else f) in
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2260
    if small_rbt t2 then RBT_Impl.fold (rbt_insert_with_key f') t2 t1
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2261
    else (case t1 of RBT_Impl.Empty \<Rightarrow> t2
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2262
      | RBT_Impl.Branch _ l1 a b r1 \<Rightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2263
        case rbt_split t2 a of (l2, \<beta>, r2) \<Rightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2264
          rbt_join (rbt_union_swap_rec f \<gamma> l1 l2) a (case \<beta> of None \<Rightarrow> b | Some b' \<Rightarrow> f' a b b') (rbt_union_swap_rec f \<gamma> r1 r2)))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2265
  by pat_completeness auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2266
termination
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2267
  using rbt_split_size
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2268
  by (relation "measure (\<lambda>(f,\<gamma>,t1,t2). size t1 + size t2)") (fastforce split: if_splits)+
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2269
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2270
declare rbt_union_swap_rec.simps[simp del]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2271
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2272
lemma rbt_union_swap_rec: "rbt_union_swap_rec f \<gamma> t1 t2 =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2273
  rbt_union_rec (if \<gamma> then (\<lambda>k v v'. f k v' v) else f) t1 t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2274
proof (induction f \<gamma> t1 t2 rule: rbt_union_swap_rec.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2275
  case (1 f \<gamma> t1 t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2276
  show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2277
    using 1[OF refl _ refl refl _ refl _ refl]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2278
    unfolding rbt_union_swap_rec.simps[of _ _ t1] rbt_union_rec.simps[of _ t1]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2279
    by (auto simp: Let_def split: rbt.splits prod.splits option.splits) (* slow *)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2280
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2281
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2282
lemma rbt_fold_rbt_insert:
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2283
  assumes "inv_12 t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2284
  shows "inv_12 (RBT_Impl.fold (rbt_insert_with_key f) t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2285
proof -
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2286
  define xs where "xs = RBT_Impl.entries t1"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2287
  from assms show ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2288
    unfolding RBT_Impl.fold_def xs_def[symmetric]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2289
    by (induct xs rule: rev_induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2290
       (auto simp: inv_12_def rbt_insert_with_key_def ins_inv1_inv2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2291
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2292
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2293
lemma rbt_union_rec: "inv_12 t1 \<Longrightarrow> inv_12 t2 \<Longrightarrow> inv_12 (rbt_union_rec f t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2294
proof (induction f t1 t2 rule: rbt_union_rec.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2295
  case (1 t1 t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2296
  thus ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2297
    by (auto simp: rbt_union_rec.simps[of t1 t2] inv_rbt_join rbt_split rbt_fold_rbt_insert
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2298
        split!: RBT_Impl.rbt.splits RBT_Impl.color.splits prod.split if_splits dest: rbt_Node)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2299
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2300
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2301
definition "map_filter_inter f t1 t2 = List.map_filter (\<lambda>(k, v).
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2302
  case rbt_lookup t1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2303
  | Some v' \<Rightarrow> Some (k, f k v' v)) (RBT_Impl.entries t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2304
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2305
function rbt_inter_rec :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2306
  "rbt_inter_rec f t1 t2 = (let (f, t2, t1) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2307
    if flip_rbt t2 t1 then (\<lambda>k v v'. f k v' v, t1, t2) else (f, t2, t1) in
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2308
    if small_rbt t2 then rbtreeify (map_filter_inter f t1 t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2309
    else case t1 of RBT_Impl.Empty \<Rightarrow> RBT_Impl.Empty
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2310
    | RBT_Impl.Branch _ l1 a b r1 \<Rightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2311
      case rbt_split t2 a of (l2, \<beta>, r2) \<Rightarrow> let l' = rbt_inter_rec f l1 l2; r' = rbt_inter_rec f r1 r2 in
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2312
      (case \<beta> of None \<Rightarrow> rbt_join2 l' r' | Some b' \<Rightarrow> rbt_join l' a (f a b b') r'))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2313
  by pat_completeness auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2314
termination
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2315
  using rbt_split_size
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2316
  by (relation "measure (\<lambda>(f,t1,t2). size t1 + size t2)") (fastforce split: if_splits)+
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2317
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2318
declare rbt_inter_rec.simps[simp del]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2319
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2320
function rbt_inter_swap_rec :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2321
  "rbt_inter_swap_rec f \<gamma> t1 t2 = (let (\<gamma>, t2, t1) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2322
    if flip_rbt t2 t1 then (\<not>\<gamma>, t1, t2) else (\<gamma>, t2, t1);
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2323
    f' = (if \<gamma> then (\<lambda>k v v'. f k v' v) else f) in
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2324
    if small_rbt t2 then rbtreeify (map_filter_inter f' t1 t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2325
    else case t1 of RBT_Impl.Empty \<Rightarrow> RBT_Impl.Empty
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2326
    | RBT_Impl.Branch _ l1 a b r1 \<Rightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2327
      case rbt_split t2 a of (l2, \<beta>, r2) \<Rightarrow> let l' = rbt_inter_swap_rec f \<gamma> l1 l2; r' = rbt_inter_swap_rec f \<gamma> r1 r2 in
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2328
      (case \<beta> of None \<Rightarrow> rbt_join2 l' r' | Some b' \<Rightarrow> rbt_join l' a (f' a b b') r'))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2329
  by pat_completeness auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2330
termination
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2331
  using rbt_split_size
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2332
  by (relation "measure (\<lambda>(f,\<gamma>,t1,t2). size t1 + size t2)") (fastforce split: if_splits)+
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2333
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2334
declare rbt_inter_swap_rec.simps[simp del]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2335
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2336
lemma rbt_inter_swap_rec: "rbt_inter_swap_rec f \<gamma> t1 t2 =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2337
  rbt_inter_rec (if \<gamma> then (\<lambda>k v v'. f k v' v) else f) t1 t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2338
proof (induction f \<gamma> t1 t2 rule: rbt_inter_swap_rec.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2339
  case (1 f \<gamma> t1 t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2340
  show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2341
    using 1[OF refl _ refl refl _ refl _ refl]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2342
    unfolding rbt_inter_swap_rec.simps[of _ _ t1] rbt_inter_rec.simps[of _ t1]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2343
    by (auto simp add: Let_def split: rbt.splits prod.splits option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2344
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2345
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2346
lemma rbt_rbtreeify[simp]: "inv_12 (rbtreeify kvs)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2347
  by (simp add: inv_12_def rbtreeify_def inv1_rbtreeify_g inv2_rbtreeify_g)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2348
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2349
lemma rbt_inter_rec: "inv_12 t1 \<Longrightarrow> inv_12 t2 \<Longrightarrow> inv_12 (rbt_inter_rec f t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2350
proof(induction f t1 t2 rule: rbt_inter_rec.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2351
  case (1 t1 t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2352
  thus ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2353
    by (auto simp: rbt_inter_rec.simps[of t1 t2] inv_rbt_join inv_rbt_join2 rbt_split Let_def
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2354
        split!: RBT_Impl.rbt.splits RBT_Impl.color.splits prod.split if_splits
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2355
        option.splits dest!: rbt_Node)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2356
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2357
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2358
definition "filter_minus t1 t2 = filter (\<lambda>(k, _). rbt_lookup t2 k = None) (RBT_Impl.entries t1)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2359
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2360
fun rbt_minus_rec :: "('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2361
  "rbt_minus_rec t1 t2 = (if small_rbt t2 then RBT_Impl.fold (\<lambda>k _ t. rbt_delete k t) t2 t1
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2362
    else if small_rbt t1 then rbtreeify (filter_minus t1 t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2363
    else case t2 of RBT_Impl.Empty \<Rightarrow> t1
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2364
      | RBT_Impl.Branch _ l2 a b r2 \<Rightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2365
        case rbt_split t1 a of (l1, _, r1) \<Rightarrow> rbt_join2 (rbt_minus_rec l1 l2) (rbt_minus_rec r1 r2))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2366
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2367
declare rbt_minus_rec.simps[simp del]
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2368
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2369
end
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2370
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2371
context linorder begin
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2372
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2373
lemma rbt_sorted_entries_right_unique:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2374
  "\<lbrakk> (k, v) \<in> set (entries t); (k, v') \<in> set (entries t); 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2375
     rbt_sorted t \<rbrakk> \<Longrightarrow> v = v'"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2376
by(auto dest!: distinct_entries inj_onD[where x="(k, v)" and y="(k, v')"] simp add: distinct_map)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2377
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2378
lemma rbt_sorted_fold_rbt_insertwk:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2379
  "rbt_sorted t \<Longrightarrow> rbt_sorted (List.fold (\<lambda>(k, v). rbt_insert_with_key f k v) xs t)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2380
by(induct xs rule: rev_induct)(auto simp add: rbt_insertwk_rbt_sorted)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2381
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2382
lemma is_rbt_fold_rbt_insertwk:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2383
  assumes "is_rbt t1"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2384
  shows "is_rbt (fold (rbt_insert_with_key f) t2 t1)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2385
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
  2386
  define xs where "xs = entries t2"
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2387
  from assms show ?thesis unfolding fold_def xs_def[symmetric]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2388
    by(induct xs rule: rev_induct)(auto simp add: rbt_insertwk_is_rbt)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2389
qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2390
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2391
lemma rbt_delete: "inv_12 t \<Longrightarrow> inv_12 (rbt_delete x t)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2392
  using rbt_del_inv1_inv2[of t x]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2393
  by (auto simp: inv_12_def rbt_delete_def rbt_del_inv1_inv2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2394
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2395
lemma rbt_sorted_delete: "rbt_sorted t \<Longrightarrow> rbt_sorted (rbt_delete x t)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2396
  by (auto simp: rbt_delete_def rbt_del_rbt_sorted)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2397
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2398
lemma rbt_fold_rbt_delete:
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2399
  assumes "inv_12 t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2400
  shows "inv_12 (RBT_Impl.fold (\<lambda>k _ t. rbt_delete k t) t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2401
proof -
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2402
  define xs where "xs = RBT_Impl.entries t1"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2403
  from assms show ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2404
    unfolding RBT_Impl.fold_def xs_def[symmetric]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2405
    by (induct xs rule: rev_induct) (auto simp: rbt_delete)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2406
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2407
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2408
lemma rbt_minus_rec: "inv_12 t1 \<Longrightarrow> inv_12 t2 \<Longrightarrow> inv_12 (rbt_minus_rec t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2409
proof(induction t1 t2 rule: rbt_minus_rec.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2410
  case (1 t1 t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2411
  thus ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2412
    by (auto simp: rbt_minus_rec.simps[of t1 t2] inv_rbt_join inv_rbt_join2 rbt_split
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2413
        rbt_fold_rbt_delete split!: RBT_Impl.rbt.splits RBT_Impl.color.splits prod.split if_splits
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2414
        dest: rbt_Node)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2415
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2416
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2417
end
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2418
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2419
context linorder begin
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2420
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2421
lemma rbt_sorted_rbt_baliL: "rbt_sorted l \<Longrightarrow> rbt_sorted r \<Longrightarrow> l |\<guillemotleft> a \<Longrightarrow> a \<guillemotleft>| r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2422
  rbt_sorted (rbt_baliL l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2423
  using rbt_greater_trans rbt_less_trans
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2424
  by (cases "(l,a,b,r)" rule: rbt_baliL.cases) fastforce+
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2425
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2426
lemma rbt_lookup_rbt_baliL: "rbt_sorted l \<Longrightarrow> rbt_sorted r \<Longrightarrow> l |\<guillemotleft> a \<Longrightarrow> a \<guillemotleft>| r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2427
  rbt_lookup (rbt_baliL l a b r) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2428
  (if k < a then rbt_lookup l k else if k = a then Some b else rbt_lookup r k)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2429
  by (cases "(l,a,b,r)" rule: rbt_baliL.cases) (auto split!: if_splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2430
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2431
lemma rbt_sorted_rbt_baliR: "rbt_sorted l \<Longrightarrow> rbt_sorted r \<Longrightarrow> l |\<guillemotleft> a \<Longrightarrow> a \<guillemotleft>| r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2432
  rbt_sorted (rbt_baliR l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2433
  using rbt_greater_trans rbt_less_trans
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2434
  by (cases "(l,a,b,r)" rule: rbt_baliR.cases) fastforce+
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2435
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2436
lemma rbt_lookup_rbt_baliR: "rbt_sorted l \<Longrightarrow> rbt_sorted r \<Longrightarrow> l |\<guillemotleft> a \<Longrightarrow> a \<guillemotleft>| r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2437
  rbt_lookup (rbt_baliR l a b r) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2438
  (if k < a then rbt_lookup l k else if k = a then Some b else rbt_lookup r k)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2439
  by (cases "(l,a,b,r)" rule: rbt_baliR.cases) (auto split!: if_splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2440
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2441
lemma rbt_sorted_rbt_joinL: "rbt_sorted (RBT_Impl.Branch c l a b r) \<Longrightarrow> bheight l \<le> bheight r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2442
  rbt_sorted (rbt_joinL l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2443
proof (induction l a b r arbitrary: c rule: rbt_joinL.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2444
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2445
  thus ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2446
    by (auto simp: rbt_set_rbt_baliL rbt_joinL.simps[of l a b r] set_rbt_joinL rbt_less_prop
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2447
        intro!: rbt_sorted_rbt_baliL split!: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2448
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2449
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2450
lemma rbt_lookup_rbt_joinL: "rbt_sorted l \<Longrightarrow> rbt_sorted r \<Longrightarrow> l |\<guillemotleft> a \<Longrightarrow> a \<guillemotleft>| r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2451
  rbt_lookup (rbt_joinL l a b r) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2452
  (if k < a then rbt_lookup l k else if k = a then Some b else rbt_lookup r k)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2453
proof (induction l a b r rule: rbt_joinL.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2454
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2455
  have less_rbt_joinL:
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2456
    "rbt_sorted r1 \<Longrightarrow> r1 |\<guillemotleft> x \<Longrightarrow> a \<guillemotleft>| r1 \<Longrightarrow> a < x \<Longrightarrow> rbt_joinL l a b r1 |\<guillemotleft> x" for x r1
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2457
    using 1(5)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2458
    by (auto simp: rbt_less_prop rbt_greater_prop set_rbt_joinL)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2459
  show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2460
    using 1 less_rbt_joinL rbt_lookup_rbt_baliL[OF rbt_sorted_rbt_joinL[of _ l a b], where ?k=k]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2461
    by (auto simp: rbt_joinL.simps[of l a b r] split!: if_splits rbt.splits color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2462
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2463
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2464
lemma rbt_sorted_rbt_joinR: "rbt_sorted l \<Longrightarrow> rbt_sorted r \<Longrightarrow> l |\<guillemotleft> a \<Longrightarrow> a \<guillemotleft>| r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2465
  rbt_sorted (rbt_joinR l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2466
proof (induction l a b r rule: rbt_joinR.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2467
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2468
  thus ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2469
    by (auto simp: rbt_set_rbt_baliR rbt_joinR.simps[of l a b r] set_rbt_joinR rbt_greater_prop
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2470
        intro!: rbt_sorted_rbt_baliR split!: RBT_Impl.rbt.splits RBT_Impl.color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2471
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2472
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2473
lemma rbt_lookup_rbt_joinR: "rbt_sorted l \<Longrightarrow> rbt_sorted r \<Longrightarrow> l |\<guillemotleft> a \<Longrightarrow> a \<guillemotleft>| r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2474
  rbt_lookup (rbt_joinR l a b r) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2475
  (if k < a then rbt_lookup l k else if k = a then Some b else rbt_lookup r k)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2476
proof (induction l a b r rule: rbt_joinR.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2477
  case (1 l a b r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2478
  have less_rbt_joinR:
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2479
    "rbt_sorted l1 \<Longrightarrow> x \<guillemotleft>| l1 \<Longrightarrow> l1 |\<guillemotleft> a \<Longrightarrow> x < a \<Longrightarrow> x \<guillemotleft>| rbt_joinR l1 a b r" for x l1
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2480
    using 1(6)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2481
    by (auto simp: rbt_less_prop rbt_greater_prop set_rbt_joinR)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2482
  show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2483
    using 1 less_rbt_joinR rbt_lookup_rbt_baliR[OF _ rbt_sorted_rbt_joinR[of _ r a b], where ?k=k]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2484
    by (auto simp: rbt_joinR.simps[of l a b r] split!: if_splits rbt.splits color.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2485
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2486
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2487
lemma rbt_sorted_paint: "rbt_sorted (paint c t) = rbt_sorted t"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2488
  by (cases t) auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2489
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2490
lemma rbt_sorted_rbt_join: "rbt_sorted (RBT_Impl.Branch c l a b r) \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2491
  rbt_sorted (rbt_join l a b r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2492
  by (auto simp: rbt_sorted_paint rbt_sorted_rbt_joinL rbt_sorted_rbt_joinR rbt_join_def Let_def)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2493
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2494
lemma rbt_lookup_rbt_join: "rbt_sorted l \<Longrightarrow> rbt_sorted r \<Longrightarrow> l |\<guillemotleft> a \<Longrightarrow> a \<guillemotleft>| r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2495
  rbt_lookup (rbt_join l a b r) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2496
  (if k < a then rbt_lookup l k else if k = a then Some b else rbt_lookup r k)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2497
  by (auto simp: rbt_join_def Let_def rbt_lookup_rbt_joinL rbt_lookup_rbt_joinR)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2498
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2499
lemma rbt_split_min_rbt_sorted: "rbt_split_min t = (a,b,t') \<Longrightarrow> rbt_sorted t \<Longrightarrow> t \<noteq> RBT_Impl.Empty \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2500
  rbt_sorted t' \<and> (\<forall>x \<in> set (RBT_Impl.keys t'). a < x)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2501
  by (induction t arbitrary: t')
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2502
     (fastforce simp: rbt_split_min_set rbt_sorted_rbt_join set_rbt_join rbt_less_prop rbt_greater_prop
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2503
      split: if_splits prod.splits)+
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2504
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2505
lemma rbt_split_min_rbt_lookup: "rbt_split_min t = (a,b,t') \<Longrightarrow> rbt_sorted t \<Longrightarrow> t \<noteq> RBT_Impl.Empty \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2506
  rbt_lookup t k = (if k < a then None else if k = a then Some b else rbt_lookup t' k)"
73526
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 73212
diff changeset
  2507
  apply (induction t arbitrary: a b t')
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 73212
diff changeset
  2508
   apply(simp_all split: if_splits prod.splits)
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 73212
diff changeset
  2509
     apply(auto simp: rbt_less_prop rbt_split_min_set rbt_lookup_rbt_join rbt_split_min_rbt_sorted)
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 73212
diff changeset
  2510
  done
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2511
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2512
lemma rbt_sorted_rbt_join2: "rbt_sorted l \<Longrightarrow> rbt_sorted r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2513
  \<forall>x \<in> set (RBT_Impl.keys l). \<forall>y \<in> set (RBT_Impl.keys r). x < y \<Longrightarrow> rbt_sorted (rbt_join2 l r)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2514
  by (simp add: rbt_join2_def rbt_sorted_rbt_join rbt_split_min_set rbt_split_min_rbt_sorted set_rbt_join
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2515
      rbt_greater_prop rbt_less_prop split: prod.split)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2516
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2517
lemma rbt_lookup_rbt_join2: "rbt_sorted l \<Longrightarrow> rbt_sorted r \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2518
  \<forall>x \<in> set (RBT_Impl.keys l). \<forall>y \<in> set (RBT_Impl.keys r). x < y \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2519
  rbt_lookup (rbt_join2 l r) k = (case rbt_lookup l k of None \<Rightarrow> rbt_lookup r k | Some v \<Rightarrow> Some v)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2520
  using rbt_lookup_keys
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2521
  by (fastforce simp: rbt_join2_def rbt_greater_prop rbt_less_prop rbt_lookup_rbt_join
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2522
      rbt_split_min_rbt_lookup rbt_split_min_rbt_sorted rbt_split_min_set split: option.splits prod.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2523
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2524
lemma rbt_split_props: "rbt_split t x = (l,\<beta>,r) \<Longrightarrow> rbt_sorted t \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2525
  set (RBT_Impl.keys l) = {a \<in> set (RBT_Impl.keys t). a < x} \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2526
  set (RBT_Impl.keys r) = {a \<in> set (RBT_Impl.keys t). x < a} \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2527
  rbt_sorted l \<and> rbt_sorted r"
73526
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 73212
diff changeset
  2528
  apply (induction t arbitrary: l r)
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 73212
diff changeset
  2529
   apply(simp_all split!: prod.splits if_splits)
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 73212
diff changeset
  2530
    apply(force simp: set_rbt_join rbt_greater_prop rbt_less_prop
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 73212
diff changeset
  2531
      intro: rbt_sorted_rbt_join)+
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 73212
diff changeset
  2532
  done
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2533
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2534
lemma rbt_split_lookup: "rbt_split t x = (l,\<beta>,r) \<Longrightarrow> rbt_sorted t \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2535
  rbt_lookup t k = (if k < x then rbt_lookup l k else if k = x then \<beta> else rbt_lookup r k)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2536
proof (induction t arbitrary: x l \<beta> r)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2537
  case (Branch c t1 a b t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2538
  have "rbt_sorted r1" "r1 |\<guillemotleft> a" if "rbt_split t1 x = (l, \<beta>, r1)" for r1
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2539
    using rbt_split_props Branch(4) that
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2540
    by (fastforce simp: rbt_less_prop)+
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2541
  moreover have "rbt_sorted l1" "a \<guillemotleft>| l1" if "rbt_split t2 x = (l1, \<beta>, r)" for l1
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2542
    using rbt_split_props Branch(4) that
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2543
    by (fastforce simp: rbt_greater_prop)+
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2544
  ultimately show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2545
    using Branch rbt_lookup_rbt_join[of t1 _ a b k] rbt_lookup_rbt_join[of _ t2 a b k]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2546
    by (auto split!: if_splits prod.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2547
qed simp
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2548
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2549
lemma rbt_sorted_fold_insertwk: "rbt_sorted t \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2550
  rbt_sorted (RBT_Impl.fold (rbt_insert_with_key f) t' t)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2551
  by (induct t' arbitrary: t)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2552
     (simp_all add: rbt_insertwk_rbt_sorted)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2553
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2554
lemma rbt_lookup_iff_keys:
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2555
  "rbt_sorted t \<Longrightarrow> set (RBT_Impl.keys t) = {k. \<exists>v. rbt_lookup t k = Some v}"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2556
  "rbt_sorted t \<Longrightarrow> rbt_lookup t k = None \<longleftrightarrow> k \<notin> set (RBT_Impl.keys t)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2557
  "rbt_sorted t \<Longrightarrow> (\<exists>v. rbt_lookup t k = Some v) \<longleftrightarrow> k \<in> set (RBT_Impl.keys t)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2558
  using entry_in_tree_keys rbt_lookup_keys[of t]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2559
  by force+
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2560
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2561
lemma rbt_lookup_fold_rbt_insertwk:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2562
  assumes t1: "rbt_sorted t1" and t2: "rbt_sorted t2"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2563
  shows "rbt_lookup (fold (rbt_insert_with_key f) t1 t2) k =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2564
  (case rbt_lookup t1 k of None \<Rightarrow> rbt_lookup t2 k
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2565
   | Some v \<Rightarrow> case rbt_lookup t2 k of None \<Rightarrow> Some v
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2566
               | Some w \<Rightarrow> Some (f k w v))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2567
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
  2568
  define xs where "xs = entries t1"
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2569
  hence dt1: "distinct (map fst xs)" using t1 by(simp add: distinct_entries)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2570
  with t2 show ?thesis
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2571
    unfolding fold_def map_of_entries[OF t1, symmetric]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2572
      xs_def[symmetric] distinct_map_of_rev[OF dt1, symmetric]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2573
    apply(induct xs rule: rev_induct)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2574
    apply(auto simp add: rbt_lookup_rbt_insertwk rbt_sorted_fold_rbt_insertwk split: option.splits)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2575
    apply(auto simp add: distinct_map_of_rev intro: rev_image_eqI)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2576
    done
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2577
qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2578
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2579
lemma rbt_lookup_union_rec: "rbt_sorted t1 \<Longrightarrow> rbt_sorted t2 \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2580
  rbt_sorted (rbt_union_rec f t1 t2) \<and> rbt_lookup (rbt_union_rec f t1 t2) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2581
  (case rbt_lookup t1 k of None \<Rightarrow> rbt_lookup t2 k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2582
  | Some v \<Rightarrow> (case rbt_lookup t2 k of None \<Rightarrow> Some v
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2583
              | Some w \<Rightarrow> Some (f k v w)))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2584
proof(induction f t1 t2 arbitrary: k rule: rbt_union_rec.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2585
  case (1 f t1 t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2586
  obtain f' t1' t2' where flip: "(f', t2', t1') =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2587
    (if flip_rbt t2 t1 then (\<lambda>k v v'. f k v' v, t1, t2) else (f, t2, t1))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2588
    by fastforce
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2589
  have rbt_sorted': "rbt_sorted t1'" "rbt_sorted t2'"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2590
    using 1(3,4) flip
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2591
    by (auto split: if_splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2592
  show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2593
  proof (cases t1')
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2594
    case Empty
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2595
    show ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2596
      unfolding rbt_union_rec.simps[of _ t1] flip[symmetric]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2597
      using flip rbt_sorted' rbt_split_props[of t2]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2598
      by (auto simp: Empty rbt_lookup_fold_rbt_insertwk
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2599
          intro!: rbt_sorted_fold_insertwk split: if_splits option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2600
  next
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2601
    case (Branch c l1 a b r1)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2602
    {
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2603
      assume not_small: "\<not>small_rbt t2'"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2604
      obtain l2 \<beta> r2 where rbt_split_t2': "rbt_split t2' a = (l2, \<beta>, r2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2605
        by (cases "rbt_split t2' a") auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2606
      have rbt_sort: "rbt_sorted l1" "rbt_sorted r1"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2607
        using 1(3,4) flip
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2608
        by (auto simp: Branch split: if_splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2609
      note rbt_split_t2'_props = rbt_split_props[OF rbt_split_t2' rbt_sorted'(2)]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2610
      have union_l1_l2: "rbt_sorted (rbt_union_rec f' l1 l2)" "rbt_lookup (rbt_union_rec f' l1 l2) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2611
        (case rbt_lookup l1 k of None \<Rightarrow> rbt_lookup l2 k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2612
        | Some v \<Rightarrow> (case rbt_lookup l2 k of None \<Rightarrow> Some v | Some w \<Rightarrow> Some (f' k v w)))" for k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2613
        using 1(1)[OF flip refl refl _ Branch rbt_split_t2'[symmetric]] rbt_sort rbt_split_t2'_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2614
        by (auto simp: not_small)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2615
      have union_r1_r2: "rbt_sorted (rbt_union_rec f' r1 r2)" "rbt_lookup (rbt_union_rec f' r1 r2) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2616
        (case rbt_lookup r1 k of None \<Rightarrow> rbt_lookup r2 k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2617
        | Some v \<Rightarrow> (case rbt_lookup r2 k of None \<Rightarrow> Some v | Some w \<Rightarrow> Some (f' k v w)))" for k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2618
        using 1(2)[OF flip refl refl _ Branch rbt_split_t2'[symmetric]] rbt_sort rbt_split_t2'_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2619
        by (auto simp: not_small)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2620
      have union_l1_l2_keys: "set (RBT_Impl.keys (rbt_union_rec f' l1 l2)) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2621
       set (RBT_Impl.keys l1) \<union> set (RBT_Impl.keys l2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2622
        using rbt_sorted'(1) rbt_split_t2'_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2623
        by (auto simp: Branch rbt_lookup_iff_keys(1) union_l1_l2 split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2624
      have union_r1_r2_keys: "set (RBT_Impl.keys (rbt_union_rec f' r1 r2)) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2625
        set (RBT_Impl.keys r1) \<union> set (RBT_Impl.keys r2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2626
        using rbt_sorted'(1) rbt_split_t2'_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2627
        by (auto simp: Branch rbt_lookup_iff_keys(1) union_r1_r2 split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2628
      have union_l1_l2_less: "rbt_union_rec f' l1 l2 |\<guillemotleft> a"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2629
        using rbt_sorted'(1) rbt_split_t2'_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2630
        by (auto simp: Branch rbt_less_prop union_l1_l2_keys)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2631
      have union_r1_r2_greater: "a \<guillemotleft>| rbt_union_rec f' r1 r2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2632
        using rbt_sorted'(1) rbt_split_t2'_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2633
        by (auto simp: Branch rbt_greater_prop union_r1_r2_keys)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2634
      have "rbt_lookup (rbt_union_rec f t1 t2) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2635
       (case rbt_lookup t1' k of None \<Rightarrow> rbt_lookup t2' k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2636
       | Some v \<Rightarrow> (case rbt_lookup t2' k of None \<Rightarrow> Some v | Some w \<Rightarrow> Some (f' k v w)))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2637
        using rbt_sorted' union_l1_l2 union_r1_r2 rbt_split_t2'_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2638
          union_l1_l2_less union_r1_r2_greater not_small
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2639
        by (auto simp: rbt_union_rec.simps[of _ t1] flip[symmetric] Branch
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2640
            rbt_split_t2' rbt_lookup_rbt_join rbt_split_lookup[OF rbt_split_t2'] split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2641
      moreover have "rbt_sorted (rbt_union_rec f t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2642
        using rbt_sorted' rbt_split_t2'_props not_small
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2643
        by (auto simp: rbt_union_rec.simps[of _ t1] flip[symmetric] Branch rbt_split_t2'
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2644
            union_l1_l2 union_r1_r2 union_l1_l2_keys union_r1_r2_keys rbt_less_prop
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2645
            rbt_greater_prop intro!: rbt_sorted_rbt_join)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2646
      ultimately have ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2647
        using flip
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2648
        by (auto split: if_splits option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2649
    }
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2650
    then show ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2651
      unfolding rbt_union_rec.simps[of _ t1] flip[symmetric]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2652
      using rbt_sorted' flip
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2653
      by (auto simp: rbt_sorted_fold_insertwk rbt_lookup_fold_rbt_insertwk split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2654
  qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2655
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2656
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2657
lemma rbtreeify_map_filter_inter:
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2658
  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2659
  assumes "rbt_sorted t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2660
  shows "rbt_sorted (rbtreeify (map_filter_inter f t1 t2))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2661
    "rbt_lookup (rbtreeify (map_filter_inter f t1 t2)) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2662
    (case rbt_lookup t1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2663
    | Some v \<Rightarrow> (case rbt_lookup t2 k of None \<Rightarrow> None | Some w \<Rightarrow> Some (f k v w)))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2664
proof -
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2665
  have map_of_map_filter: "map_of (List.map_filter (\<lambda>(k, v).
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2666
    case rbt_lookup t1 k of None \<Rightarrow> None | Some v' \<Rightarrow> Some (k, f k v' v)) xs) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2667
    (case rbt_lookup t1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2668
    | Some v \<Rightarrow> (case map_of xs k of None \<Rightarrow> None | Some w \<Rightarrow> Some (f k v w)))" for xs k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2669
    by (induction xs) (auto simp: List.map_filter_def split: option.splits) (* slow *)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2670
  have map_fst_map_filter: "map fst (List.map_filter (\<lambda>(k, v).
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2671
    case rbt_lookup t1 k of None \<Rightarrow> None | Some v' \<Rightarrow> Some (k, f k v' v)) xs) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2672
    filter (\<lambda>k. rbt_lookup t1 k \<noteq> None) (map fst xs)" for xs
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2673
    by (induction xs) (auto simp: List.map_filter_def split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2674
  have "sorted (map fst (map_filter_inter f t1 t2))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2675
    using sorted_filter[of id] rbt_sorted_entries[OF assms]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2676
    by (auto simp: map_filter_inter_def map_fst_map_filter)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2677
  moreover have "distinct (map fst (map_filter_inter f t1 t2))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2678
    using distinct_filter distinct_entries[OF assms]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2679
    by (auto simp: map_filter_inter_def map_fst_map_filter)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2680
  ultimately show
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2681
    "rbt_sorted (rbtreeify (map_filter_inter f t1 t2))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2682
    "rbt_lookup (rbtreeify (map_filter_inter f t1 t2)) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2683
      (case rbt_lookup t1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2684
      | Some v \<Rightarrow> (case rbt_lookup t2 k of None \<Rightarrow> None | Some w \<Rightarrow> Some (f k v w)))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2685
    using rbt_sorted_rbtreeify
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2686
    by (auto simp: rbt_lookup_rbtreeify map_filter_inter_def map_of_map_filter
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2687
        map_of_entries[OF assms] split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2688
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2689
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2690
lemma rbt_lookup_inter_rec: "rbt_sorted t1 \<Longrightarrow> rbt_sorted t2 \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2691
  rbt_sorted (rbt_inter_rec f t1 t2) \<and> rbt_lookup (rbt_inter_rec f t1 t2) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2692
  (case rbt_lookup t1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2693
  | Some v \<Rightarrow> (case rbt_lookup t2 k of None \<Rightarrow> None | Some w \<Rightarrow> Some (f k v w)))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2694
proof(induction f t1 t2 arbitrary: k rule: rbt_inter_rec.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2695
  case (1 f t1 t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2696
  obtain f' t1' t2' where flip: "(f', t2', t1') =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2697
    (if flip_rbt t2 t1 then (\<lambda>k v v'. f k v' v, t1, t2) else (f, t2, t1))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2698
    by fastforce
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2699
  have rbt_sorted': "rbt_sorted t1'" "rbt_sorted t2'"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2700
    using 1(3,4) flip
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2701
    by (auto split: if_splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2702
  show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2703
  proof (cases t1')
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2704
    case Empty
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2705
    show ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2706
      unfolding rbt_inter_rec.simps[of _ t1] flip[symmetric]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2707
      using flip rbt_sorted' rbt_split_props[of t2] rbtreeify_map_filter_inter[OF rbt_sorted'(2)]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2708
      by (auto simp: Empty split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2709
  next
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2710
    case (Branch c l1 a b r1)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2711
    {
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2712
      assume not_small: "\<not>small_rbt t2'"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2713
      obtain l2 \<beta> r2 where rbt_split_t2': "rbt_split t2' a = (l2, \<beta>, r2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2714
        by (cases "rbt_split t2' a") auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2715
      note rbt_split_t2'_props = rbt_split_props[OF rbt_split_t2' rbt_sorted'(2)]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2716
      have rbt_sort: "rbt_sorted l1" "rbt_sorted r1" "rbt_sorted l2" "rbt_sorted r2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2717
        using 1(3,4) flip
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2718
        by (auto simp: Branch rbt_split_t2'_props split: if_splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2719
      have inter_l1_l2: "rbt_sorted (rbt_inter_rec f' l1 l2)" "rbt_lookup (rbt_inter_rec f' l1 l2) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2720
        (case rbt_lookup l1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2721
        | Some v \<Rightarrow> (case rbt_lookup l2 k of None \<Rightarrow> None | Some w \<Rightarrow> Some (f' k v w)))" for k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2722
        using 1(1)[OF flip refl refl _ Branch rbt_split_t2'[symmetric]] rbt_sort rbt_split_t2'_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2723
        by (auto simp: not_small)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2724
      have inter_r1_r2: "rbt_sorted (rbt_inter_rec f' r1 r2)" "rbt_lookup (rbt_inter_rec f' r1 r2) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2725
        (case rbt_lookup r1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2726
        | Some v \<Rightarrow> (case rbt_lookup r2 k of None \<Rightarrow> None | Some w \<Rightarrow> Some (f' k v w)))" for k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2727
        using 1(2)[OF flip refl refl _ Branch rbt_split_t2'[symmetric]] rbt_sort rbt_split_t2'_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2728
        by (auto simp: not_small)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2729
      have inter_l1_l2_keys: "set (RBT_Impl.keys (rbt_inter_rec f' l1 l2)) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2730
        set (RBT_Impl.keys l1) \<inter> set (RBT_Impl.keys l2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2731
        using inter_l1_l2(1)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2732
        by (auto simp: rbt_lookup_iff_keys(1) inter_l1_l2(2) rbt_sort split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2733
      have inter_r1_r2_keys: "set (RBT_Impl.keys (rbt_inter_rec f' r1 r2)) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2734
        set (RBT_Impl.keys r1) \<inter> set (RBT_Impl.keys r2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2735
        using inter_r1_r2(1)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2736
        by (auto simp: rbt_lookup_iff_keys(1) inter_r1_r2(2) rbt_sort split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2737
      have inter_l1_l2_less: "rbt_inter_rec f' l1 l2 |\<guillemotleft> a"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2738
        using rbt_sorted'(1) rbt_split_t2'_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2739
        by (auto simp: Branch rbt_less_prop inter_l1_l2_keys)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2740
      have inter_r1_r2_greater: "a \<guillemotleft>| rbt_inter_rec f' r1 r2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2741
        using rbt_sorted'(1) rbt_split_t2'_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2742
        by (auto simp: Branch rbt_greater_prop inter_r1_r2_keys)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2743
      have rbt_lookup_join2: "rbt_lookup (rbt_join2 (rbt_inter_rec f' l1 l2) (rbt_inter_rec f' r1 r2)) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2744
        (case rbt_lookup (rbt_inter_rec f' l1 l2) k of None \<Rightarrow> rbt_lookup (rbt_inter_rec f' r1 r2) k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2745
        | Some v \<Rightarrow> Some v)" for k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2746
        using rbt_lookup_rbt_join2[OF inter_l1_l2(1) inter_r1_r2(1)] rbt_sorted'(1)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2747
        by (fastforce simp: Branch inter_l1_l2_keys inter_r1_r2_keys rbt_less_prop rbt_greater_prop)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2748
      have rbt_lookup_l1_k: "rbt_lookup l1 k = Some v \<Longrightarrow> k < a" for k v
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2749
        using rbt_sorted'(1) rbt_lookup_iff_keys(3)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2750
        by (auto simp: Branch rbt_less_prop)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2751
      have rbt_lookup_r1_k: "rbt_lookup r1 k = Some v \<Longrightarrow> a < k" for k v
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2752
        using rbt_sorted'(1) rbt_lookup_iff_keys(3)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2753
        by (auto simp: Branch rbt_greater_prop)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2754
      have "rbt_lookup (rbt_inter_rec f t1 t2) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2755
        (case rbt_lookup t1' k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2756
        | Some v \<Rightarrow> (case rbt_lookup t2' k of None \<Rightarrow> None | Some w \<Rightarrow> Some (f' k v w)))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2757
        by (auto simp: Let_def rbt_inter_rec.simps[of _ t1] flip[symmetric] not_small Branch
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2758
            rbt_split_t2' rbt_lookup_join2 rbt_lookup_rbt_join inter_l1_l2_less inter_r1_r2_greater
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2759
            rbt_split_lookup[OF rbt_split_t2' rbt_sorted'(2)] inter_l1_l2 inter_r1_r2
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2760
            split!: if_splits option.splits dest: rbt_lookup_l1_k rbt_lookup_r1_k)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2761
      moreover have "rbt_sorted (rbt_inter_rec f t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2762
        using rbt_sorted' inter_l1_l2 inter_r1_r2 rbt_split_t2'_props not_small
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2763
        by (auto simp: Let_def rbt_inter_rec.simps[of _ t1] flip[symmetric] Branch rbt_split_t2'
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2764
            rbt_less_prop rbt_greater_prop inter_l1_l2_less inter_r1_r2_greater
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2765
            inter_l1_l2_keys inter_r1_r2_keys intro!: rbt_sorted_rbt_join rbt_sorted_rbt_join2
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2766
            split: if_splits option.splits dest!: bspec)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2767
      ultimately have ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2768
        using flip
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2769
        by (auto split: if_splits split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2770
    }
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2771
    then show ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2772
      unfolding rbt_inter_rec.simps[of _ t1] flip[symmetric]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2773
      using rbt_sorted' flip rbtreeify_map_filter_inter[OF rbt_sorted'(2)]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2774
      by (auto split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2775
  qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2776
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2777
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2778
lemma rbt_lookup_delete:
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2779
  assumes "inv_12 t" "rbt_sorted t"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2780
  shows "rbt_lookup (rbt_delete x t) k = (if x = k then None else rbt_lookup t k)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2781
proof -
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2782
  note rbt_sorted_del = rbt_del_rbt_sorted[OF assms(2), of x]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2783
  show ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2784
    using assms rbt_sorted_del rbt_del_in_tree rbt_lookup_from_in_tree[OF assms(2) rbt_sorted_del]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2785
    by (fastforce simp: inv_12_def rbt_delete_def rbt_lookup_iff_keys(2) keys_entries)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2786
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2787
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2788
lemma fold_rbt_delete:
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2789
  assumes "inv_12 t1" "rbt_sorted t1" "rbt_sorted t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2790
  shows "inv_12 (RBT_Impl.fold (\<lambda>k _ t. rbt_delete k t) t2 t1) \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2791
    rbt_sorted (RBT_Impl.fold (\<lambda>k _ t. rbt_delete k t) t2 t1) \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2792
    rbt_lookup (RBT_Impl.fold (\<lambda>k _ t. rbt_delete k t) t2 t1) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2793
    (case rbt_lookup t1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2794
    | Some v \<Rightarrow> (case rbt_lookup t2 k of None \<Rightarrow> Some v | _ \<Rightarrow> None))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2795
proof -
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2796
  define xs where "xs = RBT_Impl.entries t2"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2797
  show "inv_12 (RBT_Impl.fold (\<lambda>k _ t. rbt_delete k t) t2 t1) \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2798
    rbt_sorted (RBT_Impl.fold (\<lambda>k _ t. rbt_delete k t) t2 t1) \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2799
    rbt_lookup (RBT_Impl.fold (\<lambda>k _ t. rbt_delete k t) t2 t1) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2800
    (case rbt_lookup t1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2801
    | Some v \<Rightarrow> (case rbt_lookup t2 k of None \<Rightarrow> Some v | _ \<Rightarrow> None))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2802
    using assms(1,2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2803
    unfolding map_of_entries[OF assms(3), symmetric] RBT_Impl.fold_def xs_def[symmetric]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2804
    by (induction xs arbitrary: t1 rule: rev_induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2805
       (auto simp: rbt_delete rbt_sorted_delete rbt_lookup_delete split!: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2806
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2807
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2808
lemma rbtreeify_filter_minus:
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2809
  assumes "rbt_sorted t1"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2810
  shows "rbt_sorted (rbtreeify (filter_minus t1 t2)) \<and>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2811
    rbt_lookup (rbtreeify (filter_minus t1 t2)) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2812
    (case rbt_lookup t1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2813
    | Some v \<Rightarrow> (case rbt_lookup t2 k of None \<Rightarrow> Some v | _ \<Rightarrow> None))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2814
proof -
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2815
  have map_of_filter: "map_of (filter (\<lambda>(k, _). rbt_lookup t2 k = None) xs) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2816
    (case map_of xs k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2817
    | Some v \<Rightarrow> (case rbt_lookup t2 k of None \<Rightarrow> Some v | Some x \<Rightarrow> Map.empty x))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2818
      for xs :: "('a \<times> 'b) list"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2819
    by (induction xs) (auto split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2820
  have map_fst_filter_minus: "map fst (filter_minus t1 t2) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2821
    filter (\<lambda>k. rbt_lookup t2 k = None) (map fst (RBT_Impl.entries t1))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2822
    by (auto simp: filter_minus_def filter_map comp_def case_prod_unfold)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2823
  have "sorted (map fst (filter_minus t1 t2))" "distinct (map fst (filter_minus t1 t2))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2824
    using distinct_filter distinct_entries[OF assms]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2825
      sorted_filter[of id] rbt_sorted_entries[OF assms]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2826
    by (auto simp: map_fst_filter_minus intro!: rbt_sorted_rbtreeify)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2827
  then show ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2828
    by (auto simp: rbt_lookup_rbtreeify filter_minus_def map_of_filter map_of_entries[OF assms]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2829
        intro!: rbt_sorted_rbtreeify)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2830
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2831
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2832
lemma rbt_lookup_minus_rec: "inv_12 t1 \<Longrightarrow> rbt_sorted t1 \<Longrightarrow> rbt_sorted t2 \<Longrightarrow>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2833
  rbt_sorted (rbt_minus_rec t1 t2) \<and> rbt_lookup (rbt_minus_rec t1 t2) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2834
  (case rbt_lookup t1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2835
  | Some v \<Rightarrow> (case rbt_lookup t2 k of None \<Rightarrow> Some v | _ \<Rightarrow> None))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2836
proof(induction t1 t2 arbitrary: k rule: rbt_minus_rec.induct)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2837
  case (1 t1 t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2838
  show ?case
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2839
  proof (cases t2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2840
    case Empty
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2841
    show ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2842
      using rbtreeify_filter_minus[OF 1(4)] 1(4)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2843
      by (auto simp: rbt_minus_rec.simps[of t1] Empty split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2844
  next
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2845
    case (Branch c l2 a b r2)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2846
    {
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2847
      assume not_small: "\<not>small_rbt t2" "\<not>small_rbt t1"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2848
      obtain l1 \<beta> r1 where rbt_split_t1: "rbt_split t1 a = (l1, \<beta>, r1)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2849
        by (cases "rbt_split t1 a") auto
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2850
      note rbt_split_t1_props = rbt_split_props[OF rbt_split_t1 1(4)]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2851
      have minus_l1_l2: "rbt_sorted (rbt_minus_rec l1 l2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2852
        "rbt_lookup (rbt_minus_rec l1 l2) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2853
        (case rbt_lookup l1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2854
        | Some v \<Rightarrow> (case rbt_lookup l2 k of None \<Rightarrow> Some v | Some x \<Rightarrow> None))" for k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2855
        using 1(1)[OF not_small Branch rbt_split_t1[symmetric] refl] 1(5) rbt_split_t1_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2856
          rbt_split[OF rbt_split_t1 1(3)]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2857
        by (auto simp: Branch)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2858
      have minus_r1_r2: "rbt_sorted (rbt_minus_rec r1 r2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2859
        "rbt_lookup (rbt_minus_rec r1 r2) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2860
        (case rbt_lookup r1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2861
        | Some v \<Rightarrow> (case rbt_lookup r2 k of None \<Rightarrow> Some v | Some x \<Rightarrow> None))" for k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2862
        using 1(2)[OF not_small Branch rbt_split_t1[symmetric] refl] 1(5) rbt_split_t1_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2863
          rbt_split[OF rbt_split_t1 1(3)]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2864
        by (auto simp: Branch)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2865
      have minus_l1_l2_keys: "set (RBT_Impl.keys (rbt_minus_rec l1 l2)) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2866
        set (RBT_Impl.keys l1) - set (RBT_Impl.keys l2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2867
        using minus_l1_l2(1) 1(5) rbt_lookup_iff_keys(3) rbt_split_t1_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2868
        by (auto simp: Branch rbt_lookup_iff_keys(1) minus_l1_l2(2) split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2869
      have minus_r1_r2_keys: "set (RBT_Impl.keys (rbt_minus_rec r1 r2)) =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2870
        set (RBT_Impl.keys r1) - set (RBT_Impl.keys r2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2871
        using minus_r1_r2(1) 1(5) rbt_lookup_iff_keys(3) rbt_split_t1_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2872
        by (auto simp: Branch rbt_lookup_iff_keys(1) minus_r1_r2(2) split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2873
      have rbt_lookup_join2: "rbt_lookup (rbt_join2 (rbt_minus_rec l1 l2) (rbt_minus_rec r1 r2)) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2874
        (case rbt_lookup (rbt_minus_rec l1 l2) k of None \<Rightarrow> rbt_lookup (rbt_minus_rec r1 r2) k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2875
        | Some v \<Rightarrow> Some v)" for k
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2876
        using rbt_lookup_rbt_join2[OF minus_l1_l2(1) minus_r1_r2(1)] rbt_split_t1_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2877
        by (fastforce simp: minus_l1_l2_keys minus_r1_r2_keys)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2878
      have lookup_l1_r1_a: "rbt_lookup l1 a = None" "rbt_lookup r1 a = None"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2879
        using rbt_split_t1_props
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2880
        by (auto simp: rbt_lookup_iff_keys(2))
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2881
      have "rbt_lookup (rbt_minus_rec t1 t2) k =
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2882
        (case rbt_lookup t1 k of None \<Rightarrow> None
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2883
        | Some v \<Rightarrow> (case rbt_lookup t2 k of None \<Rightarrow> Some v | _ \<Rightarrow> None))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2884
        using not_small rbt_lookup_iff_keys(2)[of l1] rbt_lookup_iff_keys(3)[of l1]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2885
          rbt_lookup_iff_keys(3)[of r1] rbt_split_t1_props
73526
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 73212
diff changeset
  2886
        using [[simp_depth_limit = 2]]
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2887
        by (auto simp: rbt_minus_rec.simps[of t1] Branch rbt_split_t1 rbt_lookup_join2
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2888
            minus_l1_l2(2) minus_r1_r2(2) rbt_split_lookup[OF rbt_split_t1 1(4)] lookup_l1_r1_a
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2889
            split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2890
      moreover have "rbt_sorted (rbt_minus_rec t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2891
        using not_small minus_l1_l2(1) minus_r1_r2(1) rbt_split_t1_props rbt_sorted_rbt_join2
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2892
        by (fastforce simp: rbt_minus_rec.simps[of t1] Branch rbt_split_t1 minus_l1_l2_keys minus_r1_r2_keys)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2893
      ultimately have ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2894
        by (auto split: if_splits split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2895
    }
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2896
    then show ?thesis
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2897
      using fold_rbt_delete[OF 1(3,4,5)] rbtreeify_filter_minus[OF 1(4)]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2898
      by (auto simp: rbt_minus_rec.simps[of t1])
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2899
  qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2900
qed
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2901
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2902
end
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2903
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2904
context ord begin
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2905
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2906
definition rbt_union_with_key :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2907
where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2908
  "rbt_union_with_key f t1 t2 = paint B (rbt_union_swap_rec f False t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2909
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2910
definition rbt_union_with where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2911
  "rbt_union_with f = rbt_union_with_key (\<lambda>_. f)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2912
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2913
definition rbt_union where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2914
  "rbt_union = rbt_union_with_key (%_ _ rv. rv)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2915
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2916
definition rbt_inter_with_key :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2917
where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2918
  "rbt_inter_with_key f t1 t2 = paint B (rbt_inter_swap_rec f False t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2919
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2920
definition rbt_inter_with where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2921
  "rbt_inter_with f = rbt_inter_with_key (\<lambda>_. f)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2922
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2923
definition rbt_inter where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2924
  "rbt_inter = rbt_inter_with_key (\<lambda>_ _ rv. rv)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2925
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2926
definition rbt_minus where
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2927
  "rbt_minus t1 t2 = paint B (rbt_minus_rec t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2928
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2929
end
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2930
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2931
context linorder begin
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2932
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2933
lemma is_rbt_rbt_unionwk [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2934
  "\<lbrakk> is_rbt t1; is_rbt t2 \<rbrakk> \<Longrightarrow> is_rbt (rbt_union_with_key f t1 t2)"
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2935
  using rbt_union_rec rbt_lookup_union_rec
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2936
  by (fastforce simp: rbt_union_with_key_def rbt_union_swap_rec is_rbt_def inv_12_def)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2937
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2938
lemma rbt_lookup_rbt_unionwk:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2939
  "\<lbrakk> rbt_sorted t1; rbt_sorted t2 \<rbrakk> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2940
  \<Longrightarrow> rbt_lookup (rbt_union_with_key f t1 t2) k = 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2941
  (case rbt_lookup t1 k of None \<Rightarrow> rbt_lookup t2 k 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2942
   | Some v \<Rightarrow> case rbt_lookup t2 k of None \<Rightarrow> Some v 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2943
              | Some w \<Rightarrow> Some (f k v w))"
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2944
  using rbt_lookup_union_rec
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2945
  by (auto simp: rbt_union_with_key_def rbt_union_swap_rec)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2946
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2947
lemma rbt_unionw_is_rbt: "\<lbrakk> is_rbt lt; is_rbt rt \<rbrakk> \<Longrightarrow> is_rbt (rbt_union_with f lt rt)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2948
by(simp add: rbt_union_with_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2949
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2950
lemma rbt_union_is_rbt: "\<lbrakk> is_rbt lt; is_rbt rt \<rbrakk> \<Longrightarrow> is_rbt (rbt_union lt rt)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2951
by(simp add: rbt_union_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2952
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2953
lemma rbt_lookup_rbt_union:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2954
  "\<lbrakk> rbt_sorted s; rbt_sorted t \<rbrakk> \<Longrightarrow>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2955
  rbt_lookup (rbt_union s t) = rbt_lookup s ++ rbt_lookup t"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2956
by(rule ext)(simp add: rbt_lookup_rbt_unionwk rbt_union_def map_add_def split: option.split)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2957
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2958
lemma rbt_interwk_is_rbt [simp]:
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2959
  "\<lbrakk> is_rbt t1; is_rbt t2 \<rbrakk> \<Longrightarrow> is_rbt (rbt_inter_with_key f t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2960
  using rbt_inter_rec rbt_lookup_inter_rec
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2961
  by (fastforce simp: rbt_inter_with_key_def rbt_inter_swap_rec is_rbt_def inv_12_def rbt_sorted_paint)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2962
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2963
lemma rbt_interw_is_rbt:
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2964
  "\<lbrakk> is_rbt t1; is_rbt t2 \<rbrakk> \<Longrightarrow> is_rbt (rbt_inter_with f t1 t2)"
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2965
by(simp add: rbt_inter_with_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2966
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2967
lemma rbt_inter_is_rbt:
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2968
  "\<lbrakk> is_rbt t1; is_rbt t2 \<rbrakk> \<Longrightarrow> is_rbt (rbt_inter t1 t2)"
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2969
by(simp add: rbt_inter_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2970
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2971
lemma rbt_lookup_rbt_interwk:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2972
  "\<lbrakk> rbt_sorted t1; rbt_sorted t2 \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2973
  \<Longrightarrow> rbt_lookup (rbt_inter_with_key f t1 t2) k =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2974
  (case rbt_lookup t1 k of None \<Rightarrow> None 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2975
   | Some v \<Rightarrow> case rbt_lookup t2 k of None \<Rightarrow> None
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2976
               | Some w \<Rightarrow> Some (f k v w))"
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2977
  using rbt_lookup_inter_rec
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2978
  by (auto simp: rbt_inter_with_key_def rbt_inter_swap_rec)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2979
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2980
lemma rbt_lookup_rbt_inter:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2981
  "\<lbrakk> rbt_sorted t1; rbt_sorted t2 \<rbrakk>
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2982
  \<Longrightarrow> rbt_lookup (rbt_inter t1 t2) = rbt_lookup t2 |` dom (rbt_lookup t1)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2983
by(auto simp add: rbt_inter_def rbt_lookup_rbt_interwk restrict_map_def split: option.split)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2984
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2985
lemma rbt_minus_is_rbt:
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2986
  "\<lbrakk> is_rbt t1; is_rbt t2 \<rbrakk> \<Longrightarrow> is_rbt (rbt_minus t1 t2)"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2987
  using rbt_minus_rec[of t1 t2] rbt_lookup_minus_rec[of t1 t2]
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2988
  by (auto simp: rbt_minus_def is_rbt_def inv_12_def)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2989
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2990
lemma rbt_lookup_rbt_minus:
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2991
  "\<lbrakk> is_rbt t1; is_rbt t2 \<rbrakk>
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2992
  \<Longrightarrow> rbt_lookup (rbt_minus t1 t2) = rbt_lookup t1 |` (- dom (rbt_lookup t2))"
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2993
  by (rule ext)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2994
     (auto simp: rbt_minus_def is_rbt_def inv_12_def restrict_map_def rbt_lookup_minus_rec
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2995
      split: option.splits)
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  2996
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2997
end
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2998
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2999
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  3000
subsection \<open>Code generator setup\<close>
49480
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3001
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3002
lemmas [code] =
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3003
  ord.rbt_less_prop
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3004
  ord.rbt_greater_prop
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3005
  ord.rbt_sorted.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3006
  ord.rbt_lookup.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3007
  ord.is_rbt_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3008
  ord.rbt_ins.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3009
  ord.rbt_insert_with_key_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3010
  ord.rbt_insertw_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3011
  ord.rbt_insert_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3012
  ord.rbt_del_from_left.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3013
  ord.rbt_del_from_right.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3014
  ord.rbt_del.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3015
  ord.rbt_delete_def
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  3016
  ord.rbt_split.simps
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  3017
  ord.rbt_union_swap_rec.simps
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  3018
  ord.map_filter_inter_def
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  3019
  ord.rbt_inter_swap_rec.simps
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  3020
  ord.filter_minus_def
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  3021
  ord.rbt_minus_rec.simps
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  3022
  ord.rbt_union_with_key_def
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3023
  ord.rbt_union_with_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3024
  ord.rbt_union_def
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  3025
  ord.rbt_inter_with_key_def
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  3026
  ord.rbt_inter_with_def
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  3027
  ord.rbt_inter_def
73211
bfa9f646f5ae optimize RBT_Impl
mraszyk
parents: 69593
diff changeset
  3028
  ord.rbt_minus_def
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3029
  ord.rbt_map_entry.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3030
  ord.rbt_bulkload_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3031
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3032
text \<open>More efficient implementations for \<^term>\<open>entries\<close> and \<^term>\<open>keys\<close>\<close>
49480
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3033
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3034
definition gen_entries :: 
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3035
  "(('a \<times> 'b) \<times> ('a, 'b) rbt) list \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a \<times> 'b) list"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3036
where
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  3037
  "gen_entries kvts t = entries t @ concat (map (\<lambda>(kv, t). kv # entries t) kvts)"
49480
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3038
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3039
lemma gen_entries_simps [simp, code]:
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3040
  "gen_entries [] Empty = []"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3041
  "gen_entries ((kv, t) # kvts) Empty = kv # gen_entries kvts t"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3042
  "gen_entries kvts (Branch c l k v r) = gen_entries (((k, v), r) # kvts) l"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3043
by(simp_all add: gen_entries_def)
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3044
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3045
lemma entries_code [code]:
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3046
  "entries = gen_entries []"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3047
by(simp add: gen_entries_def fun_eq_iff)
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3048
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3049
definition gen_keys :: "('a \<times> ('a, 'b) rbt) list \<Rightarrow> ('a, 'b) rbt \<Rightarrow> 'a list"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3050
where "gen_keys kts t = RBT_Impl.keys t @ concat (List.map (\<lambda>(k, t). k # keys t) kts)"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3051
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3052
lemma gen_keys_simps [simp, code]:
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3053
  "gen_keys [] Empty = []"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3054
  "gen_keys ((k, t) # kts) Empty = k # gen_keys kts t"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3055
  "gen_keys kts (Branch c l k v r) = gen_keys ((k, r) # kts) l"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3056
by(simp_all add: gen_keys_def)
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3057
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3058
lemma keys_code [code]:
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3059
  "keys = gen_keys []"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3060
by(simp add: gen_keys_def fun_eq_iff)
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  3061
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  3062
text \<open>Restore original type constraints for constants\<close>
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  3063
setup \<open>
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3064
  fold Sign.add_const_constraint
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3065
    [(\<^const_name>\<open>rbt_less\<close>, SOME \<^typ>\<open>('a :: order) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3066
     (\<^const_name>\<open>rbt_greater\<close>, SOME \<^typ>\<open>('a :: order) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3067
     (\<^const_name>\<open>rbt_sorted\<close>, SOME \<^typ>\<open>('a :: linorder, 'b) rbt \<Rightarrow> bool\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3068
     (\<^const_name>\<open>rbt_lookup\<close>, SOME \<^typ>\<open>('a :: linorder, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3069
     (\<^const_name>\<open>is_rbt\<close>, SOME \<^typ>\<open>('a :: linorder, 'b) rbt \<Rightarrow> bool\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3070
     (\<^const_name>\<open>rbt_ins\<close>, SOME \<^typ>\<open>('a::linorder \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3071
     (\<^const_name>\<open>rbt_insert_with_key\<close>, SOME \<^typ>\<open>('a::linorder \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3072
     (\<^const_name>\<open>rbt_insert_with\<close>, SOME \<^typ>\<open>('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a :: linorder) \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3073
     (\<^const_name>\<open>rbt_insert\<close>, SOME \<^typ>\<open>('a :: linorder) \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3074
     (\<^const_name>\<open>rbt_del_from_left\<close>, SOME \<^typ>\<open>('a::linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3075
     (\<^const_name>\<open>rbt_del_from_right\<close>, SOME \<^typ>\<open>('a::linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3076
     (\<^const_name>\<open>rbt_del\<close>, SOME \<^typ>\<open>('a::linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3077
     (\<^const_name>\<open>rbt_delete\<close>, SOME \<^typ>\<open>('a::linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3078
     (\<^const_name>\<open>rbt_union_with_key\<close>, SOME \<^typ>\<open>('a::linorder \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3079
     (\<^const_name>\<open>rbt_union_with\<close>, SOME \<^typ>\<open>('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a::linorder,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3080
     (\<^const_name>\<open>rbt_union\<close>, SOME \<^typ>\<open>('a::linorder,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3081
     (\<^const_name>\<open>rbt_map_entry\<close>, SOME \<^typ>\<open>'a::linorder \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt\<close>),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68450
diff changeset
  3082
     (\<^const_name>\<open>rbt_bulkload\<close>, SOME \<^typ>\<open>('a \<times> 'b) list \<Rightarrow> ('a::linorder,'b) rbt\<close>)]
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59575
diff changeset
  3083
\<close>
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  3084
73212
87e3c180044a hide the internal abbreviations MR and MB
Andreas Lochbihler <mail@andreas-lochbihler.de>
parents: 73211
diff changeset
  3085
hide_const (open) MR MB R B Empty entries keys fold gen_keys gen_entries
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  3086
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  3087
end