src/HOL/Library/RBT_Impl.thy
author blanchet
Wed, 12 Feb 2014 08:37:06 +0100
changeset 55417 01fbfb60c33e
parent 55414 eab03e9cee8a
child 55466 786edc984c98
permissions -rw-r--r--
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems * * * more transition of 'xxx_rec' to 'rec_xxx' and same for case * * * compile * * * 'rename_tac's to avoid referring to generated names * * * more robust scripts with 'rename_tac' * * * 'where' -> 'of' * * * 'where' -> 'of' * * * renamed 'xxx_rec' to 'rec_xxx'
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
36147
b43b22f63665 theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
haftmann
parents: 35618
diff changeset
     6
header {* Implementation of Red-Black Trees *}
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
36147
b43b22f63665 theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
haftmann
parents: 35618
diff changeset
    12
text {*
b43b22f63665 theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
haftmann
parents: 35618
diff changeset
    13
  For applications, you should use theory @{text RBT} which defines
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.
b43b22f63665 theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
haftmann
parents: 35618
diff changeset
    15
*}
b43b22f63665 theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
haftmann
parents: 35618
diff changeset
    16
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    17
subsection {* Datatype of RB trees *}
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    18
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
    19
datatype color = R | B
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    20
datatype ('a, 'b) rbt = Empty | Branch color "('a, 'b) rbt" 'a 'b "('a, 'b) rbt"
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
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    32
subsection {* Tree properties *}
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
    33
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    34
subsubsection {* Content of a tree *}
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
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
    69
subsubsection {* Search tree properties *}
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))"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   117
by (induct t) 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   118
  (force simp: sorted_append sorted_Cons rbt_ord_props 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   119
      dest!: entry_in_tree_keys)+
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   120
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   121
lemma distinct_entries:
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
   122
  "rbt_sorted t \<Longrightarrow> distinct (map fst (entries t))"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   123
by (induct t) 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   124
  (force simp: sorted_append sorted_Cons rbt_ord_props 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   125
      dest!: entry_in_tree_keys)+
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   126
48621
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
   127
lemma distinct_keys:
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
   128
  "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
   129
  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
   130
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
   131
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   132
subsubsection {* Tree lookup *}
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   133
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   134
primrec (in ord) rbt_lookup :: "('a, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   135
where
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   136
  "rbt_lookup Empty k = None"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   137
| "rbt_lookup (Branch _ l x y r) k = 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   138
   (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
   139
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   140
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
   141
  by (induct t) (auto simp: dom_def rbt_greater_prop rbt_less_prop)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   142
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   143
lemma dom_rbt_lookup_Branch: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   144
  "rbt_sorted (Branch c t1 k v t2) \<Longrightarrow> 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   145
    dom (rbt_lookup (Branch c t1 k v t2)) 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   146
    = Set.insert k (dom (rbt_lookup t1) \<union> dom (rbt_lookup t2))"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   147
proof -
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   148
  assume "rbt_sorted (Branch c t1 k v t2)"
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 49810
diff changeset
   149
  then show ?thesis by (simp add: rbt_lookup_keys)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   150
qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   151
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   152
lemma finite_dom_rbt_lookup [simp, intro!]: "finite (dom (rbt_lookup t))"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   153
proof (induct t)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   154
  case Empty then show ?case by simp
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   155
next
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   156
  case (Branch color t1 a b t2)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   157
  let ?A = "Set.insert a (dom (rbt_lookup t1) \<union> dom (rbt_lookup t2))"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   158
  have "dom (rbt_lookup (Branch color t1 a b t2)) \<subseteq> ?A" by (auto split: split_if_asm)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   159
  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
   160
  ultimately show ?case by (rule finite_subset)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   161
qed 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   162
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   163
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   164
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   165
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   166
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   167
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
   168
by (induct t) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   169
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   170
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
   171
by (induct t) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   172
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   173
lemma rbt_lookup_Empty: "rbt_lookup Empty = empty"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   174
by (rule ext) simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   175
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   176
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   177
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   178
context linorder begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   179
35618
b7bfd4cbcfc0 some lemma refinements
haftmann
parents: 35606
diff changeset
   180
lemma map_of_entries:
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   181
  "rbt_sorted t \<Longrightarrow> map_of (entries t) = rbt_lookup t"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   182
proof (induct t)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   183
  case Empty thus ?case by (simp add: rbt_lookup_Empty)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   184
next
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   185
  case (Branch c t1 k v t2)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   186
  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
   187
  proof (rule ext)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   188
    fix x
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   189
    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
   190
    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
   191
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   192
    have DOM_T1: "!!k'. k'\<in>dom (rbt_lookup t1) \<Longrightarrow> k>k'"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   193
    proof -
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   194
      fix k'
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   195
      from RBT_SORTED have "t1 |\<guillemotleft> k" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   196
      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
   197
      moreover assume "k'\<in>dom (rbt_lookup t1)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   198
      ultimately show "k>k'" using rbt_lookup_keys RBT_SORTED by auto
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   199
    qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   200
    
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   201
    have DOM_T2: "!!k'. k'\<in>dom (rbt_lookup t2) \<Longrightarrow> k<k'"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   202
    proof -
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   203
      fix k'
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   204
      from RBT_SORTED have "k \<guillemotleft>| t2" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   205
      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
   206
      moreover assume "k'\<in>dom (rbt_lookup t2)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   207
      ultimately show "k<k'" using rbt_lookup_keys RBT_SORTED by auto
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   208
    qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   209
    
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   210
    {
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   211
      assume C: "x<k"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   212
      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
   213
      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
   214
      moreover have "x \<notin> dom (rbt_lookup t2)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   215
      proof
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   216
        assume "x \<in> dom (rbt_lookup t2)"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   217
        with DOM_T2 have "k<x" by blast
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   218
        with C show False by simp
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   219
      qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   220
      ultimately have ?thesis by (simp add: map_add_upd_left map_add_dom_app_simps)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   221
    } moreover {
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   222
      assume [simp]: "x=k"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   223
      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
   224
      moreover have "x \<notin> dom (rbt_lookup t1)" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   225
      proof
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   226
        assume "x \<in> dom (rbt_lookup t1)"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   227
        with DOM_T1 have "k>x" by blast
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   228
        thus False by simp
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   229
      qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   230
      ultimately have ?thesis by (simp add: map_add_upd_left map_add_dom_app_simps)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   231
    } moreover {
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   232
      assume C: "x>k"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   233
      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
   234
      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
   235
      moreover have "x\<notin>dom (rbt_lookup t1)" proof
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   236
        assume "x\<in>dom (rbt_lookup t1)"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   237
        with DOM_T1 have "k>x" by simp
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   238
        with C show False by simp
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   239
      qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   240
      ultimately have ?thesis by (simp add: map_add_upd_left map_add_dom_app_simps)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   241
    } ultimately show ?thesis using less_linear by blast
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   242
  qed
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   243
  also from Branch 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   244
  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
   245
  finally show ?case by simp
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   246
qed
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   247
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   248
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
   249
  by (simp add: map_of_entries [symmetric] distinct_entries)
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   250
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   251
lemma set_entries_inject:
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   252
  assumes rbt_sorted: "rbt_sorted t1" "rbt_sorted t2" 
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   253
  shows "set (entries t1) = set (entries t2) \<longleftrightarrow> entries t1 = entries t2"
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   254
proof -
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   255
  from rbt_sorted have "distinct (map fst (entries t1))"
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   256
    "distinct (map fst (entries t2))"
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   257
    by (auto intro: distinct_entries)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   258
  with rbt_sorted show ?thesis
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   259
    by (auto intro: map_sorted_distinct_set_unique rbt_sorted_entries simp add: distinct_map)
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   260
qed
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   261
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   262
lemma entries_eqI:
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   263
  assumes rbt_sorted: "rbt_sorted t1" "rbt_sorted t2" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   264
  assumes rbt_lookup: "rbt_lookup t1 = rbt_lookup t2"
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   265
  shows "entries t1 = entries t2"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   266
proof -
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   267
  from rbt_sorted rbt_lookup have "map_of (entries t1) = map_of (entries t2)"
35618
b7bfd4cbcfc0 some lemma refinements
haftmann
parents: 35606
diff changeset
   268
    by (simp add: map_of_entries)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   269
  with rbt_sorted have "set (entries t1) = set (entries t2)"
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   270
    by (simp add: map_of_inject_set distinct_entries)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   271
  with rbt_sorted show ?thesis by (simp add: set_entries_inject)
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   272
qed
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   273
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   274
lemma entries_rbt_lookup:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   275
  assumes "rbt_sorted t1" "rbt_sorted t2" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   276
  shows "entries t1 = entries t2 \<longleftrightarrow> rbt_lookup t1 = rbt_lookup t2"
35618
b7bfd4cbcfc0 some lemma refinements
haftmann
parents: 35606
diff changeset
   277
  using assms by (auto intro: entries_eqI simp add: map_of_entries [symmetric])
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   278
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   279
lemma rbt_lookup_from_in_tree: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   280
  assumes "rbt_sorted t1" "rbt_sorted t2" 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   281
  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
   282
  shows "rbt_lookup t1 k = rbt_lookup t2 k"
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
   283
proof -
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   284
  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
   285
    by (simp add: keys_entries rbt_lookup_keys)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   286
  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
   287
qed
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   288
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   289
end
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   290
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   291
subsubsection {* Red-black properties *}
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   292
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   293
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
   294
where
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   295
  "color_of Empty = B"
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   296
| "color_of (Branch c _ _ _ _) = c"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   297
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   298
primrec bheight :: "('a,'b) rbt \<Rightarrow> nat"
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   299
where
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   300
  "bheight Empty = 0"
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   301
| "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
   302
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   303
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
   304
where
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   305
  "inv1 Empty = True"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   306
| "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
   307
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   308
primrec inv1l :: "('a, 'b) rbt \<Rightarrow> bool" -- {* Weaker version *}
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   309
where
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   310
  "inv1l Empty = True"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   311
| "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
   312
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
   313
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   314
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
   315
where
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   316
  "inv2 Empty = True"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   317
| "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
   318
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   319
context ord begin
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   320
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   321
definition is_rbt :: "('a, 'b) rbt \<Rightarrow> bool" where
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   322
  "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
   323
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   324
lemma is_rbt_rbt_sorted [simp]:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   325
  "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
   326
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   327
theorem Empty_is_rbt [simp]:
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   328
  "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
   329
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   330
end
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   331
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   332
subsection {* Insertion *}
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   333
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   334
fun (* slow, due to massive case splitting *)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   335
  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
   336
where
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   337
  "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
   338
  "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
   339
  "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
   340
  "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
   341
  "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
   342
  "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
   343
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   344
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
   345
  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
   346
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   347
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
   348
  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
   349
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   350
lemma balance_inv2: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   351
  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
   352
  shows "inv2 (balance l k v r)"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   353
  using assms
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   354
  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
   355
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   356
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   357
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   358
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
   359
  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
   360
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   361
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
   362
  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
   363
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   364
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   365
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   366
lemma (in linorder) balance_rbt_sorted: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   367
  fixes k :: "'a"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   368
  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
   369
  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
   370
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
   371
  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
   372
  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
   373
    by (auto simp add: rbt_ord_props)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   374
  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
   375
  with "2_2" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   376
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   377
  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
   378
  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
   379
    by simp
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   380
  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
   381
  with "3_2" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   382
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   383
  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
   384
  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
   385
  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
   386
  with "3_3" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   387
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   388
  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
   389
  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
   390
  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
   391
  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
   392
  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
   393
  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
   394
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   395
  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
   396
  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
   397
  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
   398
  with "4_2" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   399
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   400
  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
   401
  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
   402
  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
   403
  with "5_2" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   404
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   405
  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
   406
  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
   407
  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
   408
  with "5_3" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   409
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   410
  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
   411
  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
   412
  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
   413
  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
   414
  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
   415
  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
   416
qed simp+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   417
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   418
lemma entries_balance [simp]:
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   419
  "entries (balance l k v r) = entries l @ (k, v) # entries r"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   420
  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
   421
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   422
lemma keys_balance [simp]: 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   423
  "keys (balance l k v r) = keys l @ k # keys r"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   424
  by (simp add: keys_def)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   425
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   426
lemma balance_in_tree:  
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   427
  "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
   428
  by (auto simp add: keys_def)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   429
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   430
lemma (in linorder) rbt_lookup_balance[simp]: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   431
fixes k :: "'a"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   432
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
   433
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
   434
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
   435
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   436
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
   437
where
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   438
  "paint c Empty = Empty"
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   439
| "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
   440
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   441
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
   442
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
   443
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
   444
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
   445
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
   446
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   447
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   448
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   449
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
   450
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
   451
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
   452
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
   453
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   454
fun
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   455
  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
   456
where
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   457
  "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
   458
  "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
   459
                                       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
   460
                                       else Branch B l x (f k y v) r)" |
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   461
  "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
   462
                                       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
   463
                                       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
   464
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   465
lemma ins_inv1_inv2: 
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   466
  assumes "inv1 t" "inv2 t"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   467
  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
   468
  "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
   469
  using assms
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   470
  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
   471
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   472
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   473
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   474
context linorder begin
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   475
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   476
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
   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_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
   479
  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
   480
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
   481
  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
   482
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   483
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
   484
  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
   485
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   486
lemma rbt_lookup_ins: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   487
  fixes k :: "'a"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   488
  assumes "rbt_sorted t"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   489
  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
   490
                                                                | Some w \<Rightarrow> f k w v)) x"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   491
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
   492
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   493
end
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
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   496
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   497
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
   498
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
   499
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   500
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
   501
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   502
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
   503
  "rbt_insert = rbt_insert_with_key (\<lambda>_ _ nv. nv)"
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
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   506
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   507
context linorder begin
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   508
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   509
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
   510
  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
   511
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   512
theorem rbt_insertwk_is_rbt: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   513
  assumes inv: "is_rbt t" 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   514
  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
   515
