merged
authornipkow
Thu, 05 Nov 2015 08:27:22 +0100
changeset 61582 69492d32263a
parent 61580 c49a8ebd30cc (current diff)
parent 61581 00d9682e8dd7 (diff)
child 61583 c2b7033fa0ba
merged
--- a/src/HOL/Data_Structures/AVL_Map.thy	Thu Nov 05 00:17:13 2015 +0100
+++ b/src/HOL/Data_Structures/AVL_Map.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -8,36 +8,34 @@
   Lookup2
 begin
 
-fun update :: "'a::order \<Rightarrow> 'b \<Rightarrow> ('a*'b) avl_tree \<Rightarrow> ('a*'b) avl_tree" where
+fun update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) avl_tree \<Rightarrow> ('a*'b) avl_tree" where
 "update x y Leaf = Node 1 Leaf (x,y) Leaf" |
-"update x y (Node h l (a,b) r) = 
-   (if x = a then Node h l (x,y) r else
-    if x < a then node_bal_l (update x y l) (a,b) r
-    else node_bal_r l (a,b) (update x y r))"
+"update x y (Node h l (a,b) r) = (case cmp x a of
+   EQ \<Rightarrow> Node h l (x,y) r |
+   LT \<Rightarrow> balL (update x y l) (a,b) r |
+   GT \<Rightarrow> balR l (a,b) (update x y r))"
 
-fun delete :: "'a::order \<Rightarrow> ('a*'b) avl_tree \<Rightarrow> ('a*'b) avl_tree" where
+fun delete :: "'a::cmp \<Rightarrow> ('a*'b) avl_tree \<Rightarrow> ('a*'b) avl_tree" where
 "delete _ Leaf = Leaf" |
-"delete x (Node h l (a,b) r) = (
-   if x = a then delete_root (Node h l (a,b) r) else
-   if x < a then node_bal_r (delete x l) (a,b) r
-   else node_bal_l l (a,b) (delete x r))"
+"delete x (Node h l (a,b) r) = (case cmp x a of
+   EQ \<Rightarrow> delete_root (Node h l (a,b) r) |
+   LT \<Rightarrow> balR (delete x l) (a,b) r |
+   GT \<Rightarrow> balL l (a,b) (delete x r))"
 
 
 subsection {* Functional Correctness Proofs *}
 
 theorem inorder_update:
   "sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
-by (induct t) 
-   (auto simp: upd_list_simps inorder_node_bal_l inorder_node_bal_r)
+by (induct t) (auto simp: upd_list_simps inorder_balL inorder_balR)
 
 
 theorem inorder_delete:
   "sorted1(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
 by(induction t)
-  (auto simp: del_list_simps inorder_node_bal_l inorder_node_bal_r
+  (auto simp: del_list_simps inorder_balL inorder_balR
      inorder_delete_root inorder_delete_maxD split: prod.splits)
 
-
 interpretation Map_by_Ordered
 where empty = Leaf and lookup = lookup and update = update and delete = delete
 and inorder = inorder and wf = "\<lambda>_. True"
--- a/src/HOL/Data_Structures/AVL_Set.thy	Thu Nov 05 00:17:13 2015 +0100
+++ b/src/HOL/Data_Structures/AVL_Set.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -6,7 +6,7 @@
 section "AVL Tree Implementation of Sets"
 
 theory AVL_Set
-imports Isin2
+imports Cmp Isin2
 begin
 
 type_synonym 'a avl_tree = "('a,nat) tree"
@@ -26,8 +26,8 @@
 definition node :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
 "node l a r = Node (max (ht l) (ht r) + 1) l a r"
 
-definition node_bal_l :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
-"node_bal_l l a r = (
+definition balL :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
+"balL l a r = (
   if ht l = ht r + 2 then (case l of 
     Node _ bl b br \<Rightarrow> (if ht bl < ht br
     then case br of
@@ -35,8 +35,8 @@
     else node bl b (node br a r)))
   else node l a r)"
 
-definition node_bal_r :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
-"node_bal_r l a r = (
+definition balR :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
+"balR l a r = (
   if ht r = ht l + 2 then (case r of
     Node _ bl b br \<Rightarrow> (if ht bl > ht br
     then case bl of
@@ -44,19 +44,17 @@
     else node (node l a bl) b br))
   else node l a r)"
 
-fun insert :: "'a::order \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
+fun insert :: "'a::cmp \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
 "insert x Leaf = Node 1 Leaf x Leaf" |
-"insert x (Node h l a r) = 
-   (if x=a then Node h l a r
-    else if x<a
-      then node_bal_l (insert x l) a r
-      else node_bal_r l a (insert x r))"
+"insert x (Node h l a r) = (case cmp x a of
+   EQ \<Rightarrow> Node h l a r |
+   LT \<Rightarrow> balL (insert x l) a r |
+   GT \<Rightarrow> balR l a (insert x r))"
 
 fun delete_max :: "'a avl_tree \<Rightarrow> 'a avl_tree * 'a" where
 "delete_max (Node _ l a Leaf) = (l,a)" |
-"delete_max (Node _ l a r) = (
-  let (r',a') = delete_max r in
-  (node_bal_l l a r', a'))"
+"delete_max (Node _ l a r) =
+  (let (r',a') = delete_max r in (balL l a r', a'))"
 
 lemmas delete_max_induct = delete_max.induct[case_names Leaf Node]
 
@@ -64,16 +62,16 @@
 "delete_root (Node h Leaf a r) = r" |
 "delete_root (Node h l a Leaf) = l" |
 "delete_root (Node h l a r) =
-  (let (l', a') = delete_max l in node_bal_r l' a' r)"
+  (let (l', a') = delete_max l in balR l' a' r)"
 
 lemmas delete_root_cases = delete_root.cases[case_names Leaf_t Node_Leaf Node_Node]
 
-fun delete :: "'a::order \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
+fun delete :: "'a::cmp \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
 "delete _ Leaf = Leaf" |
-"delete x (Node h l a r) = (
-   if x = a then delete_root (Node h l a r)
-   else if x < a then node_bal_r (delete x l) a r
-   else node_bal_l l a (delete x r))"
+"delete x (Node h l a r) = (case cmp x a of
+   EQ \<Rightarrow> delete_root (Node h l a r) |
+   LT \<Rightarrow> balR (delete x l) a r |
+   GT \<Rightarrow> balL l a (delete x r))"
 
 
 subsection {* Functional Correctness Proofs *}
@@ -83,18 +81,18 @@
 
 subsubsection "Proofs for insert"
 
-lemma inorder_node_bal_l:
-  "inorder (node_bal_l l a r) = inorder l @ a # inorder r"
-by (auto simp: node_def node_bal_l_def split:tree.splits)
+lemma inorder_balL:
+  "inorder (balL l a r) = inorder l @ a # inorder r"
+by (auto simp: node_def balL_def split:tree.splits)
 
-lemma inorder_node_bal_r:
-  "inorder (node_bal_r l a r) = inorder l @ a # inorder r"
-by (auto simp: node_def node_bal_r_def split:tree.splits)
+lemma inorder_balR:
+  "inorder (balR l a r) = inorder l @ a # inorder r"
+by (auto simp: node_def balR_def split:tree.splits)
 
 theorem inorder_insert:
   "sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)"
 by (induct t) 
-   (auto simp: ins_list_simps inorder_node_bal_l inorder_node_bal_r)
+   (auto simp: ins_list_simps inorder_balL inorder_balR)
 
 
 subsubsection "Proofs for delete"
@@ -103,17 +101,17 @@
   "\<lbrakk> delete_max t = (t',a); t \<noteq> Leaf \<rbrakk> \<Longrightarrow>
    inorder t' @ [a] = inorder t"
 by(induction t arbitrary: t' rule: delete_max.induct)
-  (auto simp: inorder_node_bal_l split: prod.splits tree.split)
+  (auto simp: inorder_balL split: prod.splits tree.split)
 
 lemma inorder_delete_root:
   "inorder (delete_root (Node h l a r)) = inorder l @ inorder r"
 by(induction "Node h l a r" arbitrary: l a r h rule: delete_root.induct)
-  (auto simp: inorder_node_bal_r inorder_delete_maxD split: prod.splits)
+  (auto simp: inorder_balR inorder_delete_maxD split: prod.splits)
 
 theorem inorder_delete:
   "sorted(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
 by(induction t)
-  (auto simp: del_list_simps inorder_node_bal_l inorder_node_bal_r
+  (auto simp: del_list_simps inorder_balL inorder_balR
     inorder_delete_root inorder_delete_maxD split: prod.splits)
 
 
@@ -145,17 +143,17 @@
 lemma [simp]: "avl t \<Longrightarrow> ht t = height t"
 by (induct t) simp_all
 
-lemma height_node_bal_l:
+lemma height_balL:
   "\<lbrakk> height l = height r + 2; avl l; avl r \<rbrakk> \<Longrightarrow>
-   height (node_bal_l l a r) = height r + 2 \<or>
-   height (node_bal_l l a r) = height r + 3"
-by (cases l) (auto simp:node_def node_bal_l_def split:tree.split)
+   height (balL l a r) = height r + 2 \<or>
+   height (balL l a r) = height r + 3"
+by (cases l) (auto simp:node_def balL_def split:tree.split)
        
-lemma height_node_bal_r:
+lemma height_balR:
   "\<lbrakk> height r = height l + 2; avl l; avl r \<rbrakk> \<Longrightarrow>
-   height (node_bal_r l a r) = height l + 2 \<or>
-   height (node_bal_r l a r) = height l + 3"
-by (cases r) (auto simp add:node_def node_bal_r_def split:tree.split)
+   height (balR l a r) = height l + 2 \<or>
+   height (balR l a r) = height l + 3"
+by (cases r) (auto simp add:node_def balR_def split:tree.split)
 
 lemma [simp]: "height(node l a r) = max (height l) (height r) + 1"
 by (simp add: node_def)
@@ -166,53 +164,53 @@
    \<rbrakk> \<Longrightarrow> avl(node l a r)"
 by (auto simp add:max_def node_def)
 
-lemma height_node_bal_l2:
+lemma height_balL2:
   "\<lbrakk> avl l; avl r; height l \<noteq> height r + 2 \<rbrakk> \<Longrightarrow>
-   height (node_bal_l l a r) = (1 + max (height l) (height r))"
-by (cases l, cases r) (simp_all add: node_bal_l_def)
+   height (balL l a r) = (1 + max (height l) (height r))"
+by (cases l, cases r) (simp_all add: balL_def)
 
-lemma height_node_bal_r2:
+lemma height_balR2:
   "\<lbrakk> avl l;  avl r;  height r \<noteq> height l + 2 \<rbrakk> \<Longrightarrow>
-   height (node_bal_r l a r) = (1 + max (height l) (height r))"
-by (cases l, cases r) (simp_all add: node_bal_r_def)
+   height (balR l a r) = (1 + max (height l) (height r))"
+by (cases l, cases r) (simp_all add: balR_def)
 
-lemma avl_node_bal_l: 
+lemma avl_balL: 
   assumes "avl l" "avl r" and "height l = height r \<or> height l = height r + 1
     \<or> height r = height l + 1 \<or> height l = height r + 2" 
-  shows "avl(node_bal_l l a r)"
+  shows "avl(balL l a r)"
 proof(cases l)
   case Leaf
-  with assms show ?thesis by (simp add: node_def node_bal_l_def)
+  with assms show ?thesis by (simp add: node_def balL_def)
 next
   case (Node ln ll lr lh)
   with assms show ?thesis
   proof(cases "height l = height r + 2")
     case True
     from True Node assms show ?thesis
-      by (auto simp: node_bal_l_def intro!: avl_node split: tree.split) arith+
+      by (auto simp: balL_def intro!: avl_node split: tree.split) arith+
   next
     case False
-    with assms show ?thesis by (simp add: avl_node node_bal_l_def)
+    with assms show ?thesis by (simp add: avl_node balL_def)
   qed
 qed
 
-lemma avl_node_bal_r: 
+lemma avl_balR: 
   assumes "avl l" and "avl r" and "height l = height r \<or> height l = height r + 1
     \<or> height r = height l + 1 \<or> height r = height l + 2" 
-  shows "avl(node_bal_r l a r)"
+  shows "avl(balR l a r)"
 proof(cases r)
   case Leaf
-  with assms show ?thesis by (simp add: node_def node_bal_r_def)
+  with assms show ?thesis by (simp add: node_def balR_def)
 next
   case (Node rn rl rr rh)
   with assms show ?thesis
   proof(cases "height r = height l + 2")
     case True
       from True Node assms show ?thesis
-        by (auto simp: node_bal_r_def intro!: avl_node split: tree.split) arith+
+        by (auto simp: balR_def intro!: avl_node split: tree.split) arith+
   next
     case False
-    with assms show ?thesis by (simp add: node_bal_r_def avl_node)
+    with assms show ?thesis by (simp add: balR_def avl_node)
   qed
 qed
 
@@ -237,10 +235,10 @@
     with Node 1 show ?thesis 
     proof(cases "x<a")
       case True
-      with Node 1 show ?thesis by (auto simp add:avl_node_bal_l)
+      with Node 1 show ?thesis by (auto simp add:avl_balL)
     next
       case False
-      with Node 1 `x\<noteq>a` show ?thesis by (auto simp add:avl_node_bal_r)
+      with Node 1 `x\<noteq>a` show ?thesis by (auto simp add:avl_balR)
     qed
   qed
   case 2
@@ -255,12 +253,12 @@
       case True
       with Node 2 show ?thesis
       proof(cases "height (insert x l) = height r + 2")
-        case False with Node 2 `x < a` show ?thesis by (auto simp: height_node_bal_l2)
+        case False with Node 2 `x < a` show ?thesis by (auto simp: height_balL2)
       next
         case True 
-        hence "(height (node_bal_l (insert x l) a r) = height r + 2) \<or>
-          (height (node_bal_l (insert x l) a r) = height r + 3)" (is "?A \<or> ?B")
-          using Node 2 by (intro height_node_bal_l) simp_all
+        hence "(height (balL (insert x l) a r) = height r + 2) \<or>
+          (height (balL (insert x l) a r) = height r + 3)" (is "?A \<or> ?B")
+          using Node 2 by (intro height_balL) simp_all
         thus ?thesis
         proof
           assume ?A
@@ -275,12 +273,12 @@
       with Node 2 show ?thesis 
       proof(cases "height (insert x r) = height l + 2")
         case False
-        with Node 2 `\<not>x < a` show ?thesis by (auto simp: height_node_bal_r2)
+        with Node 2 `\<not>x < a` show ?thesis by (auto simp: height_balR2)
       next
         case True 
-        hence "(height (node_bal_r l a (insert x r)) = height l + 2) \<or>
-          (height (node_bal_r l a (insert x r)) = height l + 3)"  (is "?A \<or> ?B")
-          using Node 2 by (intro height_node_bal_r) simp_all
+        hence "(height (balR l a (insert x r)) = height l + 2) \<or>
+          (height (balR l a (insert x r)) = height l + 3)"  (is "?A \<or> ?B")
+          using Node 2 by (intro height_balR) simp_all
         thus ?thesis 
         proof
           assume ?A
@@ -306,10 +304,10 @@
   case (Node h l a rh rl b rr)
   case 1
   with Node have "avl l" "avl (fst (delete_max (Node rh rl b rr)))" by auto
-  with 1 Node have "avl (node_bal_l l a (fst (delete_max (Node rh rl b rr))))"
-    by (intro avl_node_bal_l) fastforce+
+  with 1 Node have "avl (balL l a (fst (delete_max (Node rh rl b rr))))"
+    by (intro avl_balL) fastforce+
   thus ?case 
-    by (auto simp: height_node_bal_l height_node_bal_l2
+    by (auto simp: height_balL height_balL2
       linorder_class.max.absorb1 linorder_class.max.absorb2
       split:prod.split)
 next
@@ -318,7 +316,7 @@
   let ?r = "Node rh rl b rr"
   let ?r' = "fst (delete_max ?r)"
   from `avl x` Node 2 have "avl l" and "avl ?r" by simp_all
-  thus ?case using Node 2 height_node_bal_l[of l ?r' a] height_node_bal_l2[of l ?r' a]
+  thus ?case using Node 2 height_balL[of l ?r' a] height_balL2[of l ?r' a]
     apply (auto split:prod.splits simp del:avl.simps) by arith+
 qed auto
 
@@ -337,8 +335,8 @@
          height ?l = height(?l') + 1" by (rule avl_delete_max,simp)+
   with `avl t` Node_Node have "height ?l' = height ?r \<or> height ?l' = height ?r + 1
             \<or> height ?r = height ?l' + 1 \<or> height ?r = height ?l' + 2" by fastforce
-  with `avl ?l'` `avl ?r` have "avl(node_bal_r ?l' (snd(delete_max ?l)) ?r)"
-    by (rule avl_node_bal_r)
+  with `avl ?l'` `avl ?r` have "avl(balR ?l' (snd(delete_max ?l)) ?r)"
+    by (rule avl_balR)
   with Node_Node show ?thesis by (auto split:prod.splits)
 qed simp_all
 
@@ -351,7 +349,7 @@
   let ?l = "Node lh ll ln lr"
   let ?r = "Node rh rl rn rr"
   let ?l' = "fst (delete_max ?l)"
-  let ?t' = "node_bal_r ?l' (snd(delete_max ?l)) ?r"
+  let ?t' = "balR ?l' (snd(delete_max ?l)) ?r"
   from `avl t` and Node_Node have "avl ?r" by simp
   from `avl t` and Node_Node have "avl ?l" by simp
   hence "avl(?l')"  by (rule avl_delete_max,simp)
@@ -360,11 +358,11 @@
   have "height t = height ?t' \<or> height t = height ?t' + 1" using  `avl t` Node_Node
   proof(cases "height ?r = height ?l' + 2")
     case False
-    show ?thesis using l'_height t_height False by (subst  height_node_bal_r2[OF `avl ?l'` `avl ?r` False])+ arith
+    show ?thesis using l'_height t_height False by (subst  height_balR2[OF `avl ?l'` `avl ?r` False])+ arith
   next
     case True
     show ?thesis
-    proof(cases rule: disjE[OF height_node_bal_r[OF True `avl ?l'` `avl ?r`, of "snd (delete_max ?l)"]])
+    proof(cases rule: disjE[OF height_balR[OF True `avl ?l'` `avl ?r`, of "snd (delete_max ?l)"]])
       case 1
       thus ?thesis using l'_height t_height True by arith
     next
@@ -393,10 +391,10 @@
     with Node 1 show ?thesis 
     proof(cases "x<n")
       case True
-      with Node 1 show ?thesis by (auto simp add:avl_node_bal_r)
+      with Node 1 show ?thesis by (auto simp add:avl_balR)
     next
       case False
-      with Node 1 `x\<noteq>n` show ?thesis by (auto simp add:avl_node_bal_l)
+      with Node 1 `x\<noteq>n` show ?thesis by (auto simp add:avl_balL)
     qed
   qed
   case 2
@@ -414,38 +412,38 @@
       case True
       show ?thesis
       proof(cases "height r = height (delete x l) + 2")
-        case False with Node 1 `x < n` show ?thesis by(auto simp: node_bal_r_def)
+        case False with Node 1 `x < n` show ?thesis by(auto simp: balR_def)
       next
         case True 
-        hence "(height (node_bal_r (delete x l) n r) = height (delete x l) + 2) \<or>
-          height (node_bal_r (delete x l) n r) = height (delete x l) + 3" (is "?A \<or> ?B")
-          using Node 2 by (intro height_node_bal_r) auto
+        hence "(height (balR (delete x l) n r) = height (delete x l) + 2) \<or>
+          height (balR (delete x l) n r) = height (delete x l) + 3" (is "?A \<or> ?B")
+          using Node 2 by (intro height_balR) auto
         thus ?thesis 
         proof
           assume ?A
-          with `x < n` Node 2 show ?thesis by(auto simp: node_bal_r_def)
+          with `x < n` Node 2 show ?thesis by(auto simp: balR_def)
         next
           assume ?B
-          with `x < n` Node 2 show ?thesis by(auto simp: node_bal_r_def)
+          with `x < n` Node 2 show ?thesis by(auto simp: balR_def)
         qed
       qed
     next
       case False
       show ?thesis
       proof(cases "height l = height (delete x r) + 2")
-        case False with Node 1 `\<not>x < n` `x \<noteq> n` show ?thesis by(auto simp: node_bal_l_def)
+        case False with Node 1 `\<not>x < n` `x \<noteq> n` show ?thesis by(auto simp: balL_def)
       next
         case True 
-        hence "(height (node_bal_l l n (delete x r)) = height (delete x r) + 2) \<or>
-          height (node_bal_l l n (delete x r)) = height (delete x r) + 3" (is "?A \<or> ?B")
-          using Node 2 by (intro height_node_bal_l) auto
+        hence "(height (balL l n (delete x r)) = height (delete x r) + 2) \<or>
+          height (balL l n (delete x r)) = height (delete x r) + 3" (is "?A \<or> ?B")
+          using Node 2 by (intro height_balL) auto
         thus ?thesis 
         proof
           assume ?A
-          with `\<not>x < n` `x \<noteq> n` Node 2 show ?thesis by(auto simp: node_bal_l_def)
+          with `\<not>x < n` `x \<noteq> n` Node 2 show ?thesis by(auto simp: balL_def)
         next
           assume ?B
-          with `\<not>x < n` `x \<noteq> n` Node 2 show ?thesis by(auto simp: node_bal_l_def)
+          with `\<not>x < n` `x \<noteq> n` Node 2 show ?thesis by(auto simp: balL_def)
         qed
       qed
     qed
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Data_Structures/Cmp.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -0,0 +1,21 @@
+(* Author: Tobias Nipkow *)
+
+section {* Three-Way Comparison *}
+
+theory Cmp
+imports Main
+begin
+
+datatype cmp = LT | EQ | GT
+
+class cmp = linorder +
+fixes cmp :: "'a \<Rightarrow> 'a \<Rightarrow> cmp"
+assumes LT[simp]: "cmp x y = LT \<longleftrightarrow> x < y"
+assumes EQ[simp]: "cmp x y = EQ \<longleftrightarrow> x = y"
+assumes GT[simp]: "cmp x y = GT \<longleftrightarrow> x > y"
+
+lemma case_cmp_if[simp]: "(case c of EQ \<Rightarrow> e | LT \<Rightarrow> l | GT \<Rightarrow> g) =
+  (if c = LT then l else if c = GT then g else e)"
+by(simp split: cmp.split)
+
+end
--- a/src/HOL/Data_Structures/RBT_Map.thy	Thu Nov 05 00:17:13 2015 +0100
+++ b/src/HOL/Data_Structures/RBT_Map.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -8,25 +8,26 @@
   Lookup2
 begin
 
-fun update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) rbt \<Rightarrow> ('a*'b) rbt" where
+fun update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) rbt \<Rightarrow> ('a*'b) rbt" where
 "update x y Leaf = R Leaf (x,y) Leaf" |
-"update x y (B l (a,b) r) =
-  (if x < a then bal (update x y l) (a,b) r else
-   if x > a then bal l (a,b) (update x y r)
-   else B l (x,y) r)" |
-"update x y (R l (a,b) r) =
-  (if x < a then R (update x y l) (a,b) r else
-   if x > a then R l (a,b) (update x y r)
-   else R l (x,y) r)"
+"update x y (B l (a,b) r) = (case cmp x a of
+  LT \<Rightarrow> bal (update x y l) (a,b) r |
+  GT \<Rightarrow> bal l (a,b) (update x y r) |
+  EQ \<Rightarrow> B l (x,y) r)" |
+"update x y (R l (a,b) r) = (case cmp x a of
+  LT \<Rightarrow> R (update x y l) (a,b) r |
+  GT \<Rightarrow> R l (a,b) (update x y r) |
+  EQ \<Rightarrow> R l (x,y) r)"
 