using assms
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   516
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
   517
by (auto simp: ins_inv1_inv2)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   518
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   519
lemma rbt_lookup_rbt_insertwk: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   520
  assumes "rbt_sorted t"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   521
  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
   522
                                                       | Some w \<Rightarrow> f k w v)) x"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   523
unfolding rbt_insert_with_key_def using assms
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   524
by (simp add:rbt_lookup_ins)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   525
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   526
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
   527
  by (simp add: rbt_insertwk_rbt_sorted rbt_insertw_def)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   528
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
   529
  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
   530
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   531
lemma rbt_lookup_rbt_insertw:
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   532
  assumes "is_rbt t"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   533
  shows "rbt_lookup (rbt_insert_with f k v t) = (rbt_lookup t)(k \<mapsto> (if k:dom (rbt_lookup t) then f (the (rbt_lookup t k)) v else v))"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   534
using assms
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   535
unfolding rbt_insertw_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   536
by (rule_tac ext) (cases "rbt_lookup t k", auto simp:rbt_lookup_rbt_insertwk dom_def)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   537
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   538
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
   539
  by (simp add: rbt_insertwk_rbt_sorted rbt_insert_def)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   540
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
   541
  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
   542
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   543
lemma rbt_lookup_rbt_insert: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   544
  assumes "is_rbt t"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   545
  shows "rbt_lookup (rbt_insert k v t) = (rbt_lookup t)(k\<mapsto>v)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   546
unfolding rbt_insert_def
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   547
using assms
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   548
by (rule_tac ext) (simp add: 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
   549
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   550
end
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   551
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   552
subsection {* Deletion *}
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   553
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   554
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
   555
by (cases t rule: rbt_cases) auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   556
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   557
fun
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   558
  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
   559
where
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   560
  "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
   561
  "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
   562
  "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
   563
  "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
   564
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   565
lemma balance_left_inv2_with_inv1:
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   566
  assumes "inv2 lt" "inv2 rt" "bheight lt + 1 = bheight rt" "inv1 rt"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   567
  shows "bheight (balance_left lt k v rt) = bheight lt + 1"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   568
  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
   569
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   570
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
   571
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   572
lemma balance_left_inv2_app: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   573
  assumes "inv2 lt" "inv2 rt" "bheight lt + 1 = bheight rt" "color_of rt = B"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   574
  shows "inv2 (balance_left lt k v rt)" 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   575
        "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
   576
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   577
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
   578
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   579
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
   580
  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
   581
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   582
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
   583
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
   584
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   585
lemma (in linorder) balance_left_rbt_sorted: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   586
  "\<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
   587
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
   588
apply (auto simp: balance_rbt_sorted)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   589
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
   590
by force+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   591
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   592
context order begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   593
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   594
lemma balance_left_rbt_greater: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   595
  fixes k :: "'a"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   596
  assumes "k \<guillemotleft>| a" "k \<guillemotleft>| b" "k < x" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   597
  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
   598
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   599
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
   600
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   601
lemma balance_left_rbt_less: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   602
  fixes k :: "'a"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   603
  assumes "a |\<guillemotleft> k" "b |\<guillemotleft> k" "x < k" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   604
  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
   605
using assms
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   606
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
   607
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   608
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   609
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   610
lemma balance_left_in_tree: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   611
  assumes "inv1l l" "inv1 r" "bheight l + 1 = bheight r"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   612
  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
   613
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   614
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
   615
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   616
fun
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   617
  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
   618
where
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   619
  "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
   620
  "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
   621
  "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
   622
  "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
   623
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   624
lemma balance_right_inv2_with_inv1:
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   625
  assumes "inv2 lt" "inv2 rt" "bheight lt = bheight rt + 1" "inv1 lt"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   626
  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
   627
using assms
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   628
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
   629
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   630
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
   631
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
   632
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   633
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
   634
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
   635
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   636
lemma (in linorder) balance_right_rbt_sorted:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   637
  "\<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
   638
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
   639
apply (auto simp:balance_rbt_sorted)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   640
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
   641
by force+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   642
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   643
context order begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   644
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   645
lemma balance_right_rbt_greater: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   646
  fixes k :: "'a"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   647
  assumes "k \<guillemotleft>| a" "k \<guillemotleft>| b" "k < x" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   648
  shows "k \<guillemotleft>| balance_right a x t b"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   649
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
   650
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   651
lemma balance_right_rbt_less: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   652
  fixes k :: "'a"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   653
  assumes "a |\<guillemotleft> k" "b |\<guillemotleft> k" "x < k" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   654
  shows "balance_right a x t b |\<guillemotleft> k"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   655
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
   656
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   657
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   658
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   659
lemma balance_right_in_tree:
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   660
  assumes "inv1 l" "inv1l r" "bheight l = bheight r + 1" "inv2 l" "inv2 r"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   661
  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
   662
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
   663
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   664
fun
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   665
  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
   666
where
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   667
  "combine Empty x = x" 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   668
| "combine x Empty = x" 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   669
| "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
   670
                                    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
   671
                                    bc \<Rightarrow> Branch R a k x (Branch R bc s y d))" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   672