-fun delete :: "'a::linorder \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
-and deleteL :: "'a::linorder \<Rightarrow> ('a*'b)rbt \<Rightarrow> 'a*'b \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
-and deleteR :: "'a::linorder \<Rightarrow> ('a*'b)rbt \<Rightarrow> 'a*'b \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
+fun delete :: "'a::cmp \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
+and deleteL :: "'a::cmp \<Rightarrow> ('a*'b)rbt \<Rightarrow> 'a*'b \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
+and deleteR :: "'a::cmp \<Rightarrow> ('a*'b)rbt \<Rightarrow> 'a*'b \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
 where
 "delete x Leaf = Leaf" |
-"delete x (Node c t1 (a,b) t2) = 
-  (if x < a then deleteL x t1 (a,b) t2 else
-   if x > a then deleteR x t1 (a,b) t2 else combine t1 t2)" |
+"delete x (Node c t1 (a,b) t2) = (case cmp x a of
+  LT \<Rightarrow> deleteL x t1 (a,b) t2 |
+  GT \<Rightarrow> deleteR x t1 (a,b) t2 |
+  EQ \<Rightarrow> combine t1 t2)" |
 "deleteL x (B t1 a t2) b t3 = balL (delete x (B t1 a t2)) b t3" |
 "deleteL x t1 a t2 = R (delete x t1) a t2" |
 "deleteR x t1 a (B t2 b t3) = balR t1 a (delete x (B t2 b t3))" | 
@@ -50,7 +51,6 @@
 by(induction x t1 and x t1 a t2 and x t1 a t2 rule: delete_deleteL_deleteR.induct)
   (auto simp: del_list_simps inorder_combine inorder_balL inorder_balR)
 
-
 interpretation Map_by_Ordered
 where empty = Leaf and lookup = lookup and update = update and delete = delete
 and inorder = inorder and wf = "\<lambda>_. True"
--- a/src/HOL/Data_Structures/RBT_Set.thy	Thu Nov 05 00:17:13 2015 +0100
+++ b/src/HOL/Data_Structures/RBT_Set.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -5,26 +5,30 @@
 theory RBT_Set
 imports
   RBT
+  Cmp
   Isin2
 begin
 
-fun insert :: "'a::linorder \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
+fun insert :: "'a::cmp \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
 "insert x Leaf = R Leaf x Leaf" |
-"insert x (B l a r) =
-  (if x < a then bal (insert x l) a r else
-   if x > a then bal l a (insert x r) else B l a r)" |
-"insert x (R l a r) =
-  (if x < a then R (insert x l) a r
-   else if x > a then R l a (insert x r) else R l a r)"
+"insert x (B l a r) = (case cmp x a of
+  LT \<Rightarrow> bal (insert x l) a r |
+  GT \<Rightarrow> bal l a (insert x r) |
+  EQ \<Rightarrow> B l a r)" |
+"insert x (R l a r) = (case cmp x a of
+  LT \<Rightarrow> R (insert x l) a r |
+  GT \<Rightarrow> R l a (insert x r) |
+  EQ \<Rightarrow> R l a r)"
 
-fun delete :: "'a::linorder \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"
-and deleteL :: "'a::linorder \<Rightarrow> 'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"
-and deleteR :: "'a::linorder \<Rightarrow> 'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"
+fun delete :: "'a::cmp \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"
+and deleteL :: "'a::cmp \<Rightarrow> 'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"
+and deleteR :: "'a::cmp \<Rightarrow> 'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"
 where
 "delete x Leaf = Leaf" |
-"delete x (Node _ l a r) = 
-  (if x < a then deleteL x l a r 
-   else if x > a then deleteR x l a r else combine l r)" |
+"delete x (Node _ l a r) = (case cmp x a of
+  LT \<Rightarrow> deleteL x l a r |
+  GT \<Rightarrow> deleteR x l a r |
+  EQ \<Rightarrow> combine l r)" |
 "deleteL x (B t1 a t2) b t3 = balL (delete x (B t1 a t2)) b t3" |
 "deleteL x l a r = R (delete x l) a r" |
 "deleteR x t1 a (B t2 b t3) = balR t1 a (delete x (B t2 b t3))" | 
@@ -66,6 +70,7 @@
 by(induction x t and x l a r and x l a r rule: delete_deleteL_deleteR.induct)
   (auto simp: del_list_simps inorder_combine inorder_balL inorder_balR)
 
+
 interpretation Set_by_Ordered
 where empty = Leaf and isin = isin and insert = insert and delete = delete
 and inorder = inorder and wf = "\<lambda>_. True"
--- a/src/HOL/Data_Structures/Splay_Map.thy	Thu Nov 05 00:17:13 2015 +0100
+++ b/src/HOL/Data_Structures/Splay_Map.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -42,35 +42,30 @@
 termination splay
 by lexicographic_order
 
-lemma splay_code: "splay x t = (case t of Leaf \<Rightarrow> Leaf |
-  Node al a ar \<Rightarrow>
-  (if x = fst a then t else
-   if x < fst a then
-     case al of
-       Leaf \<Rightarrow> t |
-       Node bl b br \<Rightarrow>
-         (if x = fst b then Node bl b (Node br a ar) else
-          if x < fst b then
-            if bl = Leaf then Node bl b (Node br a ar)
-            else case splay x bl of
-                   Node bll y blr \<Rightarrow> Node bll y (Node blr b (Node br a ar))
-          else
-          if br = Leaf then Node bl b (Node br a ar)
-          else case splay x br of
-                 Node brl y brr \<Rightarrow> Node (Node bl b brl) y (Node brr a ar))
-   else
-   case ar of
-     Leaf \<Rightarrow> t |
-     Node bl b br \<Rightarrow>
-       (if x = fst b then Node (Node al a bl) b br else
-        if x < fst b then
-          if bl = Leaf then Node (Node al a bl) b br
-          else case splay x bl of
-                 Node bll y blr \<Rightarrow> Node (Node al a bll) y (Node blr b br)
-        else if br=Leaf then Node (Node al a bl) b br
-             else case splay x br of
-                    Node bll y blr \<Rightarrow> Node (Node (Node al a bl) b bll) y blr)))"
-by(auto split: tree.split)
+lemma splay_code: "splay (x::_::cmp) t = (case t of Leaf \<Rightarrow> Leaf |
+  Node al a ar \<Rightarrow> (case cmp x (fst a) of
+    EQ \<Rightarrow> t |
+    LT \<Rightarrow> (case al of
+      Leaf \<Rightarrow> t |
+      Node bl b br \<Rightarrow> (case cmp x (fst b) of
+        EQ \<Rightarrow> Node bl b (Node br a ar) |
+        LT \<Rightarrow> if bl = Leaf then Node bl b (Node br a ar)
+              else case splay x bl of
+                Node bll y blr \<Rightarrow> Node bll y (Node blr b (Node br a ar)) |
+        GT \<Rightarrow> if br = Leaf then Node bl b (Node br a ar)
+              else case splay x br of
+                Node brl y brr \<Rightarrow> Node (Node bl b brl) y (Node brr a ar))) |
+    GT \<Rightarrow> (case ar of
+      Leaf \<Rightarrow> t |
+      Node bl b br \<Rightarrow> (case cmp x (fst b) of
+        EQ \<Rightarrow> Node (Node al a bl) b br |
+        LT \<Rightarrow> if bl = Leaf then Node (Node al a bl) b br
+              else case splay x bl of
+                Node bll y blr \<Rightarrow> Node (Node al a bll) y (Node blr b br) |
+        GT \<Rightarrow> if br=Leaf then Node (Node al a bl) b br
+              else case splay x br of
+                Node bll y blr \<Rightarrow> Node (Node (Node al a bl) b bll) y blr))))"
+by(auto cong: case_tree_cong split: tree.split)
 
 definition lookup :: "('a*'b)tree \<Rightarrow> 'a::linorder \<Rightarrow> 'b option" where "lookup t x =
   (case splay x t of Leaf \<Rightarrow> None | Node _ (a,b) _ \<Rightarrow> if x=a then Some b else None)"
--- a/src/HOL/Data_Structures/Splay_Set.thy	Thu Nov 05 00:17:13 2015 +0100
+++ b/src/HOL/Data_Structures/Splay_Set.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -1,6 +1,6 @@
 (*
 Author: Tobias Nipkow
-Function defs follows AFP entry Splay_Tree, proofs are new.
+Function defs follow AFP entry Splay_Tree, proofs are new.
 *)
 
 section "Splay Tree Implementation of Sets"
@@ -9,6 +9,7 @@
 imports
   "~~/src/HOL/Library/Tree"
   Set_by_Ordered
+  Cmp
 begin
 
 function splay :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
@@ -45,35 +46,35 @@
 termination splay
 by lexicographic_order
 
-lemma splay_code: "splay x t = (case t of Leaf \<Rightarrow> Leaf |
-  Node al a ar \<Rightarrow>
-  (if x=a then t else
-   if x < a then
-     case al of
-       Leaf \<Rightarrow> t |
-       Node bl b br \<Rightarrow>
-         (if x=b then Node bl b (Node br a ar) else
-          if x < b then
-            if bl = Leaf then Node bl b (Node br a ar)
-            else case splay x bl of
-                   Node bll y blr \<Rightarrow> Node bll y (Node blr b (Node br a ar))
-          else
-          if br = Leaf then Node bl b (Node br a ar)
-          else case splay x br of
-                 Node brl y brr \<Rightarrow> Node (Node bl b brl) y (Node brr a ar))
-   else
-   case ar of
-     Leaf \<Rightarrow> t |
-     Node bl b br \<Rightarrow>
-       (if x=b then Node (Node al a bl) b br else
-        if x < b then
-          if bl = Leaf then Node (Node al a bl) b br
-          else case splay x bl of
-                 Node bll y blr \<Rightarrow> Node (Node al a bll) y (Node blr b br)
-        else if br=Leaf then Node (Node al a bl) b br
-             else case splay x br of
-                    Node bll y blr \<Rightarrow> Node (Node (Node al a bl) b bll) y blr)))"
-by(auto split: tree.split)
+(* no idea why this speeds things up below *)
+lemma case_tree_cong:
+  "\<lbrakk> x = x'; y = y'; z = z' \<rbrakk> \<Longrightarrow> case_tree x y z = case_tree x' y' z'"
+by auto
+
+lemma splay_code: "splay (x::_::cmp) t = (case t of Leaf \<Rightarrow> Leaf |
+  Node al a ar \<Rightarrow> (case cmp x a of
+    EQ \<Rightarrow> t |
+    LT \<Rightarrow> (case al of
+      Leaf \<Rightarrow> t |
+      Node bl b br \<Rightarrow> (case cmp x b of
+        EQ \<Rightarrow> Node bl b (Node br a ar) |
+        LT \<Rightarrow> if bl = Leaf then Node bl b (Node br a ar)
+              else case splay x bl of
+                Node bll y blr \<Rightarrow> Node bll y (Node blr b (Node br a ar)) |
+        GT \<Rightarrow> if br = Leaf then Node bl b (Node br a ar)
+              else case splay x br of
+                Node brl y brr \<Rightarrow> Node (Node bl b brl) y (Node brr a ar))) |
+    GT \<Rightarrow> (case ar of
+      Leaf \<Rightarrow> t |
+      Node bl b br \<Rightarrow> (case cmp x b of
+        EQ \<Rightarrow> Node (Node al a bl) b br |
+        LT \<Rightarrow> if bl = Leaf then Node (Node al a bl) b br
+              else case splay x bl of
+                Node bll y blr \<Rightarrow> Node (Node al a bll) y (Node blr b br) |
+        GT \<Rightarrow> if br=Leaf then Node (Node al a bl) b br
+              else case splay x br of
+                Node bll y blr \<Rightarrow> Node (Node (Node al a bl) b bll) y blr))))"
+by(auto cong: case_tree_cong split: tree.split)
 
 definition is_root :: "'a \<Rightarrow> 'a tree \<Rightarrow> bool" where
 "is_root a t = (case t of Leaf \<Rightarrow> False | Node _ x _ \<Rightarrow> x = a)"
--- a/src/HOL/Data_Structures/Tree234_Map.thy	Thu Nov 05 00:17:13 2015 +0100
+++ b/src/HOL/Data_Structures/Tree234_Map.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -10,118 +10,105 @@
 
 subsection \<open>Map operations on 2-3-4 trees\<close>
 
-fun lookup :: "('a::linorder * 'b) tree234 \<Rightarrow> 'a \<Rightarrow> 'b option" where
+fun lookup :: "('a::cmp * 'b) tree234 \<Rightarrow> 'a \<Rightarrow> 'b option" where
 "lookup Leaf x = None" |
-"lookup (Node2 l (a,b) r) x =
-  (if x < a then lookup l x else
-  if a < x then lookup r x else Some b)" |
-"lookup (Node3 l (a1,b1) m (a2,b2) r) x =
-  (if x < a1 then lookup l x else
-   if x = a1 then Some b1 else
-   if x < a2 then lookup m x else
-   if x = a2 then Some b2
-   else lookup r x)" |
-"lookup (Node4 l (a1,b1) m (a2,b2) n (a3,b3) r) x =
-  (if x < a2 then
-     if x = a1 then Some b1 else
-     if x < a1 then lookup l x else lookup m x
-   else
-     if x = a2 then Some b2 else
-     if x = a3 then Some b3 else
-     if x < a3 then lookup n x
-     else lookup r x)"
+"lookup (Node2 l (a,b) r) x = (case cmp x a of
+  LT \<Rightarrow> lookup l x |
+  GT \<Rightarrow> lookup r x |
+  EQ \<Rightarrow> Some b)" |
+"lookup (Node3 l (a1,b1) m (a2,b2) r) x = (case cmp x a1 of
+  LT \<Rightarrow> lookup l x |
+  EQ \<Rightarrow> Some b1 |
+  GT \<Rightarrow> (case cmp x a2 of
+          LT \<Rightarrow> lookup m x |
+          EQ \<Rightarrow> Some b2 |
+          GT \<Rightarrow> lookup r x))" |
+"lookup (Node4 t1 (a1,b1) t2 (a2,b2) t3 (a3,b3) t4) x = (case cmp x a2 of
+  LT \<Rightarrow> (case cmp x a1 of
+           LT \<Rightarrow> lookup t1 x | EQ \<Rightarrow> Some b1 | GT \<Rightarrow> lookup t2 x) |
+  EQ \<Rightarrow> Some b2 |
+  GT \<Rightarrow> (case cmp x a3 of
+           LT \<Rightarrow> lookup t3 x | EQ \<Rightarrow> Some b3 | GT \<Rightarrow> lookup t4 x))"
 
-fun upd :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) up\<^sub>i" where
+fun upd :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) up\<^sub>i" where
 "upd x y Leaf = Up\<^sub>i Leaf (x,y) Leaf" |
-"upd x y (Node2 l ab r) =
-   (if x < fst ab then
-        (case upd x y l of
+"upd x y (Node2 l ab r) = (case cmp x (fst ab) of
+   LT \<Rightarrow> (case upd x y l of
            T\<^sub>i l' => T\<^sub>i (Node2 l' ab r)
-         | Up\<^sub>i l1 q l2 => T\<^sub>i (Node3 l1 q l2 ab r))
-    else if x = fst ab then T\<^sub>i (Node2 l (x,y) r)
-    else
-        (case upd x y r of
+         | Up\<^sub>i l1 ab' l2 => T\<^sub>i (Node3 l1 ab' l2 ab r)) |
+   EQ \<Rightarrow> T\<^sub>i (Node2 l (x,y) r) |
+   GT \<Rightarrow> (case upd x y r of
            T\<^sub>i r' => T\<^sub>i (Node2 l ab r')
-         | Up\<^sub>i r1 q r2 => T\<^sub>i (Node3 l ab r1 q r2)))" |
-"upd x y (Node3 l ab1 m ab2 r) =
-   (if x < fst ab1 then
-        (case upd x y l of
+         | Up\<^sub>i r1 ab' r2 => T\<^sub>i (Node3 l ab r1 ab' r2)))" |
+"upd x y (Node3 l ab1 m ab2 r) = (case cmp x (fst ab1) of
+   LT \<Rightarrow> (case upd x y l of
            T\<^sub>i l' => T\<^sub>i (Node3 l' ab1 m ab2 r)
-         | Up\<^sub>i l1 q l2 => Up\<^sub>i (Node2 l1 q l2) ab1 (Node2 m ab2 r))
-    else if x = fst ab1 then T\<^sub>i (Node3 l (x,y) m ab2 r)
-    else if x < fst ab2 then
-             (case upd x y m of
-                T\<^sub>i m' => T\<^sub>i (Node3 l ab1 m' ab2 r)
-              | Up\<^sub>i m1 q m2 => Up\<^sub>i (Node2 l ab1 m1) q (Node2 m2 ab2 r))
-         else if x = fst ab2 then T\<^sub>i (Node3 l ab1 m (x,y) r)
-         else
-             (case upd x y r of
-                T\<^sub>i r' => T\<^sub>i (Node3 l ab1 m ab2 r')
-              | Up\<^sub>i r1 q r2 => Up\<^sub>i (Node2 l ab1 m) ab2 (Node2 r1 q r2)))" |
-"upd x y (Node4 l ab1 m ab2 n ab3 r) =
-   (if x < fst ab2 then
-      if x < fst ab1 then
-        (case upd x y l of
-           T\<^sub>i l' => T\<^sub>i (Node4 l' ab1 m ab2 n ab3 r)
-         | Up\<^sub>i l1 q l2 => Up\<^sub>i (Node2 l1 q l2) ab1 (Node3 m ab2 n ab3 r))
-      else
-      if x = fst ab1 then T\<^sub>i (Node4 l (x,y) m ab2 n ab3 r)
-      else
-        (case upd x y m of
-           T\<^sub>i m' => T\<^sub>i (Node4 l ab1 m' ab2 n ab3 r)
-         | Up\<^sub>i m1 q m2 => Up\<^sub>i (Node2 l ab1 m1) q (Node3 m2 ab2 n ab3 r))
-    else
-    if x = fst ab2 then T\<^sub>i (Node4 l ab1 m (x,y) n ab3 r) else
-    if x < fst ab3 then
-      (case upd x y n of
-         T\<^sub>i n' => T\<^sub>i (Node4 l ab1 m ab2 n' ab3 r)
-       | Up\<^sub>i n1 q n2 => Up\<^sub>i (Node2 l ab1 m) ab2(*q*) (Node3 n1 q n2 ab3 r))
-    else
-    if x = fst ab3 then T\<^sub>i (Node4 l ab1 m ab2 n (x,y) r)
-    else
-      (case upd x y r of
-         T\<^sub>i r' => T\<^sub>i (Node4 l ab1 m ab2 n ab3 r')
-       | Up\<^sub>i r1 q r2 => Up\<^sub>i (Node2 l ab1 m) ab2 (Node3 n ab3 r1 q r2)))"
+         | Up\<^sub>i l1 ab' l2 => Up\<^sub>i (Node2 l1 ab' l2) ab1 (Node2 m ab2 r)) |
+   EQ \<Rightarrow> T\<^sub>i (Node3 l (x,y) m ab2 r) |
+   GT \<Rightarrow> (case cmp x (fst ab2) of
+           LT \<Rightarrow> (case upd x y m of
+                   T\<^sub>i m' => T\<^sub>i (Node3 l ab1 m' ab2 r)
+                 | Up\<^sub>i m1 ab' m2 => Up\<^sub>i (Node2 l ab1 m1) ab' (Node2 m2 ab2 r)) |
+           EQ \<Rightarrow> T\<^sub>i (Node3 l ab1 m (x,y) r) |
+           GT \<Rightarrow> (case upd x y r of
+                   T\<^sub>i r' => T\<^sub>i (Node3 l ab1 m ab2 r')
+                 | Up\<^sub>i r1 ab' r2 => Up\<^sub>i (Node2 l ab1 m) ab2 (Node2 r1 ab' r2))))" |
+"upd x y (Node4 t1 ab1 t2 ab2 t3 ab3 t4) = (case cmp x (fst ab2) of
+   LT \<Rightarrow> (case cmp x (fst ab1) of
+            LT \<Rightarrow> (case upd x y t1 of
+                     T\<^sub>i t1' => T\<^sub>i (Node4 t1' ab1 t2 ab2 t3 ab3 t4)
+                  | Up\<^sub>i t11 q t12 => Up\<^sub>i (Node2 t11 q t12) ab1 (Node3 t2 ab2 t3 ab3 t4)) |
+            EQ \<Rightarrow> T\<^sub>i (Node4 t1 (x,y) t2 ab2 t3 ab3 t4) |
+            GT \<Rightarrow> (case upd x y t2 of
+                    T\<^sub>i t2' => T\<^sub>i (Node4 t1 ab1 t2' ab2 t3 ab3 t4)
+                  | Up\<^sub>i t21 q t22 => Up\<^sub>i (Node2 t1 ab1 t21) q (Node3 t22 ab2 t3 ab3 t4))) |
+   EQ \<Rightarrow> T\<^sub>i (Node4 t1 ab1 t2 (x,y) t3 ab3 t4) |
+   GT \<Rightarrow> (case cmp x (fst ab3) of
+            LT \<Rightarrow> (case upd x y t3 of
+                    T\<^sub>i t3' \<Rightarrow> T\<^sub>i (Node4 t1 ab1 t2 ab2 t3' ab3 t4)
+                  | Up\<^sub>i t31 q t32 => Up\<^sub>i (Node2 t1 ab1 t2) ab2(*q*) (Node3 t31 q t32 ab3 t4)) |
+            EQ \<Rightarrow> T\<^sub>i (Node4 t1 ab1 t2 ab2 t3 (x,y) t4) |
+            GT \<Rightarrow> (case upd x y t4 of
+                    T\<^sub>i t4' => T\<^sub>i (Node4 t1 ab1 t2 ab2 t3 ab3 t4')
+                  | Up\<^sub>i t41 q t42 => Up\<^sub>i (Node2 t1 ab1 t2) ab2 (Node3 t3 ab3 t41 q t42))))"
+
+definition update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) tree234" where
+"update x y t = tree\<^sub>i(upd x y t)"
 