| "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
   673
                                    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
   674
                                    bc \<Rightarrow> balance_left a k x (Branch B bc s y d))" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   675
| "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
   676
| "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
   677
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   678
lemma combine_inv2:
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   679
  assumes "inv2 lt" "inv2 rt" "bheight lt = bheight rt"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   680
  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
   681
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   682
by (induct lt rt rule: combine.induct) 
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   683
   (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
   684
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   685
lemma combine_inv1: 
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   686
  assumes "inv1 lt" "inv1 rt"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   687
  shows "color_of lt = B \<Longrightarrow> color_of rt = B \<Longrightarrow> inv1 (combine lt rt)"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   688
         "inv1l (combine lt rt)"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   689
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   690
by (induct lt rt rule: combine.induct)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   691
   (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
   692
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   693
context linorder begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   694
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   695
lemma combine_rbt_greater[simp]: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   696
  fixes k :: "'a"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   697
  assumes "k \<guillemotleft>| l" "k \<guillemotleft>| r" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   698
  shows "k \<guillemotleft>| combine l r"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   699
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   700
by (induct l r rule: combine.induct)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   701
   (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
   702
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   703
lemma combine_rbt_less[simp]: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   704
  fixes k :: "'a"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   705
  assumes "l |\<guillemotleft> k" "r |\<guillemotleft> k" 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   706
  shows "combine l r |\<guillemotleft> k"
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   707
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   708
by (induct l r rule: combine.induct)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   709
   (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
   710
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   711
lemma combine_rbt_sorted: 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   712
  fixes k :: "'a"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   713
  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
   714
  shows "rbt_sorted (combine l r)"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   715
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
   716
  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
   717
  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
   718
    by auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   719
  with 3
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   720
  show ?case
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   721
    by (cases "combine b c" rule: rbt_cases)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   722
      (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
   723
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   724
  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
   725
  hence "x < k \<and> rbt_greater k c" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   726
  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
   727
  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
   728
  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
   729
  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
   730
  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
   731
  show ?case
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   732
  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
   733
    case Empty
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   734
    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
   735
    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
   736
    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
   737
      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
   738
    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
   739
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   740
    case (Red lta va ka rta)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   741
    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
   742
    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
   743
    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
   744
    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
   745
    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
   746
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   747
    case (Black lta va ka rta)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   748
    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
   749
    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
   750
    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
   751
      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
   752
    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
   753
  qed
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   754
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   755
  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
   756
  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
   757
  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
   758
  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
   759
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   760
  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
   761
  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
   762
  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
   763
  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
   764
qed simp+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   765
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   766
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   767
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   768
lemma combine_in_tree: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   769
  assumes "inv2 l" "inv2 r" "bheight l = bheight r" "inv1 l" "inv1 r"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   770
  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
   771
using assms 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   772
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
   773
  case (4 _ _ _ b c)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   774
  hence a: "bheight (combine b c) = bheight b" by (simp add: combine_inv2)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   775
  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
   776
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   777
  show ?case
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   778
  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
   779
    case Empty
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   780
    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
   781
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   782
    case (Red lta ka va rta)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   783
    with 4 show ?thesis by auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   784
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   785
    case (Black lta ka va rta)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   786
    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
   787
  qed 
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   788
qed (auto split: rbt.splits color.splits)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   789
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   790
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   791
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   792
fun
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   793
  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
   794
  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
   795
  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
   796
where
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   797
  "rbt_del x Empty = Empty" |
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   798
  "rbt_del x (Branch c a y s b) = 
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   799
   (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
   800
    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
   801
  "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
   802
  "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
   803
  "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
   804
  "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
   805
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   806
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   807
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   808
context linorder begin
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   809
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   810
lemma 
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   811
  assumes "inv2 lt" "inv1 lt"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   812
  shows
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   813
  "\<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
   814
   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
   815
   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
   816
   (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
   817
    (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
   818
  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
   819
  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
   820
  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
   821
  (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
   822
   (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
   823
  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
   824
  \<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
   825
using assms
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   826
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
   827
case (2 y c _ y')
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   828
  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
   829
  thus ?case proof (elim disjE)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   830
    assume "y = y'"
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   831
    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
   832
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   833
    assume "y < y'"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   834
    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
   835
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   836
    assume "y' < y"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   837
    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
   838
  qed
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   839
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   840
  case (3 y lt z v rta y' ss bb) 
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   841
  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
   842
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   843
  case (5 y a y' ss lt z v rta)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   844
  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
   845
next
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   846
  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
   847
qed auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   848
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   849
lemma 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   850
  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
   851
  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
   852
  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
   853
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
   854
   (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
   855
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   856
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
   857
  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
   858
  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
   859
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
   860
   (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
   861
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   862
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
   863
  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
   864
  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
   865
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
   866
  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
   867
  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
   868
  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
   869
  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
   870
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   871
  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
   872
  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
   873
  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
   874
  with "4_2" show ?case by simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   875
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   876
  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
   877
  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
   878
  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
   879
  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
   880
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   881
  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
   882
  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
   883
  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
   884
  with "6_2" show ?case by simp
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   885
qed (auto simp: combine_rbt_sorted)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   886
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   887
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
   888
  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
   889
  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
   890
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
   891
  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
   892
  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
   893
  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
   894
    assume "xx = yy"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   895
    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
   896
      case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   897
      from 2 `xx = yy` `xx = k` have "rbt_sorted (Branch c aa yy ss bb) \<and> k = yy" by simp
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   898
      hence "\<not> entry_in_tree k v aa" "\<not> entry_in_tree k v bb" by (auto simp: rbt_less_nit rbt_greater_prop)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   899
      with `xx = yy` 2 `xx = k` show ?thesis by (simp add: combine_in_tree)
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   900
    qed (simp add: combine_in_tree)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   901
  qed simp+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   902
next    
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   903
  case (3 xx lta zz vv rta yy ss bb)
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   904
  def mt[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
   905
  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
   906
  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
   907
  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
   908
  thus ?case proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   909
    case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   910
    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
   911
    hence "k \<guillemotleft>| bb" by (blast dest: rbt_greater_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   912
    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
   913
  qed auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   914
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   915
  case ("4_1" xx yy ss bb)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   916
  show ?case proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   917
    case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   918
    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
   919
    hence "k \<guillemotleft>| bb" by (blast dest: rbt_greater_trans)
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   920
    with "4_1" `xx = k` 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   921
   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
   922
    thus ?thesis by auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   923
  qed simp+
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   924
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   925
  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
   926
  thus ?case proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   927
    case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   928
    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
   929
    hence "k \<guillemotleft>| bb" by (blast dest: rbt_greater_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   930
    with True "4_2" show ?thesis by (auto simp: rbt_greater_nit)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   931
  qed auto
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   932
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   933
  case (5 xx aa yy ss lta zz vv rta)
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   934
  def mt[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
   935
  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
   936
  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
   937
  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
   938
  thus ?case proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   939
    case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   940
    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
   941
    hence "aa |\<guillemotleft> k" by (blast dest: rbt_less_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   942
    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
   943
  qed auto
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   944
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   945
  case ("6_1" xx aa yy ss)
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   946
  show ?case proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   947
    case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   948
    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
   949
    hence "aa |\<guillemotleft> k" by (blast dest: rbt_less_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   950
    with "6_1" `xx = k` 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
   951
  qed simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   952
next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   953
  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
   954
  thus ?case proof (cases "xx = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   955
    case True
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   956
    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
   957
    hence "aa |\<guillemotleft> k" by (blast dest: rbt_less_trans)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   958
    with True "6_2" show ?thesis by (auto simp: rbt_less_nit)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   959
  qed auto
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   960
qed simp
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   961
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   962
definition (in ord) rbt_delete where
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   963
  "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
   964
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   965
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
   966
proof -
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   967
  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
   968
  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
   969
  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
   970
  with assms show ?thesis
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   971
    unfolding is_rbt_def rbt_delete_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   972
    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
   973
qed
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   974
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   975
lemma rbt_delete_in_tree: 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   976
  assumes "is_rbt t" 
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   977
  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
   978
  using assms unfolding is_rbt_def rbt_delete_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   979
  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
   980
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   981
lemma rbt_lookup_rbt_delete:
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   982
  assumes is_rbt: "is_rbt t"
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   983
  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
   984
proof
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   985
  fix x
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   986
  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
   987
  proof (cases "x = k")
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   988
    assume "x = k" 
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
   989
    with is_rbt show ?thesis
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   990
      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
   991
  next
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   992
    assume "x \<noteq> k"
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   993
    thus ?thesis
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   994
      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
   995
  qed
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   996
qed
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
   997
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
   998
end
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
   999
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1000
subsection {* Modifying existing entries *}
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1001
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1002
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1003
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1004
primrec
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1005
  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
  1006
where
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1007
  "rbt_map_entry k f Empty = Empty"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1008
| "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
  1009
    (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
  1010
    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
  1011
    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
  1012
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1013
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1014
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
  1015
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
  1016
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
  1017
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
  1018
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
  1019
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
  1020
  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
  1021
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1022
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
  1023
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
  1024
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1025
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1026
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1027
theorem (in linorder) rbt_lookup_rbt_map_entry:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1028
  "rbt_lookup (rbt_map_entry k f t) = (rbt_lookup t)(k := Option.map 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
  1029
  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
  1030
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1031
subsection {* Mapping all entries *}
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1032
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1033
primrec
35602
e814157560e8 various refinements
haftmann
parents: 35550
diff changeset
  1034
  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
  1035
where
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1036
  "map f Empty = Empty"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1037
| "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
  1038
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1039
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
  1040
  by (induct t) auto
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1041
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
  1042
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
  1043
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
  1044
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
  1045
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1046
context ord begin
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1047
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1048
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
  1049
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
  1050
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
  1051
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
  1052
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
  1053
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1054
end
26192
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  1055
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1056
theorem (in linorder) rbt_lookup_map: "rbt_lookup (map f t) x = Option.map (f x) (rbt_lookup t x)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1057
  apply(induct t)
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1058
  apply auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1059
  apply(subgoal_tac "x = a")
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1060
  apply auto
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1061
  done
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1062
 (* 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
  1063
    by (induct t) auto *)
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1064
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1065
hide_const (open) map
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1066
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1067
subsection {* Folding over entries *}
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1068
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1069
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
  1070
  "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
  1071
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1072
lemma fold_simps [simp]:
35550
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1073
  "fold f Empty = id"
e2bc7f8d8d51 restructured RBT theory
haftmann
parents: 35534
diff changeset
  1074
  "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
  1075
  by (simp_all add: fold_def fun_eq_iff)
35534
14d8d72f8b1f more explicit naming scheme
haftmann
parents: 32245
diff changeset
  1076
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1077
lemma fold_code [code]:
49810
53f14f62cca2 fix code equation for RBT_Impl.fold
Andreas Lochbihler
parents: 49807
diff changeset
  1078
  "fold f Empty x = x"
53f14f62cca2 fix code equation for RBT_Impl.fold
Andreas Lochbihler
parents: 49807
diff changeset
  1079
  "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
  1080
by(simp_all)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1081
48621
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1082
(* fold with continuation predicate *)
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1083
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1084
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
  1085
  where
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1086
  "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
  1087
  "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
  1088
    if (c s) then
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1089
      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
  1090
        if (c s') then
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1091
          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
  1092
        else s'
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1093
    else 
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1094
      s
877df57629e3 a couple of additions to RBT formalization to allow us to implement RBT_Set
kuncar
parents: 47455
diff changeset
  1095
  )"
35606
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1096
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1097
subsection {* Bulkloading a tree *}
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1098
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1099
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
  1100
  "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
  1101
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1102
context linorder begin
35606
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1103
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1104
lemma rbt_bulkload_is_rbt [simp, intro]:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1105
  "is_rbt (rbt_bulkload xs)"
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1106
  unfolding rbt_bulkload_def by (induct xs) auto
35606
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1107
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1108
lemma rbt_lookup_rbt_bulkload:
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1109
  "rbt_lookup (rbt_bulkload xs) = map_of xs"
35606
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1110
proof -
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1111
  obtain ys where "ys = rev xs" by simp
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1112
  have "\<And>t. is_rbt t \<Longrightarrow>
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55412
diff changeset
  1113
    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
  1114
      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
  1115
  from this Empty_is_rbt have
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55412
diff changeset
  1116
    "rbt_lookup (List.fold (case_prod rbt_insert) (rev xs) Empty) = rbt_lookup Empty ++ map_of xs"
35606
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1117
     by (simp add: `ys = rev xs`)
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1118
  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
  1119
qed
7c5b40c7e8c4 added bulkload; tuned document
haftmann
parents: 35603
diff changeset
  1120
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1121
end
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  1122
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1123
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
subsection {* Building a RBT from a sorted list *}
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1126
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1127
text {* 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1128
  These functions have been adapted from 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1129
  Andrew W. Appel, Efficient Verified Red-Black Trees (September 2011) 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1130
*}
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1131
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1132
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
  1133
  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
  1134
where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1135
  "rbtreeify_f n kvs =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1136
   (if n = 0 then (Empty, kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1137
    else if n = 1 then
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1138
      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
  1139
    else if (n mod 2 = 0) then
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1140
      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
  1141
        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
  1142
    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
  1143
        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
  1144
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1145
| "rbtreeify_g n kvs =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1146
   (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
  1147
    else if n mod 2 = 0 then
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1148
      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
  1149
        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
  1150
    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
  1151
        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
  1152
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1153
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
  1154
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
  1155
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1156
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
  1157
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1158
lemma rbtreeify_f_code [code]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1159
  "rbtreeify_f n kvs =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1160
   (if n = 0 then (Empty, kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1161
    else if n = 1 then
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1162
      case kvs of (k, v) # kvs' \<Rightarrow> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1163
        (Branch R Empty k v Empty, kvs')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1164
    else let (n', r) = divmod_nat n 2 in
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1165
      if r = 0 then
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1166
        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
  1167
          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
  1168
      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
  1169
          apfst (Branch B t1 k v) (rbtreeify_f n' kvs'))"
55412
eb2caacf3ba4 avoid old 'prod.simps' -- better be more specific
blanchet
parents: 53374
diff changeset
  1170
by (subst rbtreeify_f.simps) (simp only: Let_def divmod_nat_div_mod prod.case)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1171
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1172
lemma rbtreeify_g_code [code]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1173
  "rbtreeify_g n kvs =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1174
   (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
  1175
    else let (n', r) = divmod_nat n 2 in
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1176
      if r = 0 then
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1177
        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
  1178
          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
  1179
      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
  1180
          apfst (Branch B t1 k v) (rbtreeify_g n' kvs'))"
55412
eb2caacf3ba4 avoid old 'prod.simps' -- better be more specific
blanchet
parents: 53374
diff changeset
  1181
by(subst rbtreeify_g.simps)(simp only: Let_def divmod_nat_div_mod prod.case)
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1182
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1183
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
  1184
by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1185
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1186
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
  1187
by arith
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1188
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1189
lemma rbtreeify_f_rec_aux_lemma:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1190
  "\<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
  1191
  \<Longrightarrow> k' - n div 2 = k - n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1192
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
  1193
apply(subst add_diff_assoc2, arith)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1194
apply(simp add: div2_plus_div2)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1195
done
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1196
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1197
lemma rbtreeify_f_simps:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1198
  "rbtreeify_f 0 kvs = (RBT_Impl.Empty, kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1199
  "rbtreeify_f (Suc 0) ((k, v) # kvs) = 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1200
  (Branch R Empty k v Empty, kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1201
  "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
  1202
   (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
  1203
     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
  1204
  "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
  1205
   (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
  1206
     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
  1207
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
  1208
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1209
lemma rbtreeify_g_simps:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1210
  "rbtreeify_g 0 kvs = (Empty, kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1211
  "rbtreeify_g (Suc 0) kvs = (Empty, kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1212
  "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
  1213
   (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
  1214
     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
  1215
  "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
  1216
   (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
  1217
     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
  1218
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
  1219
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1220
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
  1221
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1222
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
  1223
  \<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
  1224
  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
  1225
  \<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
  1226
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
  1227
  case (1 n kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1228
  show ?case
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1229
  proof(cases "n \<le> 1")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1230
    case True thus ?thesis using "1.prems"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1231
      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
  1232
  next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1233
    case False
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1234
    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
  1235
    note IH = "1.IH"[OF this]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1236
    show ?thesis
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1237
    proof(cases "n mod 2 = 0")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1238
      case True
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1239
      hence "length (snd (rbtreeify_f n kvs)) = 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1240
        length (snd (rbtreeify_f (2 * (n div 2)) kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1241
        by(metis minus_nat.diff_0 mult_div_cancel)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1242
      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
  1243
        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
  1244
      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
  1245
      note rbtreeify_f_simps(3)[OF this]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1246
      also note kvs[symmetric] 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1247
      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
  1248
      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
  1249
      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
  1250
      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
  1251
        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
  1252
         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
  1253
      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
  1254
      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
  1255
      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
  1256
        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
  1257
      with True kvs''[symmetric] refl refl
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1258
      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
  1259
        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
  1260
      finally show ?thesis using len[unfolded kvs''] "1.prems" True
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1261
        by(simp add: Suc_diff_le[symmetric] mult_2[symmetric] mult_div_cancel)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1262
    next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1263
      case False
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1264
      hence "length (snd (rbtreeify_f n kvs)) = 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1265
        length (snd (rbtreeify_f (Suc (2 * (n div 2))) kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1266
        by(metis Suc_eq_plus1_left comm_semiring_1_class.normalizing_semiring_rules(7)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1267
             mod_2_not_eq_zero_eq_one_nat semiring_div_class.mod_div_equality')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1268
      also from "1.prems" `\<not> n \<le> 1` obtain k v kvs' 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1269
        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
  1270
      also have "0 < n div 2" using `\<not> n \<le> 1` by(simp) 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1271
      note rbtreeify_f_simps(4)[OF this]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1272
      also note kvs[symmetric] 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1273
      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
  1274
      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
  1275
      with False 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
  1276
      with "1.prems" `\<not> n \<le> 1` obtain t1 k' v' kvs''
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1277
        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
  1278
        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
  1279
      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
  1280
      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
  1281
      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
  1282
        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
  1283
      with False kvs''[symmetric] refl refl
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1284
      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
  1285
        by(rule IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1286
      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
  1287
        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
  1288
    qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1289
  qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1290
next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1291
  case (2 n kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1292
  show ?case
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1293
  proof(cases "n > 1")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1294
    case False with `0 < n` show ?thesis
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1295
      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
  1296
  next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1297
    case True
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1298
    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
  1299
    note IH = "2.IH"[OF this]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1300
    show ?thesis
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1301
    proof(cases "n mod 2 = 0")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1302
      case True
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1303
      hence "length (snd (rbtreeify_g n kvs)) =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1304
        length (snd (rbtreeify_g (2 * (n div 2)) kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1305
        by(metis minus_nat.diff_0 mult_div_cancel)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1306
      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
  1307
        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
  1308
      also have "0 < n div 2" using `1 < n` by(simp) 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1309
      note rbtreeify_g_simps(3)[OF this]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1310
      also note kvs[symmetric] 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1311
      also let ?rest1 = "snd (rbtreeify_g (n div 2) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1312
      from "2.prems" `1 < n`
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1313
      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
  1314
      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
  1315
      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
  1316
        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
  1317
        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
  1318
      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
  1319
      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
  1320
      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
  1321
        using len "2.prems" unfolding kvs'' by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1322
      with True kvs''[symmetric] refl refl `0 < n div 2`
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1323
      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
  1324
        by(rule IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1325
      finally show ?thesis using len[unfolded kvs''] "2.prems" True
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1326
        by(simp add: Suc_diff_le[symmetric] mult_2[symmetric] mult_div_cancel)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1327
    next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1328
      case False
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1329
      hence "length (snd (rbtreeify_g n kvs)) = 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1330
        length (snd (rbtreeify_g (Suc (2 * (n div 2))) kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1331
        by(metis Suc_eq_plus1_left comm_semiring_1_class.normalizing_semiring_rules(7) 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1332
            mod_2_not_eq_zero_eq_one_nat semiring_div_class.mod_div_equality')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1333
      also from "2.prems" `1 < n` obtain k v kvs'
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1334
        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
  1335
      also have "0 < n div 2" using `1 < n` by(simp)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1336
      note rbtreeify_g_simps(4)[OF this]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1337
      also note kvs[symmetric] 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1338
      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
  1339
      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
  1340
      with False 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
  1341
      with "2.prems" `1 < n` False obtain t1 k' v' kvs'' 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1342
        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
  1343
        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
  1344
      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
  1345
      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
  1346
      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
  1347
        using len "2.prems" False unfolding kvs'' by simp arith
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1348
      with False kvs''[symmetric] refl refl `0 < n div 2`
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1349
      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
  1350
        by(rule IH)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1351
      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
  1352
        by(simp add: div2_plus_div2)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1353
    qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1354
  qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1355
qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1356
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1357
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
  1358
  fixes P Q
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1359
  defines "f0 == (\<And>kvs. P 0 kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1360
  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
  1361
  and "feven ==
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1362
    (\<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
  1363
       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
  1364
     \<Longrightarrow> P (2 * n) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1365
  and "fodd == 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1366
    (\<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
  1367
       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
  1368
    \<Longrightarrow> P (Suc (2 * n)) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1369
  and "g0 == (\<And>kvs. Q 0 kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1370
  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
  1371
  and "geven == 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1372
    (\<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
  1373
       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
  1374
    \<Longrightarrow> Q (2 * n) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1375
  and "godd == 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1376
    (\<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
  1377
       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
  1378
    \<Longrightarrow> Q (Suc (2 * n)) kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1379
  shows "\<lbrakk> n \<le> length kvs; 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1380
           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
  1381
           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
  1382
         \<Longrightarrow> P n kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1383
  and "\<lbrakk> n \<le> Suc (length kvs);
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1384
          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
  1385
          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
  1386
       \<Longrightarrow> Q n kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1387
proof -
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1388
  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
  1389
    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
  1390
  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
  1391
  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
  1392
    case (1 n kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1393
    show ?case
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1394
    proof(cases "n \<le> 1")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1395
      case True thus ?thesis using "1.prems"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1396
        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
  1397
          (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
  1398
    next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1399
      case False 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1400
      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
  1401
      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
  1402
      note IH = "1.IH"[OF ns]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1403
      show ?thesis
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1404
      proof(cases "n mod 2 = 0")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1405
        case True note ge0 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1406
        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
  1407
        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
  1408
        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
  1409
          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
  1410
          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
  1411
            (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
  1412
        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
  1413
        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
  1414
        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
  1415
        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
  1416
        moreover note feven[unfolded feven_def]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1417
          (* 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
  1418
        ultimately have "P (2 * (n div 2)) kvs" by -
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1419
        thus ?thesis using True by (metis div_mod_equality' minus_nat.diff_0 nat_mult_commute)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1420
      next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1421
        case False note ge0
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1422
        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
  1423
        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
  1424
        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
  1425
          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
  1426
          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
  1427
            (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
  1428
        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
  1429
        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
  1430
        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
  1431
        moreover note fodd[unfolded fodd_def]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1432
        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
  1433
        thus ?thesis using False 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1434
          by simp (metis One_nat_def Suc_eq_plus1_left le_add_diff_inverse mod_less_eq_dividend mult_div_cancel)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1435
      qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1436
    qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1437
  next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1438
    case (2 n kvs)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1439
    show ?case
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1440
    proof(cases "n \<le> 1")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1441
      case True thus ?thesis using "2.prems"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1442
        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
  1443
          (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
  1444
    next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1445
      case False 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1446
      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
  1447
      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
  1448
      note IH = "2.IH"[OF ns]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1449
      show ?thesis
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1450
      proof(cases "n mod 2 = 0")
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1451
        case True note ge0
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1452
        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
  1453
        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
  1454
        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
  1455
          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
  1456
          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
  1457
            (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
  1458
        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
  1459
        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
  1460
        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
  1461
        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
  1462
        moreover note geven[unfolded geven_def]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1463
        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
  1464
        thus ?thesis using True 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1465
          by(metis div_mod_equality' minus_nat.diff_0 nat_mult_commute)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1466
      next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1467
        case False note ge0
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1468
        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
  1469
        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
  1470
        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
  1471
          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
  1472
          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
  1473
            (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
  1474
        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
  1475
        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
  1476
        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
  1477
        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
  1478
        moreover note godd[unfolded godd_def]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1479
        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
  1480
        thus ?thesis using False 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1481
          by simp (metis One_nat_def Suc_eq_plus1_left le_add_diff_inverse mod_less_eq_dividend mult_div_cancel)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1482
      qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1483
    qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1484
  qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1485
qed
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
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
  1488
  \<Longrightarrow> inv1 (fst (rbtreeify_f n kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1489
  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
  1490
  \<Longrightarrow> inv1 (fst (rbtreeify_g n kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1491
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
  1492
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1493
fun plog2 :: "nat \<Rightarrow> nat" 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1494
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
  1495
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1496
declare plog2.simps [simp del]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1497
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1498
lemma plog2_simps [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1499
  "plog2 0 = 0" "plog2 (Suc 0) = 0"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1500
  "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
  1501
  "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
  1502
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
  1503
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1504
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
  1505
  \<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
  1506
  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
  1507
  \<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
  1508
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
  1509
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1510
lemma bheight_rbtreeify_f_eq_plog2I:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1511
  "\<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
  1512
  \<Longrightarrow> bheight t = plog2 n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1513
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
  1514
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1515
lemma bheight_rbtreeify_g_eq_plog2I: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1516
  "\<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
  1517
  \<Longrightarrow> bheight t = plog2 n"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1518
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
  1519
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1520
hide_const (open) plog2
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1521
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1522
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
  1523
  \<Longrightarrow> inv2 (fst (rbtreeify_f n kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1524
  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
  1525
  \<Longrightarrow> inv2 (fst (rbtreeify_g n kvs))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1526
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
  1527
  (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
  1528
        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
  1529
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1530
lemma "n \<le> length kvs \<Longrightarrow> True"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1531
  and color_of_rbtreeify_g:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1532
  "\<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
  1533
  \<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
  1534
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
  1535
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1536
lemma entries_rbtreeify_f_append:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1537
  "n \<le> length kvs 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1538
  \<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
  1539
  and entries_rbtreeify_g_append: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1540
  "n \<le> Suc (length kvs) 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1541
  \<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
  1542
by(induction rule: rbtreeify_induct) simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1543
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1544
lemma length_entries_rbtreeify_f:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1545
  "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
  1546
  and length_entries_rbtreeify_g: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1547
  "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
  1548
by(induct rule: rbtreeify_induct) simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1549
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1550
lemma rbtreeify_f_conv_drop: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1551
  "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
  1552
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
  1553
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
  1554
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1555
lemma rbtreeify_g_conv_drop: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1556
  "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
  1557
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
  1558
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
  1559
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1560
lemma entries_rbtreeify_f [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1561
  "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
  1562
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
  1563
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
  1564
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1565
lemma entries_rbtreeify_g [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1566
  "n \<le> Suc (length kvs) \<Longrightarrow> 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1567
  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
  1568
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
  1569
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
  1570
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1571
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
  1572
  \<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
  1573
by(simp add: keys_def take_map)
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 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
  1576
  \<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
  1577
by(simp add: keys_def take_map)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1578
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1579
lemma rbtreeify_fD: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1580
  "\<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
  1581
  \<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
  1582
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
  1583
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1584
lemma rbtreeify_gD: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1585
  "\<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
  1586
  \<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
  1587
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
  1588
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1589
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
  1590
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
  1591
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1592
context linorder begin
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1593
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1594
lemma rbt_sorted_rbtreeify_f: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1595
  "\<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
  1596
  \<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
  1597
  and rbt_sorted_rbtreeify_g: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1598
  "\<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
  1599
  \<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
  1600
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
  1601
  case (f_even n kvs t k v kvs')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1602
  from rbtreeify_fD[OF `rbtreeify_f n kvs = (t, (k, v) # kvs')` `n \<le> length kvs`]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1603
  have "entries t = take n kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1604
    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
  1605
  hence unfold: "kvs = take n kvs @ (k, v) # kvs'" by(metis append_take_drop_id)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1606
  from `sorted (map fst kvs)` kvs'
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1607
  have "(\<forall>(x, y) \<in> set (take n kvs). x \<le> k) \<and> (\<forall>(x, y) \<in> set kvs'. k \<le> x)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1608
    by(subst (asm) unfold)(auto simp add: sorted_append sorted_Cons)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1609
  moreover from `distinct (map fst kvs)` kvs'
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1610
  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
  1611
    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
  1612
  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
  1613
    by fastforce
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1614
  hence "fst (rbtreeify_f n kvs) |\<guillemotleft> k" "k \<guillemotleft>| fst (rbtreeify_g n kvs')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1615
    using `n \<le> Suc (length kvs')` `n \<le> length kvs` set_take_subset[of "n - 1" kvs']
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1616
    by(auto simp add: ord.rbt_greater_prop ord.rbt_less_prop take_map split_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1617
  moreover from `sorted (map fst kvs)` `distinct (map fst kvs)`
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1618
  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
  1619
  moreover have "sorted (map fst kvs')" "distinct (map fst kvs')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1620
    using `sorted (map fst kvs)` `distinct (map fst kvs)`
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1621
    by(subst (asm) (1 2) unfold, simp add: sorted_append sorted_Cons)+
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1622
  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
  1623
  ultimately show ?case
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1624
    using `0 < n` `rbtreeify_f n kvs = (t, (k, v) # kvs')` by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1625
next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1626
  case (f_odd n kvs t k v kvs')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1627
  from rbtreeify_fD[OF `rbtreeify_f n kvs = (t, (k, v) # kvs')` `n \<le> length kvs`]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1628
  have "entries t = take n kvs" 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1629
    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
  1630
  hence unfold: "kvs = take n kvs @ (k, v) # kvs'" by(metis append_take_drop_id)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1631
  from `sorted (map fst kvs)` kvs'
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1632
  have "(\<forall>(x, y) \<in> set (take n kvs). x \<le> k) \<and> (\<forall>(x, y) \<in> set kvs'. k \<le> x)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1633
    by(subst (asm) unfold)(auto simp add: sorted_append sorted_Cons)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1634
  moreover from `distinct (map fst kvs)` kvs'
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1635
  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
  1636
    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
  1637
  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
  1638
    by fastforce
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1639
  hence "fst (rbtreeify_f n kvs) |\<guillemotleft> k" "k \<guillemotleft>| fst (rbtreeify_f n kvs')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1640
    using `n \<le> length kvs'` `n \<le> length kvs` set_take_subset[of n kvs']
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1641
    by(auto simp add: rbt_greater_prop rbt_less_prop take_map split_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1642
  moreover from `sorted (map fst kvs)` `distinct (map fst kvs)`
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1643
  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
  1644
  moreover have "sorted (map fst kvs')" "distinct (map fst kvs')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1645
    using `sorted (map fst kvs)` `distinct (map fst kvs)`
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1646
    by(subst (asm) (1 2) unfold, simp add: sorted_append sorted_Cons)+
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1647
  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
  1648
  ultimately show ?case 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1649
    using `0 < n` `rbtreeify_f n kvs = (t, (k, v) # kvs')` by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1650
next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1651
  case (g_even n kvs t k v kvs')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1652
  from rbtreeify_gD[OF `rbtreeify_g n kvs = (t, (k, v) # kvs')` `n \<le> Suc (length kvs)`]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1653
  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
  1654
    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
  1655
  hence unfold: "kvs = take (n - 1) kvs @ (k, v) # kvs'" by(metis append_take_drop_id)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1656
  from `sorted (map fst kvs)` kvs'
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1657
  have "(\<forall>(x, y) \<in> set (take (n - 1) kvs). x \<le> k) \<and> (\<forall>(x, y) \<in> set kvs'. k \<le> x)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1658
    by(subst (asm) unfold)(auto simp add: sorted_append sorted_Cons)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1659
  moreover from `distinct (map fst kvs)` kvs'
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1660
  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
  1661
    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
  1662
  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
  1663
    by fastforce
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1664
  hence "fst (rbtreeify_g n kvs) |\<guillemotleft> k" "k \<guillemotleft>| fst (rbtreeify_g n kvs')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1665
    using `n \<le> Suc (length kvs')` `n \<le> Suc (length kvs)` set_take_subset[of "n - 1" kvs']
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1666
    by(auto simp add: rbt_greater_prop rbt_less_prop take_map split_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1667
  moreover from `sorted (map fst kvs)` `distinct (map fst kvs)`
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1668
  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
  1669
  moreover have "sorted (map fst kvs')" "distinct (map fst kvs')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1670
    using `sorted (map fst kvs)` `distinct (map fst kvs)`
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1671
    by(subst (asm) (1 2) unfold, simp add: sorted_append sorted_Cons)+
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1672
  hence "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
  1673
  ultimately show ?case using `0 < n` `rbtreeify_g n kvs = (t, (k, v) # kvs')` by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1674
next
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1675
  case (g_odd n kvs t k v kvs')
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1676
  from rbtreeify_fD[OF `rbtreeify_f n kvs = (t, (k, v) # kvs')` `n \<le> length kvs`]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1677
  have "entries t = take n kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1678
    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
  1679
  hence unfold: "kvs = take n kvs @ (k, v) # kvs'" by(metis append_take_drop_id)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1680
  from `sorted (map fst kvs)` kvs'
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1681
  have "(\<forall>(x, y) \<in> set (take n kvs). x \<le> k) \<and> (\<forall>(x, y) \<in> set kvs'. k \<le> x)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1682
    by(subst (asm) unfold)(auto simp add: sorted_append sorted_Cons)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1683
  moreover from `distinct (map fst kvs)` kvs'
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1684
  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
  1685
    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
  1686
  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
  1687
    by fastforce
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1688
  hence "fst (rbtreeify_f n kvs) |\<guillemotleft> k" "k \<guillemotleft>| fst (rbtreeify_g n kvs')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1689
    using `n \<le> Suc (length kvs')` `n \<le> length kvs` set_take_subset[of "n - 1" kvs']
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1690
    by(auto simp add: rbt_greater_prop rbt_less_prop take_map split_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1691
  moreover from `sorted (map fst kvs)` `distinct (map fst kvs)`
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1692
  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
  1693
  moreover have "sorted (map fst kvs')" "distinct (map fst kvs')"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1694
    using `sorted (map fst kvs)` `distinct (map fst kvs)`
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1695
    by(subst (asm) (1 2) unfold, simp add: sorted_append sorted_Cons)+
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1696
  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
  1697
  ultimately show ?case
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1698
    using `0 < n` `rbtreeify_f n kvs = (t, (k, v) # kvs')` by simp
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1699
qed simp_all
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_sorted_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> rbt_sorted (rbtreeify kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1703
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
  1704
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1705
lemma is_rbt_rbtreeify: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1706
  "\<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
  1707
  \<Longrightarrow> is_rbt (rbtreeify kvs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1708
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
  1709
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1710
lemma rbt_lookup_rbtreeify:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1711
  "\<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
  1712
  rbt_lookup (rbtreeify kvs) = map_of kvs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1713
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
  1714
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1715
end
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1716
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1717
text {* 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1718
  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
  1719
  Andrew W. Appel, Efficient Verified Red-Black Trees (September 2011)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1720
*}
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1721
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1722
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
  1723
where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1724
  "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
  1725
| "skip_red t = t"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1726
49807
9a0843995fd3 correct definition for skip_black
Andreas Lochbihler
parents: 49770
diff changeset
  1727
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
  1728
where
49807
9a0843995fd3 correct definition for skip_black
Andreas Lochbihler
parents: 49770
diff changeset
  1729
  "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
  1730
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1731
datatype compare = LT | GT | EQ
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1732
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1733
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
  1734
where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1735
  "compare_height sx s t tx =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1736
  (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
  1737
     (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
  1738
       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
  1739
   | (_, rbt.Empty, _, Branch _ _ _ _ _) \<Rightarrow> LT
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1740
   | (Branch _ _ _ _ _, _, rbt.Empty, _) \<Rightarrow> GT
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1741
   | (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
  1742
       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
  1743
   | (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
  1744
       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
  1745
   | _ \<Rightarrow> EQ)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1746
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1747
declare compare_height.simps [code]
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1748
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1749
hide_type (open) compare
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1750
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
  1751
  compare_height skip_black skip_red LT GT EQ case_compare rec_compare
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 55414
diff changeset
  1752
  Abs_compare Rep_compare rep_set_compare
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1753
hide_fact (open)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1754
  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
  1755
  Rep_compare Rep_compare_cases Rep_compare_induct Rep_compare_inject Rep_compare_inverse
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1756
  compare.simps compare.exhaust compare.induct compare.recs compare.simps
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1757
  compare.size compare.case_cong compare.weak_case_cong compare.cases 
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 55414
diff changeset
  1758
  compare.nchotomy compare.split compare.split_asm rec_compare_def
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1759
  compare.eq.refl compare.eq.simps
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1760
  compare.EQ_def compare.GT_def compare.LT_def
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1761
  equal_compare_def
49807
9a0843995fd3 correct definition for skip_black
Andreas Lochbihler
parents: 49770
diff changeset
  1762
  skip_red_def skip_red.simps skip_red.cases skip_red.induct 
9a0843995fd3 correct definition for skip_black
Andreas Lochbihler
parents: 49770
diff changeset
  1763
  skip_black_def
9a0843995fd3 correct definition for skip_black
Andreas Lochbihler
parents: 49770
diff changeset
  1764
  compare_height_def compare_height.simps
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1765
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1766
subsection {* union and intersection of sorted associative lists *}
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1767
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1768
context ord begin
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 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
  1771
where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1772
  "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
  1773
   (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
  1774
    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
  1775
    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
  1776
| "sunion_with f [] bs = bs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1777
| "sunion_with f as [] = as"
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
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
  1782
where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1783
  "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
  1784
  (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
  1785
   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
  1786
   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
  1787
| "sinter_with f [] _ = []"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1788
| "sinter_with f _ [] = []"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1789
by pat_completeness auto
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1790
termination by lexicographic_order
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1791
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1792
end
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1793
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1794
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
  1795
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1796
context linorder begin
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1797
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1798
lemma set_fst_sunion_with: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1799
  "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
  1800
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
  1801
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1802
lemma sorted_sunion_with [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1803
  "\<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
  1804
  \<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
  1805
by(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
  1806
  (auto simp add: sorted_Cons set_fst_sunion_with simp del: set_map)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1807
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1808
lemma distinct_sunion_with [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1809
  "\<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
  1810
  \<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
  1811
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
  1812
  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
  1813
  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
  1814
  thus ?case using "1"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1815
    by(auto simp add: set_fst_sunion_with sorted_Cons simp del: set_map)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1816
qed simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1817
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1818
lemma map_of_sunion_with: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1819
  "\<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
  1820
  \<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
  1821
  (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
  1822
  | 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
  1823
              | Some w \<Rightarrow> Some (f k v w))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1824
by(induct f xs ys rule: sunion_with.induct)(auto simp add: sorted_Cons split: option.split dest: map_of_SomeD bspec)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1825
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1826
lemma set_fst_sinter_with [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1827
  "\<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
  1828
  \<Longrightarrow> set (map fst (sinter_with f xs ys)) = set (map fst xs) \<inter> set (map fst ys)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1829
by(induct f xs ys rule: sinter_with.induct)(auto simp add: sorted_Cons simp del: set_map)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1830
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1831
lemma set_fst_sinter_with_subset1:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1832
  "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
  1833
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
  1834
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1835
lemma set_fst_sinter_with_subset2:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1836
  "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
  1837
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
  1838
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1839
lemma sorted_sinter_with [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1840
  "\<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
  1841
  \<Longrightarrow> sorted (map fst (sinter_with f xs ys))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1842
by(induct f xs ys rule: sinter_with.induct)(auto simp add: sorted_Cons simp del: set_map)
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 distinct_sinter_with [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1845
  "\<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
  1846
  \<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
  1847
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
  1848
  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
  1849
  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
  1850
  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
  1851
    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
  1852
    by(auto simp del: set_map)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1853
qed simp_all
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1854
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1855
lemma map_of_sinter_with:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1856
  "\<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
  1857
  \<Longrightarrow> map_of (sinter_with f xs ys) k = 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1858
  (case map_of xs k of None \<Rightarrow> None | Some v \<Rightarrow> Option.map (f k v) (map_of ys k))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1859
apply(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
  1860
apply(auto simp add: sorted_Cons Option.map_def split: option.splits dest: map_of_SomeD bspec)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1861
done
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1862
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1863
end
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 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
  1866
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
  1867
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1868
lemma map_map_filter: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1869
  "map f (List.map_filter g xs) = List.map_filter (Option.map f \<circ> g) xs"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1870
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
  1871
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1872
lemma map_filter_option_map_const: 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1873
  "List.map_filter (\<lambda>x. Option.map (\<lambda>y. f x) (g (f x))) xs = filter (\<lambda>x. g x \<noteq> None) (map f xs)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1874
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
  1875
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1876
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
  1877
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
  1878
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1879
context ord begin
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1880
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1881
definition rbt_union_with_key :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1882
where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1883
  "rbt_union_with_key f t1 t2 =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1884
  (case RBT_Impl.compare_height t1 t1 t2 t2
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1885
   of compare.EQ \<Rightarrow> rbtreeify (sunion_with f (entries t1) (entries t2))
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1886
    | compare.LT \<Rightarrow> fold (rbt_insert_with_key (\<lambda>k v w. f k w v)) t1 t2
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1887
    | compare.GT \<Rightarrow> 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
  1888
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1889
definition rbt_union_with where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1890
  "rbt_union_with f = rbt_union_with_key (\<lambda>_. f)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1891
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1892
definition rbt_union where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1893
  "rbt_union = rbt_union_with_key (%_ _ rv. rv)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1894
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1895
definition rbt_inter_with_key :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1896
where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1897
  "rbt_inter_with_key f t1 t2 =
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1898
  (case RBT_Impl.compare_height t1 t1 t2 t2 
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1899
   of compare.EQ \<Rightarrow> rbtreeify (sinter_with f (entries t1) (entries t2))
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1900
    | compare.LT \<Rightarrow> rbtreeify (List.map_filter (\<lambda>(k, v). Option.map (\<lambda>w. (k, f k v w)) (rbt_lookup t2 k)) (entries t1))
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1901
    | compare.GT \<Rightarrow> rbtreeify (List.map_filter (\<lambda>(k, v). Option.map (\<lambda>w. (k, f k w v)) (rbt_lookup t1 k)) (entries t2)))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1902
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1903
definition rbt_inter_with where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1904
  "rbt_inter_with f = rbt_inter_with_key (\<lambda>_. f)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1905
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1906
definition rbt_inter where
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1907
  "rbt_inter = rbt_inter_with_key (\<lambda>_ _ rv. rv)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1908
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1909
end
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1910
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1911
context linorder begin
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1912
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1913
lemma rbt_sorted_entries_right_unique:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1914
  "\<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
  1915
     rbt_sorted t \<rbrakk> \<Longrightarrow> v = v'"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1916
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
  1917
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1918
lemma rbt_sorted_fold_rbt_insertwk:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1919
  "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
  1920
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
  1921
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1922
lemma is_rbt_fold_rbt_insertwk:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1923
  assumes "is_rbt t1"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1924
  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
  1925
proof -
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1926
  def xs \<equiv> "entries t2"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1927
  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
  1928
    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
  1929
qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1930
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1931
lemma rbt_lookup_fold_rbt_insertwk:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1932
  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
  1933
  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
  1934
  (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
  1935
   | 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
  1936
               | 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
  1937
proof -
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1938
  def xs \<equiv> "entries t1"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1939
  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
  1940
  with t2 show ?thesis
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1941
    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
  1942
      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
  1943
    apply(induct xs rule: rev_induct)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1944
    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
  1945
    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
  1946
    done
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1947
qed
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1948
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1949
lemma is_rbt_rbt_unionwk [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1950
  "\<lbrakk> is_rbt t1; is_rbt t2 \<rbrakk> \<Longrightarrow> is_rbt (rbt_union_with_key f t1 t2)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1951
by(simp add: rbt_union_with_key_def Let_def is_rbt_fold_rbt_insertwk is_rbt_rbtreeify rbt_sorted_entries distinct_entries split: compare.split)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1952
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1953
lemma rbt_lookup_rbt_unionwk:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1954
  "\<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
  1955
  \<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
  1956
  (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
  1957
   | 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
  1958
              | Some w \<Rightarrow> Some (f k v w))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1959
by(auto simp add: rbt_union_with_key_def Let_def rbt_lookup_fold_rbt_insertwk rbt_sorted_entries distinct_entries map_of_sunion_with map_of_entries rbt_lookup_rbtreeify split: option.split compare.split)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1960
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1961
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
  1962
by(simp add: rbt_union_with_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1963
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1964
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
  1965
by(simp add: rbt_union_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1966
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1967
lemma rbt_lookup_rbt_union:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1968
  "\<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
  1969
  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
  1970
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
  1971
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1972
lemma rbt_interwk_is_rbt [simp]:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1973
  "\<lbrakk> rbt_sorted t1; rbt_sorted t2 \<rbrakk> \<Longrightarrow> is_rbt (rbt_inter_with_key f t1 t2)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1974
by(auto simp add: rbt_inter_with_key_def Let_def map_map_filter split_def o_def option_map_comp map_filter_option_map_const sorted_filter[where f=id, simplified] rbt_sorted_entries distinct_entries intro: is_rbt_rbtreeify split: compare.split)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1975
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1976
lemma rbt_interw_is_rbt:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1977
  "\<lbrakk> rbt_sorted t1; rbt_sorted t2 \<rbrakk> \<Longrightarrow> is_rbt (rbt_inter_with f t1 t2)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1978
by(simp add: rbt_inter_with_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1979
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1980
lemma rbt_inter_is_rbt:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1981
  "\<lbrakk> rbt_sorted t1; rbt_sorted t2 \<rbrakk> \<Longrightarrow> is_rbt (rbt_inter t1 t2)"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1982
by(simp add: rbt_inter_def)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1983
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1984
lemma rbt_lookup_rbt_interwk:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1985
  "\<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
  1986
  \<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
  1987
  (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
  1988
   | 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
  1989
               | Some w \<Rightarrow> Some (f k v w))"
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1990
by(auto 4 3 simp add: rbt_inter_with_key_def Let_def map_of_entries[symmetric] rbt_lookup_rbtreeify map_map_filter split_def o_def option_map_comp map_filter_option_map_const sorted_filter[where f=id, simplified] rbt_sorted_entries distinct_entries map_of_sinter_with map_of_eq_None_iff set_map_filter split: option.split compare.split intro: rev_image_eqI dest: rbt_sorted_entries_right_unique)
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1991
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1992
lemma rbt_lookup_rbt_inter:
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1993
  "\<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
  1994
  \<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
  1995
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
  1996
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1997
end
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1998
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  1999
49480
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2000
subsection {* Code generator setup *}
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2001
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2002
lemmas [code] =
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2003
  ord.rbt_less_prop
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2004
  ord.rbt_greater_prop
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2005
  ord.rbt_sorted.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2006
  ord.rbt_lookup.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2007
  ord.is_rbt_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2008
  ord.rbt_ins.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2009
  ord.rbt_insert_with_key_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2010
  ord.rbt_insertw_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2011
  ord.rbt_insert_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2012
  ord.rbt_del_from_left.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2013
  ord.rbt_del_from_right.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2014
  ord.rbt_del.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2015
  ord.rbt_delete_def
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2016
  ord.sunion_with.simps
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2017
  ord.sinter_with.simps
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2018
  ord.rbt_union_with_key_def
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2019
  ord.rbt_union_with_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2020
  ord.rbt_union_def
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2021
  ord.rbt_inter_with_key_def
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2022
  ord.rbt_inter_with_def
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2023
  ord.rbt_inter_def
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2024
  ord.rbt_map_entry.simps
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2025
  ord.rbt_bulkload_def
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2026
49480
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2027
text {* More efficient implementations for @{term entries} and @{term keys} *}
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2028
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2029
definition gen_entries :: 
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2030
  "(('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
  2031
where
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2032
  "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
  2033
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2034
lemma gen_entries_simps [simp, code]:
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2035
  "gen_entries [] Empty = []"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2036
  "gen_entries ((kv, t) # kvts) Empty = kv # gen_entries kvts t"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2037
  "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
  2038
by(simp_all add: gen_entries_def)
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2039
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2040
lemma entries_code [code]:
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2041
  "entries = gen_entries []"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2042
by(simp add: gen_entries_def fun_eq_iff)
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2043
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2044
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
  2045
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
  2046
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2047
lemma gen_keys_simps [simp, code]:
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2048
  "gen_keys [] Empty = []"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2049
  "gen_keys ((k, t) # kts) Empty = k # gen_keys kts t"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2050
  "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
  2051
by(simp_all add: gen_keys_def)
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2052
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2053
lemma keys_code [code]:
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2054
  "keys = gen_keys []"
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2055
by(simp add: gen_keys_def fun_eq_iff)
4632b867fba7 more efficient code setup
Andreas Lochbihler
parents: 48621
diff changeset
  2056
47450
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2057
text {* Restore original type constraints for constants *}
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2058
setup {*
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2059
  fold Sign.add_const_constraint
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2060
    [(@{const_name rbt_less}, SOME @{typ "('a :: order) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2061
     (@{const_name rbt_greater}, SOME @{typ "('a :: order) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2062
     (@{const_name rbt_sorted}, SOME @{typ "('a :: linorder, 'b) rbt \<Rightarrow> bool"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2063
     (@{const_name rbt_lookup}, SOME @{typ "('a :: linorder, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2064
     (@{const_name is_rbt}, SOME @{typ "('a :: linorder, 'b) rbt \<Rightarrow> bool"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2065
     (@{const_name rbt_ins}, SOME @{typ "('a\<Colon>linorder \<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
  2066
     (@{const_name rbt_insert_with_key}, SOME @{typ "('a\<Colon>linorder \<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
  2067
     (@{const_name rbt_insert_with}, SOME @{typ "('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a :: linorder) \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2068
     (@{const_name rbt_insert}, SOME @{typ "('a :: linorder) \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2069
     (@{const_name rbt_del_from_left}, SOME @{typ "('a\<Colon>linorder) \<Rightarrow> ('a,'b) rbt \<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
  2070
     (@{const_name rbt_del_from_right}, SOME @{typ "('a\<Colon>linorder) \<Rightarrow> ('a,'b) rbt \<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
  2071
     (@{const_name rbt_del}, SOME @{typ "('a\<Colon>linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2072
     (@{const_name rbt_delete}, SOME @{typ "('a\<Colon>linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2073
     (@{const_name rbt_union_with_key}, SOME @{typ "('a\<Colon>linorder \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2074
     (@{const_name rbt_union_with}, SOME @{typ "('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2075
     (@{const_name rbt_union}, SOME @{typ "('a\<Colon>linorder,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2076
     (@{const_name rbt_map_entry}, SOME @{typ "'a\<Colon>linorder \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2077
     (@{const_name rbt_bulkload}, SOME @{typ "('a \<times> 'b) list \<Rightarrow> ('a\<Colon>linorder,'b) rbt"})]
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2078
*}
2ada2be850cb move RBT implementation into type class contexts
Andreas Lochbihler
parents: 47397
diff changeset
  2079
49770
cf6a78acf445 efficient construction of red black trees from sorted associative lists
Andreas Lochbihler
parents: 49480
diff changeset
  2080
hide_const (open) 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
  2081
52617dca8386 new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
diff changeset
  2082
end