-definition update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) tree234" where
-"update a b t = tree\<^sub>i(upd a b t)"
-
-fun del :: "'a::linorder \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) up\<^sub>d"
-where
-"del k Leaf = T\<^sub>d Leaf" |
-"del k (Node2 Leaf p Leaf) = (if k=fst p then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf p Leaf))" |
-"del k (Node3 Leaf p Leaf q Leaf) =
-  T\<^sub>d(if k=fst p then Node2 Leaf q Leaf else
-     if k=fst q then Node2 Leaf p Leaf
-     else Node3 Leaf p Leaf q Leaf)" |
-"del k (Node4 Leaf ab1 Leaf ab2 Leaf ab3 Leaf) =
-  T\<^sub>d(if k=fst ab1 then Node3 Leaf ab2 Leaf ab3 Leaf else
-     if k=fst ab2 then Node3 Leaf ab1 Leaf ab3 Leaf else
-     if k=fst ab3 then Node3 Leaf ab1 Leaf ab2 Leaf
+fun del :: "'a::cmp \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) up\<^sub>d" where
+"del x Leaf = T\<^sub>d Leaf" |
+"del x (Node2 Leaf ab1 Leaf) = (if x=fst ab1 then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf ab1 Leaf))" |
+"del x (Node3 Leaf ab1 Leaf ab2 Leaf) = T\<^sub>d(if x=fst ab1 then Node2 Leaf ab2 Leaf
+  else if x=fst ab2 then Node2 Leaf ab1 Leaf else Node3 Leaf ab1 Leaf ab2 Leaf)" |
+"del x (Node4 Leaf ab1 Leaf ab2 Leaf ab3 Leaf) =
+  T\<^sub>d(if x = fst ab1 then Node3 Leaf ab2 Leaf ab3 Leaf else
+     if x = fst ab2 then Node3 Leaf ab1 Leaf ab3 Leaf else
+     if x = fst ab3 then Node3 Leaf ab1 Leaf ab2 Leaf
      else Node4 Leaf ab1 Leaf ab2 Leaf ab3 Leaf)" |
-"del k (Node2 l a r) =
-  (if k<fst a then node21 (del k l) a r else
-   if k > fst a then node22 l a (del k r)
-   else let (a',t) = del_min r in node22 l a' t)" |
-"del k (Node3 l a m b r) =
-  (if k<fst a then node31 (del k l) a m b r else
-   if k = fst a then let (a',m') = del_min m in node32 l a' m' b r else
-   if k < fst b then node32 l a (del k m) b r else
-   if k = fst b then let (b',r') = del_min r in node33 l a m b' r'
-   else node33 l a m b (del k r))" |
-"del x (Node4 l ab1 m ab2 n ab3 r) =
-  (if x < fst ab2 then
-     if x < fst ab1 then node41 (del x l) ab1 m ab2 n ab3 r else
-     if x = fst ab1 then let (ab',m') = del_min m in node42 l ab' m' ab2 n ab3 r
-     else node42 l ab1 (del x m) ab2 n ab3 r
-   else
-     if x = fst ab2 then let (ab',n') = del_min n in node43 l ab1 m ab' n' ab3 r else
-     if x < fst ab3 then node43 l ab1 m ab2 (del x n) ab3 r else
-     if x = fst ab3 then let (ab',r') = del_min r in node44 l ab1 m ab2 n ab' r'
-     else node44 l ab1 m ab2 n ab3 (del x r))"
+"del x (Node2 l ab1 r) = (case cmp x (fst ab1) of
+  LT \<Rightarrow> node21 (del x l) ab1 r |
+  GT \<Rightarrow> node22 l ab1 (del x r) |
+  EQ \<Rightarrow> let (ab1',t) = del_min r in node22 l ab1' t)" |
+"del x (Node3 l ab1 m ab2 r) = (case cmp x (fst ab1) of
+  LT \<Rightarrow> node31 (del x l) ab1 m ab2 r |
+  EQ \<Rightarrow> let (ab1',m') = del_min m in node32 l ab1' m' ab2 r |
+  GT \<Rightarrow> (case cmp x (fst ab2) of
+           LT \<Rightarrow> node32 l ab1 (del x m) ab2 r |
+           EQ \<Rightarrow> let (ab2',r') = del_min r in node33 l ab1 m ab2' r' |
+           GT \<Rightarrow> node33 l ab1 m ab2 (del x r)))" |
+"del x (Node4 t1 ab1 t2 ab2 t3 ab3 t4) = (case cmp x (fst ab2) of
+  LT \<Rightarrow> (case cmp x (fst ab1) of
+           LT \<Rightarrow> node41 (del x t1) ab1 t2 ab2 t3 ab3 t4 |
+           EQ \<Rightarrow> let (ab',t2') = del_min t2 in node42 t1 ab' t2' ab2 t3 ab3 t4 |
+           GT \<Rightarrow> node42 t1 ab1 (del x t2) ab2 t3 ab3 t4) |
+  EQ \<Rightarrow> let (ab',t3') = del_min t3 in node43 t1 ab1 t2 ab' t3' ab3 t4 |
+  GT \<Rightarrow> (case cmp x (fst ab3) of
+          LT \<Rightarrow> node43 t1 ab1 t2 ab2 (del x t3) ab3 t4 |
+          EQ \<Rightarrow> let (ab',t4') = del_min t4 in node44 t1 ab1 t2 ab2 t3 ab' t4' |
+          GT \<Rightarrow> node44 t1 ab1 t2 ab2 t3 ab3 (del x t4)))"
 
-definition delete :: "'a::linorder \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) tree234" where
-"delete k t = tree\<^sub>d(del k t)"
+definition delete :: "'a::cmp \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) tree234" where
+"delete x t = tree\<^sub>d(del x t)"
 
 
 subsection "Functional correctness"
@@ -144,7 +131,7 @@
   inorder(tree\<^sub>d (del x t)) = del_list x (inorder t)"
 by(induction t rule: del.induct)
   ((auto simp: del_list_simps inorder_nodes del_minD split: prod.splits)[1])+
-(* 290 secs (2015) *)
+(* 200 secs (2015) *)
 
 lemma inorder_delete: "\<lbrakk> bal t ; sorted1(inorder t) \<rbrakk> \<Longrightarrow>
   inorder(delete x t) = del_list x (inorder t)"
@@ -154,7 +141,7 @@
 subsection \<open>Balancedness\<close>
 
 lemma bal_upd: "bal t \<Longrightarrow> bal (tree\<^sub>i(upd x y t)) \<and> height(upd x y t) = height t"
-by (induct t) (auto, auto split: up\<^sub>i.split) (* 33 secs (2015) *)
+by (induct t) (auto, auto split: up\<^sub>i.split) (* 20 secs (2015) *)
 
 lemma bal_update: "bal t \<Longrightarrow> bal (update x y t)"
 by (simp add: update_def bal_upd)
@@ -163,11 +150,12 @@
 lemma height_del: "bal t \<Longrightarrow> height(del x t) = height t"
 by(induction x t rule: del.induct)
   (auto simp add: heights height_del_min split: prod.split)
+(* 20 secs (2015) *)
 
 lemma bal_tree\<^sub>d_del: "bal t \<Longrightarrow> bal(tree\<^sub>d(del x t))"
 by(induction x t rule: del.induct)
   (auto simp: bals bal_del_min height_del height_del_min split: prod.split)
-(* 110 secs (2015) *)
+(* 100 secs (2015) *)
 
 corollary bal_delete: "bal t \<Longrightarrow> bal(delete x t)"
 by(simp add: delete_def bal_tree\<^sub>d_del)
--- a/src/HOL/Data_Structures/Tree234_Set.thy	Thu Nov 05 00:17:13 2015 +0100
+++ b/src/HOL/Data_Structures/Tree234_Set.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -5,19 +5,29 @@
 theory Tree234_Set
 imports
   Tree234
+  Cmp
   "../Data_Structures/Set_by_Ordered"
 begin
 
 subsection \<open>Set operations on 2-3-4 trees\<close>
 
-fun isin :: "'a::linorder tree234 \<Rightarrow> 'a \<Rightarrow> bool" where
+fun isin :: "'a::cmp tree234 \<Rightarrow> 'a \<Rightarrow> bool" where
 "isin Leaf x = False" |
-"isin (Node2 l a r) x = (x < a \<and> isin l x \<or> x=a \<or> isin r x)" |
+"isin (Node2 l a r) x =
+  (case cmp x a of LT \<Rightarrow> isin l x | EQ \<Rightarrow> True | GT \<Rightarrow> isin r x)" |
 "isin (Node3 l a m b r) x =
-  (x < a \<and> isin l x \<or> x = a \<or> x < b \<and> isin m x \<or> x = b \<or> isin r x)" |
-"isin (Node4 l a m b n c r) x =
-  (x < b \<and> (x < a \<and> isin l x \<or> x = a \<or> isin m x) \<or> x = b \<or>
-   x > b \<and> (x < c \<and> isin n x \<or> x=c \<or> isin r x))"
+  (case cmp x a of LT \<Rightarrow> isin l x | EQ \<Rightarrow> True | GT \<Rightarrow> (case cmp x b of
+   LT \<Rightarrow> isin m x | EQ \<Rightarrow> True | GT \<Rightarrow> isin r x))" |
+"isin (Node4 t1 a t2 b t3 c t4) x = (case cmp x b of
+  LT \<Rightarrow> (case cmp x a of
+          LT \<Rightarrow> isin t1 x |
+          EQ \<Rightarrow> True |
+          GT \<Rightarrow> isin t2 x) |
+  EQ \<Rightarrow> True |
+  GT \<Rightarrow> (case cmp x c of
+          LT \<Rightarrow> isin t3 x |
+          EQ \<Rightarrow> True |
+          GT \<Rightarrow> isin t4 x))"
 
 datatype 'a up\<^sub>i = T\<^sub>i "'a tree234" | Up\<^sub>i "'a tree234" 'a "'a tree234"
 
@@ -25,33 +35,31 @@
 "tree\<^sub>i (T\<^sub>i t) = t" |
 "tree\<^sub>i (Up\<^sub>i l p r) = Node2 l p r"
 
-fun ins :: "'a::linorder \<Rightarrow> 'a tree234 \<Rightarrow> 'a up\<^sub>i" where
-"ins a Leaf = Up\<^sub>i Leaf a Leaf" |
-"ins a (Node2 l x r) =
-   (if a < x then
-        (case ins a l of
-           T\<^sub>i l' => T\<^sub>i (Node2 l' x r)
-         | Up\<^sub>i l1 q l2 => T\<^sub>i (Node3 l1 q l2 x r))
-    else if a=x then T\<^sub>i (Node2 l x r)
-    else
-        (case ins a r of
-           T\<^sub>i r' => T\<^sub>i (Node2 l x r')
-         | Up\<^sub>i r1 q r2 => T\<^sub>i (Node3 l x r1 q r2)))" |
-"ins a (Node3 l x1 m x2 r) =
-   (if a < x1 then
-        (case ins a l of
-           T\<^sub>i l' => T\<^sub>i (Node3 l' x1 m x2 r)
-         | Up\<^sub>i l1 q l2 => T\<^sub>i (Node4 l1 q l2 x1 m x2 r))
-    else if a=x1 then T\<^sub>i (Node3 l x1 m x2 r)
-    else if a < x2 then
-             (case ins a m of
-                T\<^sub>i m' => T\<^sub>i (Node3 l x1 m' x2 r)
-              | Up\<^sub>i m1 q m2 => T\<^sub>i (Node4 l x1 m1 q m2 x2 r))
-         else if a=x2 then T\<^sub>i (Node3 l x1 m x2 r)
-         else
-             (case ins a r of
-                T\<^sub>i r' => T\<^sub>i (Node3 l x1 m x2 r')
-              | Up\<^sub>i r1 q r2 => T\<^sub>i (Node4 l x1 m x2 r1 q r2)))" |
+fun ins :: "'a::cmp \<Rightarrow> 'a tree234 \<Rightarrow> 'a up\<^sub>i" where
+"ins x Leaf = Up\<^sub>i Leaf x Leaf" |
+"ins x (Node2 l a r) =
+   (case cmp x a of
+      LT \<Rightarrow> (case ins x l of
+              T\<^sub>i l' => T\<^sub>i (Node2 l' a r)
+            | Up\<^sub>i l1 b l2 => T\<^sub>i (Node3 l1 b l2 a r)) |
+      EQ \<Rightarrow> T\<^sub>i (Node2 l x r) |
+      GT \<Rightarrow> (case ins x r of
+              T\<^sub>i r' => T\<^sub>i (Node2 l a r')
+            | Up\<^sub>i r1 b r2 => T\<^sub>i (Node3 l a r1 b r2)))" |
+"ins x (Node3 l a m b r) =
+   (case cmp x a of
+      LT \<Rightarrow> (case ins x l of
+              T\<^sub>i l' => T\<^sub>i (Node3 l' a m b r)
+            | Up\<^sub>i l1 c l2 => Up\<^sub>i (Node2 l1 c l2) a (Node2 m b r)) |
+      EQ \<Rightarrow> T\<^sub>i (Node3 l a m b r) |
+      GT \<Rightarrow> (case cmp x b of
+               GT \<Rightarrow> (case ins x r of
+                       T\<^sub>i r' => T\<^sub>i (Node3 l a m b r')
+                     | Up\<^sub>i r1 c r2 => Up\<^sub>i (Node2 l a m) b (Node2 r1 c r2)) |
+               EQ \<Rightarrow> T\<^sub>i (Node3 l a m b r) |
+               LT \<Rightarrow> (case ins x m of
+                       T\<^sub>i m' => T\<^sub>i (Node3 l a m' b r)
+                     | Up\<^sub>i m1 c m2 => Up\<^sub>i (Node2 l a m1) c (Node2 m2 b r))))" |
 "ins a (Node4 l x1 m x2 n x3 r) =
    (if a < x2 then
       if a < x1 then
@@ -75,8 +83,8 @@
 
 hide_const insert
 
-definition insert :: "'a::linorder \<Rightarrow> 'a tree234 \<Rightarrow> 'a tree234" where
-"insert a t = tree\<^sub>i(ins a t)"
+definition insert :: "'a::cmp \<Rightarrow> 'a tree234 \<Rightarrow> 'a tree234" where
+"insert x t = tree\<^sub>i(ins x t)"
 
 datatype 'a up\<^sub>d = T\<^sub>d "'a tree234" | Up\<^sub>d "'a tree234"
 
@@ -146,7 +154,7 @@
 "del_min (Node3 l a m b r) = (let (x,l') = del_min l in (x, node31 l' a m b r))" |
 "del_min (Node4 l a m b n c r) = (let (x,l') = del_min l in (x, node41 l' a m b n c r))"
 
-fun del :: "'a::linorder \<Rightarrow> 'a tree234 \<Rightarrow> 'a up\<^sub>d" where
+fun del :: "'a::cmp \<Rightarrow> 'a tree234 \<Rightarrow> 'a up\<^sub>d" where
 "del k Leaf = T\<^sub>d Leaf" |
 "del k (Node2 Leaf p Leaf) = (if k=p then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf p Leaf))" |
 "del k (Node3 Leaf p Leaf q Leaf) = T\<^sub>d(if k=p then Node2 Leaf q Leaf
@@ -156,36 +164,38 @@
      if k=b then Node3 Leaf a Leaf c Leaf else
      if k=c then Node3 Leaf a Leaf b Leaf
      else Node4 Leaf a Leaf b Leaf c Leaf)" |
-"del k (Node2 l a r) = (if k<a then node21 (del k l) a r else
-  if k > a then node22 l a (del k r) else
-  let (a',t) = del_min r in node22 l a' t)" |
-"del k (Node3 l a m b r) = (if k<a then node31 (del k l) a m b r else
-  if k = a then let (a',m') = del_min m in node32 l a' m' b r else
-  if k < b then node32 l a (del k m) b r else
-  if k = b then let (b',r') = del_min r in node33 l a m b' r'
-  else node33 l a m b (del k r))" |
-"del k (Node4 l a m b n c r) =
-  (if k < b then
-     if k < a then node41 (del k l) a m b n c r else
-     if k = a then let (a',m') = del_min m in node42 l a' m' b n c r
-     else node42 l a (del k m) b n c r
-   else
-     if k = b then let (b',n') = del_min n in node43 l a m b' n' c r else
-     if k < c then node43 l a m b (del k n) c r else
-     if k = c then let (c',r') = del_min r in node44 l a m b n c' r'
-     else node44 l a m b n c (del k r))"
+"del k (Node2 l a r) = (case cmp k a of
+  LT \<Rightarrow> node21 (del k l) a r |
+  GT \<Rightarrow> node22 l a (del k r) |
+  EQ \<Rightarrow> let (a',t) = del_min r in node22 l a' t)" |
+"del k (Node3 l a m b r) = (case cmp k a of
+  LT \<Rightarrow> node31 (del k l) a m b r |
+  EQ \<Rightarrow> let (a',m') = del_min m in node32 l a' m' b r |
+  GT \<Rightarrow> (case cmp k b of
+           LT \<Rightarrow> node32 l a (del k m) b r |
+           EQ \<Rightarrow> let (b',r') = del_min r in node33 l a m b' r' |
+           GT \<Rightarrow> node33 l a m b (del k r)))" |
+"del k (Node4 l a m b n c r) = (case cmp k b of
+  LT \<Rightarrow> (case cmp k a of
+          LT \<Rightarrow> node41 (del k l) a m b n c r |
+          EQ \<Rightarrow> let (a',m') = del_min m in node42 l a' m' b n c r |
+          GT \<Rightarrow> node42 l a (del k m) b n c r) |
+  EQ \<Rightarrow> let (b',n') = del_min n in node43 l a m b' n' c r |
+  GT \<Rightarrow> (case cmp k c of
+           LT \<Rightarrow> node43 l a m b (del k n) c r |
+           EQ \<Rightarrow> let (c',r') = del_min r in node44 l a m b n c' r' |
+           GT \<Rightarrow> node44 l a m b n c (del k r)))"
 
-definition delete :: "'a::linorder \<Rightarrow> 'a tree234 \<Rightarrow> 'a tree234" where
-"delete k t = tree\<^sub>d(del k t)"
+definition delete :: "'a::cmp \<Rightarrow> 'a tree234 \<Rightarrow> 'a tree234" where
+"delete x t = tree\<^sub>d(del x t)"
 
 
 subsection "Functional correctness"
 
-
 subsubsection \<open>Functional correctness of isin:\<close>
 
 lemma "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
-by (induction t) (auto simp: elems_simps1)
+by (induction t) (auto simp: elems_simps1 ball_Un)
 
 lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
 by (induction t) (auto simp: elems_simps2)
@@ -252,12 +262,9 @@
 
 lemma inorder_del: "\<lbrakk> bal t ; sorted(inorder t) \<rbrakk> \<Longrightarrow>
   inorder(tree\<^sub>d (del x t)) = del_list x (inorder t)"
-apply(induction t rule: del.induct)
-apply(simp_all add: del_list_simps inorder_nodes)
-apply(auto simp: del_list_simps;
-      auto simp: inorder_nodes del_list_simps del_minD split: prod.splits)+
-(* takes 285 s (2015); the last line alone would do it but takes hours *)
-done
+by(induction t rule: del.induct)
+  (auto simp: inorder_nodes del_list_simps del_minD split: prod.splits)
+  (* 150 secs (2015) *)
 
 lemma inorder_delete: "\<lbrakk> bal t ; sorted(inorder t) \<rbrakk> \<Longrightarrow>
   inorder(delete x t) = del_list x (inorder t)"
@@ -282,7 +289,7 @@
 end
 
 lemma bal_ins: "bal t \<Longrightarrow> bal (tree\<^sub>i(ins a t)) \<and> height(ins a t) = height t"
-by (induct t) (auto, auto split: up\<^sub>i.split) (* 29 secs (2015) *)
+by (induct t) (auto, auto split: up\<^sub>i.split) (* 20 secs (2015) *)
 
 
 text{* Now an alternative proof (by Brian Huffman) that runs faster because
@@ -344,9 +351,7 @@
 "full\<^sub>i n (Up\<^sub>i l p r) \<longleftrightarrow> full n l \<and> full n r"
 
 lemma full\<^sub>i_ins: "full n t \<Longrightarrow> full\<^sub>i n (ins a t)"
-apply (induct rule: full.induct)
-apply (auto, auto split: up\<^sub>i.split)
-done
+by (induct rule: full.induct) (auto, auto split: up\<^sub>i.split)
 
 text {* The @{const insert} operation preserves balance. *}
 
@@ -482,13 +487,12 @@
 
 lemma bal_tree\<^sub>d_del: "bal t \<Longrightarrow> bal(tree\<^sub>d(del x t))"
 by(induction x t rule: del.induct)
-  ((auto simp: bals bal_del_min height_del height_del_min split: prod.split)[1])+
-(* 64 secs (2015) *)
+  (auto simp: bals bal_del_min height_del height_del_min split: prod.split)
+(* 60 secs (2015) *)
 
 corollary bal_delete: "bal t \<Longrightarrow> bal(delete x t)"
 by(simp add: delete_def bal_tree\<^sub>d_del)
 
-
 subsection \<open>Overall Correctness\<close>
 
 interpretation Set_by_Ordered
--- a/src/HOL/Data_Structures/Tree23_Map.thy	Thu Nov 05 00:17:13 2015 +0100
+++ b/src/HOL/Data_Structures/Tree23_Map.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -8,65 +8,65 @@
   Map_by_Ordered
 begin
 
-fun lookup :: "('a::linorder * 'b) tree23 \<Rightarrow> 'a \<Rightarrow> 'b option" where
+fun lookup :: "('a::cmp * 'b) tree23 \<Rightarrow> 'a \<Rightarrow> 'b option" where
 "lookup Leaf x = None" |
-"lookup (Node2 l (a,b) r) x =
-  (if x < a then lookup l x else
-  if a < x then lookup r x else Some b)" |
-"lookup (Node3 l (a1,b1) m (a2,b2) r) x =
-  (if x < a1 then lookup l x else
-   if x = a1 then Some b1 else
-   if x < a2 then lookup m x else
-   if x = a2 then Some b2
-   else lookup r x)"
+"lookup (Node2 l (a,b) r) x = (case cmp x a of
+  LT \<Rightarrow> lookup l x |
+  GT \<Rightarrow> lookup r x |
+  EQ \<Rightarrow> Some b)" |
+"lookup (Node3 l (a1,b1) m (a2,b2) r) x = (case cmp x a1 of
+  LT \<Rightarrow> lookup l x |
+  EQ \<Rightarrow> Some b1 |
+  GT \<Rightarrow> (case cmp x a2 of
+          LT \<Rightarrow> lookup m x |
+          EQ \<Rightarrow> Some b2 |
+          GT \<Rightarrow> lookup r x))"
 
-fun upd :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) up\<^sub>i" where
+fun upd :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) up\<^sub>i" where
 "upd x y Leaf = Up\<^sub>i Leaf (x,y) Leaf" |
-"upd x y (Node2 l ab r) =
-   (if x < fst ab then
-        (case upd x y l of
+"upd x y (Node2 l ab r) = (case cmp x (fst ab) of
+   LT \<Rightarrow> (case upd x y l of
            T\<^sub>i l' => T\<^sub>i (Node2 l' ab r)
-         | Up\<^sub>i l1 ab' l2 => T\<^sub>i (Node3 l1 ab' l2 ab r))
-    else if x = fst ab then T\<^sub>i (Node2 l (x,y) r)
-    else
-        (case upd x y r of
+         | Up\<^sub>i l1 ab' l2 => T\<^sub>i (Node3 l1 ab' l2 ab r)) |
+   EQ \<Rightarrow> T\<^sub>i (Node2 l (x,y) r) |
+   GT \<Rightarrow> (case upd x y r of
            T\<^sub>i r' => T\<^sub>i (Node2 l ab r')
          | Up\<^sub>i r1 ab' r2 => T\<^sub>i (Node3 l ab r1 ab' r2)))" |
-"upd x y (Node3 l ab1 m ab2 r) =
-   (if x < fst ab1 then
-        (case upd x y l of
+"upd x y (Node3 l ab1 m ab2 r) = (case cmp x (fst ab1) of
+   LT \<Rightarrow> (case upd x y l of
            T\<^sub>i l' => T\<^sub>i (Node3 l' ab1 m ab2 r)
-         | Up\<^sub>i l1 ab' l2 => Up\<^sub>i (Node2 l1 ab' l2) ab1 (Node2 m ab2 r))
-    else if x = fst ab1 then T\<^sub>i (Node3 l (x,y) m ab2 r)
-    else if x < fst ab2 then
-             (case upd x y m of
-                T\<^sub>i m' => T\<^sub>i (Node3 l ab1 m' ab2 r)
-              | Up\<^sub>i m1 ab' m2 => Up\<^sub>i (Node2 l ab1 m1) ab' (Node2 m2 ab2 r))
-         else if x = fst ab2 then T\<^sub>i (Node3 l ab1 m (x,y) r)
-         else
-             (case upd x y r of
-                T\<^sub>i r' => T\<^sub>i (Node3 l ab1 m ab2 r')
-              | Up\<^sub>i r1 ab' r2 => Up\<^sub>i (Node2 l ab1 m) ab2 (Node2 r1 ab' r2)))"
+         | Up\<^sub>i l1 ab' l2 => Up\<^sub>i (Node2 l1 ab' l2) ab1 (Node2 m ab2 r)) |
+   EQ \<Rightarrow> T\<^sub>i (Node3 l (x,y) m ab2 r) |
+   GT \<Rightarrow> (case cmp x (fst ab2) of
+           LT \<Rightarrow> (case upd x y m of
+                   T\<^sub>i m' => T\<^sub>i (Node3 l ab1 m' ab2 r)
+                 | Up\<^sub>i m1 ab' m2 => Up\<^sub>i (Node2 l ab1 m1) ab' (Node2 m2 ab2 r)) |
+           EQ \<Rightarrow> T\<^sub>i (Node3 l ab1 m (x,y) r) |
+           GT \<Rightarrow> (case upd x y r of
+                   T\<^sub>i r' => T\<^sub>i (Node3 l ab1 m ab2 r')
+                 | Up\<^sub>i r1 ab' r2 => Up\<^sub>i (Node2 l ab1 m) ab2 (Node2 r1 ab' r2))))"
 
-definition update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) tree23" where
+definition update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) tree23" where
 "update a b t = tree\<^sub>i(upd a b t)"
 
-fun del :: "'a::linorder \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) up\<^sub>d"
-where
+fun del :: "'a::cmp \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) up\<^sub>d" where
 "del x Leaf = T\<^sub>d Leaf" |
 "del x (Node2 Leaf ab1 Leaf) = (if x=fst ab1 then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf ab1 Leaf))" |
 "del x (Node3 Leaf ab1 Leaf ab2 Leaf) = T\<^sub>d(if x=fst ab1 then Node2 Leaf ab2 Leaf
   else if x=fst ab2 then Node2 Leaf ab1 Leaf else Node3 Leaf ab1 Leaf ab2 Leaf)" |
-"del x (Node2 l ab1 r) = (if x<fst ab1 then node21 (del x l) ab1 r else
-  if x > fst ab1 then node22 l ab1 (del x r) else
-  let (ab1',t) = del_min r in node22 l ab1' t)" |
-"del x (Node3 l ab1 m ab2 r) = (if x<fst ab1 then node31 (del x l) ab1 m ab2 r else
-  if x = fst ab1 then let (ab1',m') = del_min m in node32 l ab1' m' ab2 r else
-  if x < fst ab2 then node32 l ab1 (del x m) ab2 r else
-  if x = fst ab2 then let (ab2',r') = del_min r in node33 l ab1 m ab2' r'
-  else node33 l ab1 m ab2 (del x r))"
+"del x (Node2 l ab1 r) = (case cmp x (fst ab1) of
+  LT \<Rightarrow> node21 (del x l) ab1 r |
+  GT \<Rightarrow> node22 l ab1 (del x r) |
+  EQ \<Rightarrow> let (ab1',t) = del_min r in node22 l ab1' t)" |
+"del x (Node3 l ab1 m ab2 r) = (case cmp x (fst ab1) of
+  LT \<Rightarrow> node31 (del x l) ab1 m ab2 r |
+  EQ \<Rightarrow> let (ab1',m') = del_min m in node32 l ab1' m' ab2 r |
+  GT \<Rightarrow> (case cmp x (fst ab2) of
+           LT \<Rightarrow> node32 l ab1 (del x m) ab2 r |
+           EQ \<Rightarrow> let (ab2',r') = del_min r in node33 l ab1 m ab2' r' |
+           GT \<Rightarrow> node33 l ab1 m ab2 (del x r)))"
 
-definition delete :: "'a::linorder \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) tree23" where
+definition delete :: "'a::cmp \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) tree23" where
 "delete x t = tree\<^sub>d(del x t)"
 
 
@@ -98,7 +98,7 @@
 subsection \<open>Balancedness\<close>
 
 lemma bal_upd: "bal t \<Longrightarrow> bal (tree\<^sub>i(upd a b t)) \<and> height(upd a b t) = height t"
-by (induct t) (auto split: up\<^sub>i.split)(* 30 secs in 2015 *)
+by (induct t) (auto split: up\<^sub>i.split)(* 16 secs in 2015 *)
 
 corollary bal_update: "bal t \<Longrightarrow> bal (update a b t)"
 by (simp add: update_def bal_upd)
--- a/src/HOL/Data_Structures/Tree23_Set.thy	Thu Nov 05 00:17:13 2015 +0100
+++ b/src/HOL/Data_Structures/Tree23_Set.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -5,14 +5,17 @@
 theory Tree23_Set
 imports
   Tree23
+  Cmp
   Set_by_Ordered
 begin
 
-fun isin :: "'a::linorder tree23 \<Rightarrow> 'a \<Rightarrow> bool" where
+fun isin :: "'a::cmp tree23 \<Rightarrow> 'a \<Rightarrow> bool" where
 "isin Leaf x = False" |
-"isin (Node2 l a r) x = (x < a \<and> isin l x \<or> x=a \<or> isin r x)" |
+"isin (Node2 l a r) x =
+  (case cmp x a of LT \<Rightarrow> isin l x | EQ \<Rightarrow> True | GT \<Rightarrow> isin r x)" |
 "isin (Node3 l a m b r) x =
-  (x < a \<and> isin l x \<or> x > b \<and> isin r x \<or> x = a \<or> x = b \<or> isin m x)"
+  (case cmp x a of LT \<Rightarrow> isin l x | EQ \<Rightarrow> True | GT \<Rightarrow> (case cmp x b of
+   LT \<Rightarrow> isin m x | EQ \<Rightarrow> True | GT \<Rightarrow> isin r x))"
 
 datatype 'a up\<^sub>i = T\<^sub>i "'a tree23" | Up\<^sub>i "'a tree23" 'a "'a tree23"
 
@@ -20,38 +23,35 @@
 "tree\<^sub>i (T\<^sub>i t) = t" |
 "tree\<^sub>i (Up\<^sub>i l p r) = Node2 l p r"
 
-fun ins :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>i" where
+fun ins :: "'a::cmp \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>i" where
 "ins x Leaf = Up\<^sub>i Leaf x Leaf" |
 "ins x (Node2 l a r) =
-   (if x < a then
-      case ins x l of
-         T\<^sub>i l' => T\<^sub>i (Node2 l' a r)
-       | Up\<^sub>i l1 b l2 => T\<^sub>i (Node3 l1 b l2 a r)
-    else if x=a then T\<^sub>i (Node2 l x r)
-    else
-      case ins x r of
-        T\<^sub>i r' => T\<^sub>i (Node2 l a r')
-      | Up\<^sub>i r1 b r2 => T\<^sub>i (Node3 l a r1 b r2))" |
+   (case cmp x a of
+      LT \<Rightarrow> (case ins x l of
+              T\<^sub>i l' => T\<^sub>i (Node2 l' a r)
+            | Up\<^sub>i l1 b l2 => T\<^sub>i (Node3 l1 b l2 a r)) |
+      EQ \<Rightarrow> T\<^sub>i (Node2 l x r) |
+      GT \<Rightarrow> (case ins x r of
+              T\<^sub>i r' => T\<^sub>i (Node2 l a r')
+            | Up\<^sub>i r1 b r2 => T\<^sub>i (Node3 l a r1 b r2)))" |
 "ins x (Node3 l a m b r) =
-   (if x < a then
-      case ins x l of
-        T\<^sub>i l' => T\<^sub>i (Node3 l' a m b r)
-      | Up\<^sub>i l1 c l2 => Up\<^sub>i (Node2 l1 c l2) a (Node2 m b r)
-    else
-    if x > b then
-      case ins x r of
-        T\<^sub>i r' => T\<^sub>i (Node3 l a m b r')
-      | Up\<^sub>i r1 c r2 => Up\<^sub>i (Node2 l a m) b (Node2 r1 c r2)
-    else
-    if x=a \<or> x = b then T\<^sub>i (Node3 l a m b r)
-    else
-      case ins x m of
-        T\<^sub>i m' => T\<^sub>i (Node3 l a m' b r)
-      | Up\<^sub>i m1 c m2 => Up\<^sub>i (Node2 l a m1) c (Node2 m2 b r))"
+   (case cmp x a of
+      LT \<Rightarrow> (case ins x l of
+              T\<^sub>i l' => T\<^sub>i (Node3 l' a m b r)
+            | Up\<^sub>i l1 c l2 => Up\<^sub>i (Node2 l1 c l2) a (Node2 m b r)) |
+      EQ \<Rightarrow> T\<^sub>i (Node3 l a m b r) |
+      GT \<Rightarrow> (case cmp x b of
+               GT \<Rightarrow> (case ins x r of
+                       T\<^sub>i r' => T\<^sub>i (Node3 l a m b r')
+                     | Up\<^sub>i r1 c r2 => Up\<^sub>i (Node2 l a m) b (Node2 r1 c r2)) |
+               EQ \<Rightarrow> T\<^sub>i (Node3 l a m b r) |
+               LT \<Rightarrow> (case ins x m of
+                       T\<^sub>i m' => T\<^sub>i (Node3 l a m' b r)
+                     | Up\<^sub>i m1 c m2 => Up\<^sub>i (Node2 l a m1) c (Node2 m2 b r))))"
 
 hide_const insert
 
-definition insert :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
+definition insert :: "'a::cmp \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
 "insert x t = tree\<^sub>i(ins x t)"
 
 datatype 'a up\<^sub>d = T\<^sub>d "'a tree23" | Up\<^sub>d "'a tree23"
@@ -93,32 +93,34 @@
 "del_min (Node2 l a r) = (let (x,l') = del_min l in (x, node21 l' a r))" |
 "del_min (Node3 l a m b r) = (let (x,l') = del_min l in (x, node31 l' a m b r))"
 
-fun del :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>d"
+fun del :: "'a::cmp \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>d"
 where
 "del x Leaf = T\<^sub>d Leaf" |
 "del x (Node2 Leaf a Leaf) = (if x = a then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf a Leaf))" |
 "del x (Node3 Leaf a Leaf b Leaf) = T\<^sub>d(if x = a then Node2 Leaf b Leaf
   else if x = b then Node2 Leaf a Leaf else Node3 Leaf a Leaf b Leaf)" |
-"del x (Node2 l a r) = (if x<a then node21 (del x l) a r else
-  if x > a then node22 l a (del x r) else
-  let (a',t) = del_min r in node22 l a' t)" |
-"del x (Node3 l a m b r) = (if x<a then node31 (del x l) a m b r else
-  if x = a then let (a',m') = del_min m in node32 l a' m' b r else
-  if x < b then node32 l a (del x m) b r else
-  if x = b then let (b',r') = del_min r in node33 l a m b' r'
-  else node33 l a m b (del x r))"
+"del x (Node2 l a r) = (case cmp x a of
+  LT \<Rightarrow> node21 (del x l) a r |
+  GT \<Rightarrow> node22 l a (del x r) |
+  EQ \<Rightarrow> let (a',t) = del_min r in node22 l a' t)" |
+"del x (Node3 l a m b r) = (case cmp x a of
+  LT \<Rightarrow> node31 (del x l) a m b r |
+  EQ \<Rightarrow> let (a',m') = del_min m in node32 l a' m' b r |
+  GT \<Rightarrow> (case cmp x b of
+          LT \<Rightarrow> node32 l a (del x m) b r |
+          EQ \<Rightarrow> let (b',r') = del_min r in node33 l a m b' r' |
+          GT \<Rightarrow> node33 l a m b (del x r)))"
 
-definition delete :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
+definition delete :: "'a::cmp \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
 "delete x t = tree\<^sub>d(del x t)"
 
 
 subsection "Functional Correctness"
 
-
 subsubsection "Proofs for isin"
 
 lemma "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
-by (induction t) (auto simp: elems_simps1)
+by (induction t) (auto simp: elems_simps1 ball_Un)
 
 lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
 by (induction t) (auto simp: elems_simps2)
@@ -128,7 +130,7 @@
 
 lemma inorder_ins:
   "sorted(inorder t) \<Longrightarrow> inorder(tree\<^sub>i(ins x t)) = ins_list x (inorder t)"
-by(induction t) (auto simp: ins_list_simps split: up\<^sub>i.splits) (* 38 secs in 2015 *)
+by(induction t) (auto simp: ins_list_simps split: up\<^sub>i.splits)
 
 lemma inorder_insert:
   "sorted(inorder t) \<Longrightarrow> inorder(insert a t) = ins_list a (inorder t)"
@@ -195,7 +197,7 @@
 end
 
 lemma bal_ins: "bal t \<Longrightarrow> bal (tree\<^sub>i(ins a t)) \<and> height(ins a t) = height t"
-by (induct t) (auto split: up\<^sub>i.split) (* 87 secs in 2015 *)
+by (induct t) (auto split: up\<^sub>i.split) (* 15 secs in 2015 *)
 
 text{* Now an alternative proof (by Brian Huffman) that runs faster because
 two properties (balance and height) are combined in one predicate. *}
--- a/src/HOL/Data_Structures/Tree_Map.thy	Thu Nov 05 00:17:13 2015 +0100
+++ b/src/HOL/Data_Structures/Tree_Map.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -8,23 +8,24 @@
   Map_by_Ordered
 begin
 
-fun lookup :: "('a::linorder*'b) tree \<Rightarrow> 'a \<Rightarrow> 'b option" where
+fun lookup :: "('a::cmp*'b) tree \<Rightarrow> 'a \<Rightarrow> 'b option" where
 "lookup Leaf x = None" |
-"lookup (Node l (a,b) r) x = (if x < a then lookup l x else
-  if x > a then lookup r x else Some b)"
+"lookup (Node l (a,b) r) x =
+  (case cmp x a of LT \<Rightarrow> lookup l x | GT \<Rightarrow> lookup r x | EQ \<Rightarrow> Some b)"
 
-fun update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree \<Rightarrow> ('a*'b) tree" where
+fun update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree \<Rightarrow> ('a*'b) tree" where
 "update x y Leaf = Node Leaf (x,y) Leaf" |
-"update x y (Node l (a,b) r) =
-   (if x < a then Node (update x y l) (a,b) r
-    else if x = a then Node l (x,y) r
-    else Node l (a,b) (update x y r))"
+"update x y (Node l (a,b) r) = (case cmp x a of
+   LT \<Rightarrow> Node (update x y l) (a,b) r |
+   EQ \<Rightarrow> Node l (x,y) r |
+   GT \<Rightarrow> Node l (a,b) (update x y r))"
 
-fun delete :: "'a::linorder \<Rightarrow> ('a*'b) tree \<Rightarrow> ('a*'b) tree" where
+fun delete :: "'a::cmp \<Rightarrow> ('a*'b) tree \<Rightarrow> ('a*'b) tree" where
 "delete x Leaf = Leaf" |
-"delete x (Node l (a,b) r) = (if x < a then Node (delete x l) (a,b) r else
-  if x > a then Node l (a,b) (delete x r) else
-  if r = Leaf then l else let (ab',r') = del_min r in Node l ab' r')"
+"delete x (Node l (a,b) r) = (case cmp x a of
+  LT \<Rightarrow> Node (delete x l) (a,b) r |
+  GT \<Rightarrow> Node l (a,b) (delete x r) |
+  EQ \<Rightarrow> if r = Leaf then l else let (ab',r') = del_min r in Node l ab' r')"
 
 
 subsection "Functional Correctness Proofs"
@@ -49,7 +50,6 @@
   "sorted1(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
 by(induction t) (auto simp: del_list_simps del_minD split: prod.splits)
 
-
 interpretation Map_by_Ordered
 where empty = Leaf and lookup = lookup and update = update and delete = delete
 and inorder = inorder and wf = "\<lambda>_. True"
--- a/src/HOL/Data_Structures/Tree_Set.thy	Thu Nov 05 00:17:13 2015 +0100
+++ b/src/HOL/Data_Structures/Tree_Set.thy	Thu Nov 05 08:27:22 2015 +0100
@@ -5,31 +5,34 @@
 theory Tree_Set
 imports
   "~~/src/HOL/Library/Tree"
+  Cmp
   Set_by_Ordered
 begin
 
-fun isin :: "'a::linorder tree \<Rightarrow> 'a \<Rightarrow> bool" where
+fun isin :: "'a::cmp tree \<Rightarrow> 'a \<Rightarrow> bool" where
 "isin Leaf x = False" |
-"isin (Node l a r) x = (x < a \<and> isin l x \<or> x=a \<or> isin r x)"
+"isin (Node l a r) x =
+  (case cmp x a of LT \<Rightarrow> isin l x | EQ \<Rightarrow> True | GT \<Rightarrow> isin r x)"
 
 hide_const (open) insert
 
-fun insert :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
+fun insert :: "'a::cmp \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
 "insert x Leaf = Node Leaf x Leaf" |
-"insert x (Node l a r) =
-   (if x < a then Node (insert x l) a r else
-    if x = a then Node l a r
-    else Node l a (insert x r))"
+"insert x (Node l a r) = (case cmp x a of
+      LT \<Rightarrow> Node (insert x l) a r |
+      EQ \<Rightarrow> Node l a r |
+      GT \<Rightarrow> Node l a (insert x r))"
 
 fun del_min :: "'a tree \<Rightarrow> 'a * 'a tree" where
 "del_min (Node Leaf a r) = (a, r)" |
 "del_min (Node l a r) = (let (x,l') = del_min l in (x, Node l' a r))"
 
-fun delete :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
+fun delete :: "'a::cmp \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
 "delete x Leaf = Leaf" |
-"delete x (Node l a r) = (if x < a then Node (delete x l) a r else
-  if x > a then Node l a (delete x r) else
-  if r = Leaf then l else let (a',r') = del_min r in Node l a' r')"
+"delete x (Node l a r) = (case cmp x a of
+  LT \<Rightarrow>  Node (delete x l) a r |
+  GT \<Rightarrow>  Node l a (delete x r) |
+  EQ \<Rightarrow> if r = Leaf then l else let (a',r') = del_min r in Node l a' r')"
 
 
 subsection "Functional Correctness Proofs"
@@ -56,7 +59,6 @@
   "sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
 by(induction t) (auto simp: del_list_simps del_minD split: prod.splits)
 
-
 interpretation Set_by_Ordered
 where empty = Leaf and isin = isin and insert = insert and delete = delete
 and inorder = inorder and wf = "\<lambda>_. True"