more abstract naming
authornipkow
Tue, 12 Jun 2018 17:18:40 +0200
changeset 68431 b294e095f64c
parent 68423 c1db7503dbaa
child 68432 6f3bd27ba389
more abstract naming
src/HOL/Data_Structures/AA_Map.thy
src/HOL/Data_Structures/AA_Set.thy
src/HOL/Data_Structures/AVL_Map.thy
src/HOL/Data_Structures/AVL_Set.thy
src/HOL/Data_Structures/Brother12_Map.thy
src/HOL/Data_Structures/Brother12_Set.thy
src/HOL/Data_Structures/Map_Specs.thy
src/HOL/Data_Structures/RBT_Map.thy
src/HOL/Data_Structures/RBT_Set.thy
src/HOL/Data_Structures/Tree234_Map.thy
src/HOL/Data_Structures/Tree234_Set.thy
src/HOL/Data_Structures/Tree23_Map.thy
src/HOL/Data_Structures/Tree23_Set.thy
src/HOL/Data_Structures/Tree_Map.thy
src/HOL/Data_Structures/Tree_Set.thy
--- a/src/HOL/Data_Structures/AA_Map.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/AA_Map.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -207,10 +207,10 @@
               post_split_max post_delete split_maxD split: prod.splits)
 
 interpretation I: Map_by_Ordered
-where empty = Leaf and lookup = lookup and update = update and delete = delete
+where empty = empty and lookup = lookup and update = update and delete = delete
 and inorder = inorder and inv = invar
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: lookup_map_of)
 next
@@ -218,7 +218,7 @@
 next
   case 4 thus ?case by(simp add: inorder_delete)
 next
-  case 5 thus ?case by(simp)
+  case 5 thus ?case by(simp add: empty_def)
 next
   case 6 thus ?case by(simp add: invar_update)
 next
--- a/src/HOL/Data_Structures/AA_Set.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/AA_Set.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -12,6 +12,9 @@
 
 type_synonym 'a aa_tree = "('a,nat) tree"
 
+definition empty :: "'a aa_tree" where
+"empty = Leaf"
+
 fun lvl :: "'a aa_tree \<Rightarrow> nat" where
 "lvl Leaf = 0" |
 "lvl (Node _ _ lv _) = lv"
@@ -483,10 +486,10 @@
               post_split_max post_delete split_maxD split: prod.splits)
 
 interpretation I: Set_by_Ordered
-where empty = Leaf and isin = isin and insert = insert and delete = delete
+where empty = empty and isin = isin and insert = insert and delete = delete
 and inorder = inorder and inv = invar
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: isin_set_inorder)
 next
@@ -494,7 +497,7 @@
 next
   case 4 thus ?case by(simp add: inorder_delete)
 next
-  case 5 thus ?case by(simp)
+  case 5 thus ?case by(simp add: empty_def)
 next
   case 6 thus ?case by(simp add: invar_insert)
 next
--- a/src/HOL/Data_Structures/AVL_Map.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/AVL_Map.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -25,12 +25,12 @@
 
 subsection \<open>Functional Correctness\<close>
 
-theorem inorder_update:
+theorem inorder_update_avl:
   "sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
 by (induct t) (auto simp: upd_list_simps inorder_balL inorder_balR)
 
 
-theorem inorder_delete:
+theorem inorder_delete_avl:
   "sorted1(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
 by(induction t)
   (auto simp: del_list_simps inorder_balL inorder_balR
@@ -181,18 +181,18 @@
 
 
 interpretation Map_by_Ordered
-where empty = Leaf and lookup = lookup and update = update and delete = delete
+where empty = empty and lookup = lookup and update = update and delete = delete
 and inorder = inorder and inv = avl
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: lookup_map_of)
 next
-  case 3 thus ?case by(simp add: inorder_update)
+  case 3 thus ?case by(simp add: inorder_update_avl)
 next
-  case 4 thus ?case by(simp add: inorder_delete)
+  case 4 thus ?case by(simp add: inorder_delete_avl)
 next
-  case 5 show ?case by simp
+  case 5 show ?case by (simp add: empty_def)
 next
   case 6 thus ?case by(simp add: avl_update(1))
 next
--- a/src/HOL/Data_Structures/AVL_Set.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/AVL_Set.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -14,6 +14,9 @@
 
 type_synonym 'a avl_tree = "('a,nat) tree"
 
+definition empty :: "'a avl_tree" where
+"empty = Leaf"
+
 text \<open>Invariant:\<close>
 
 fun avl :: "'a avl_tree \<Rightarrow> bool" where
@@ -421,10 +424,10 @@
 subsection "Overall correctness"
 
 interpretation Set_by_Ordered
-where empty = Leaf and isin = isin and insert = insert and delete = delete
+where empty = empty and isin = isin and insert = insert and delete = delete
 and inorder = inorder and inv = avl
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: isin_set_inorder)
 next
@@ -432,7 +435,7 @@
 next
   case 4 thus ?case by(simp add: inorder_delete)
 next
-  case 5 thus ?case by simp
+  case 5 thus ?case by (simp add: empty_def)
 next
   case 6 thus ?case by (simp add: avl_insert(1))
 next
--- a/src/HOL/Data_Structures/Brother12_Map.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/Brother12_Map.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -192,7 +192,7 @@
 subsection "Overall correctness"
 
 interpretation Map_by_Ordered
-where empty = N0 and lookup = lookup and update = update.update
+where empty = empty and lookup = lookup and update = update.update
 and delete = delete.delete and inorder = inorder and inv = "\<lambda>t. \<exists>h. t \<in> B h"
 proof (standard, goal_cases)
   case 2 thus ?case by(auto intro!: lookup_map_of)
@@ -204,6 +204,6 @@
   case 6 thus ?case using update.update_type by (metis Un_iff)
 next
   case 7 thus ?case using delete.delete_type by blast
-qed auto
+qed (auto simp: empty_def)
 
 end
--- a/src/HOL/Data_Structures/Brother12_Set.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/Brother12_Set.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -19,6 +19,9 @@
   L2 'a |
   N3 "'a bro" 'a "'a bro" 'a "'a bro"
 
+definition empty :: "'a bro" where
+"empty = N0"
+
 fun inorder :: "'a bro \<Rightarrow> 'a list" where
 "inorder N0 = []" |
 "inorder (N1 t) = inorder t" |
@@ -465,7 +468,7 @@
 subsection "Overall correctness"
 
 interpretation Set_by_Ordered
-where empty = N0 and isin = isin and insert = insert.insert
+where empty = empty and isin = isin and insert = insert.insert
 and delete = delete.delete and inorder = inorder and inv = "\<lambda>t. \<exists>h. t \<in> B h"
 proof (standard, goal_cases)
   case 2 thus ?case by(auto intro!: isin_set)
@@ -477,7 +480,7 @@
   case 6 thus ?case using insert.insert_type by blast
 next
   case 7 thus ?case using delete.delete_type by blast
-qed auto
+qed (auto simp: empty_def)
 
 
 subsection \<open>Height-Size Relation\<close>
--- a/src/HOL/Data_Structures/Map_Specs.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/Map_Specs.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -31,36 +31,41 @@
 fixes lookup :: "'t \<Rightarrow> 'a \<Rightarrow> 'b option"
 fixes inorder :: "'t \<Rightarrow> ('a * 'b) list"
 fixes inv :: "'t \<Rightarrow> bool"
-assumes empty: "inorder empty = []"
-and lookup: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
+assumes inorder_empty: "inorder empty = []"
+and inorder_lookup: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
   lookup t a = map_of (inorder t) a"
-and update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
+and inorder_update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
   inorder(update a b t) = upd_list a b (inorder t)"
-and delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
+and inorder_delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
   inorder(delete a t) = del_list a (inorder t)"
-and inv_empty:  "inv empty"
-and inv_update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(update a b t)"
-and inv_delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(delete a t)"
+and inorder_inv_empty:  "inv empty"
+and inorder_inv_update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(update a b t)"
+and inorder_inv_delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(delete a t)"
 begin
 
 text \<open>It implements the traditional specification:\<close>
 
+definition invar :: "'t \<Rightarrow> bool" where
+"invar t == inv t \<and> sorted1 (inorder t)"
+
 sublocale Map
-  empty update delete lookup "\<lambda>t. inv t \<and> sorted1 (inorder t)"
+  empty update delete lookup invar
 proof(standard, goal_cases)
-  case 1 show ?case by (auto simp: lookup empty inv_empty)
+  case 1 show ?case by (auto simp: inorder_lookup inorder_empty inorder_inv_empty)
 next
   case 2 thus ?case
-    by(simp add: fun_eq_iff update inv_update map_of_ins_list lookup sorted_upd_list)
+    by(simp add: fun_eq_iff inorder_update inorder_inv_update map_of_ins_list inorder_lookup
+         sorted_upd_list invar_def)
 next
   case 3 thus ?case
-    by(simp add: fun_eq_iff delete inv_delete map_of_del_list lookup sorted_del_list)
-next
-  case 4 thus ?case by(simp add: empty inv_empty)
+    by(simp add: fun_eq_iff inorder_delete inorder_inv_delete map_of_del_list inorder_lookup
+         sorted_del_list invar_def)
 next
-  case 5 thus ?case by(simp add: update inv_update sorted_upd_list)
+  case 4 thus ?case by(simp add: inorder_empty inorder_inv_empty invar_def)
 next
-  case 6 thus ?case by (auto simp: delete inv_delete sorted_del_list)
+  case 5 thus ?case by(simp add: inorder_update inorder_inv_update sorted_upd_list invar_def)
+next
+  case 6 thus ?case by (auto simp: inorder_delete inorder_inv_delete sorted_del_list invar_def)
 qed
 
 end
--- a/src/HOL/Data_Structures/RBT_Map.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/RBT_Map.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -42,7 +42,7 @@
 by(induction x y t rule: upd.induct)
   (auto simp: upd_list_simps inorder_baliL inorder_baliR)
 
-lemma inorder_update:
+lemma inorder_update_rbt:
   "sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
 by(simp add: update_def inorder_upd inorder_paint)
 
@@ -51,7 +51,7 @@
 by(induction x t rule: del.induct)
   (auto simp: del_list_simps inorder_combine inorder_baldL inorder_baldR)
 
-lemma inorder_delete:
+lemma inorder_delete_rbt:
   "sorted1(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
 by(simp add: delete_def inorder_del inorder_paint)
 
@@ -105,18 +105,18 @@
 by (metis delete_def rbt_def color_paint_Black del_invc_invh invc_paint_Black invc2I invh_paint)
 
 interpretation Map_by_Ordered
-where empty = Leaf and lookup = lookup and update = update and delete = delete
+where empty = empty and lookup = lookup and update = update and delete = delete
 and inorder = inorder and inv = rbt
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: lookup_map_of)
 next
-  case 3 thus ?case by(simp add: inorder_update)
+  case 3 thus ?case by(simp add: inorder_update_rbt)
 next
-  case 4 thus ?case by(simp add: inorder_delete)
+  case 4 thus ?case by(simp add: inorder_delete_rbt)
 next
-  case 5 thus ?case by (simp add: rbt_Leaf) 
+  case 5 thus ?case by (simp add: rbt_def empty_def) 
 next
   case 6 thus ?case by (simp add: rbt_update) 
 next
--- a/src/HOL/Data_Structures/RBT_Set.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/RBT_Set.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -10,6 +10,9 @@
   Isin2
 begin
 
+definition empty :: "'a rbt" where
+"empty = Leaf"
+
 fun ins :: "'a::linorder \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
 "ins x Leaf = R Leaf x Leaf" |
 "ins x (B l a r) =
@@ -121,9 +124,6 @@
 lemma color_paint_Black: "color (paint Black t) = Black"
 by (cases t) auto
 
-theorem rbt_Leaf: "rbt Leaf"
-by (simp add: rbt_def)
-
 lemma paint_invc2: "invc2 t \<Longrightarrow> invc2 (paint c t)"
 by (cases t) auto
 
@@ -257,10 +257,10 @@
 text \<open>Overall correctness:\<close>
 
 interpretation Set_by_Ordered
-where empty = Leaf and isin = isin and insert = insert and delete = delete
+where empty = empty and isin = isin and insert = insert and delete = delete
 and inorder = inorder and inv = rbt
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: isin_set_inorder)
 next
@@ -268,7 +268,7 @@
 next
   case 4 thus ?case by(simp add: inorder_delete)
 next
-  case 5 thus ?case by (simp add: rbt_Leaf) 
+  case 5 thus ?case by (simp add: rbt_def empty_def) 
 next
   case 6 thus ?case by (simp add: rbt_insert) 
 next
--- a/src/HOL/Data_Structures/Tree234_Map.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/Tree234_Map.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -123,7 +123,7 @@
 by(induction t)
   (auto simp: upd_list_simps, auto simp: upd_list_simps split: up\<^sub>i.splits)
 
-lemma inorder_update:
+lemma inorder_update_234:
   "sorted1(inorder t) \<Longrightarrow> inorder(update a b t) = upd_list a b (inorder t)"
 by(simp add: update_def inorder_upd)
 
@@ -133,7 +133,7 @@
   (auto simp: del_list_simps inorder_nodes split_minD split!: if_splits prod.splits)
 (* 30 secs (2016) *)
 
-lemma inorder_delete: "\<lbrakk> bal t ; sorted1(inorder t) \<rbrakk> \<Longrightarrow>
+lemma inorder_delete_234: "\<lbrakk> bal t ; sorted1(inorder t) \<rbrakk> \<Longrightarrow>
   inorder(delete x t) = del_list x (inorder t)"
 by(simp add: delete_def inorder_del)
 
@@ -161,18 +161,18 @@
 subsection \<open>Overall Correctness\<close>
 
 interpretation Map_by_Ordered
-where empty = Leaf and lookup = lookup and update = update and delete = delete
+where empty = empty and lookup = lookup and update = update and delete = delete
 and inorder = inorder and inv = bal
 proof (standard, goal_cases)
   case 2 thus ?case by(simp add: lookup_map_of)
 next
-  case 3 thus ?case by(simp add: inorder_update)
+  case 3 thus ?case by(simp add: inorder_update_234)
 next
-  case 4 thus ?case by(simp add: inorder_delete)
+  case 4 thus ?case by(simp add: inorder_delete_234)
 next
   case 6 thus ?case by(simp add: bal_update)
 next
   case 7 thus ?case by(simp add: bal_delete)
-qed simp+
+qed (simp add: empty_def)+
 
 end
--- a/src/HOL/Data_Structures/Tree234_Set.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/Tree234_Set.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -13,6 +13,9 @@
 
 subsection \<open>Set operations on 2-3-4 trees\<close>
 
+definition empty :: "'a tree234" where
+"empty = Leaf"
+
 fun isin :: "'a::linorder tree234 \<Rightarrow> 'a \<Rightarrow> bool" where
 "isin Leaf x = False" |
 "isin (Node2 l a r) x =
@@ -502,7 +505,7 @@
 subsection \<open>Overall Correctness\<close>
 
 interpretation Set_by_Ordered
-where empty = Leaf and isin = isin and insert = insert and delete = delete
+where empty = empty and isin = isin and insert = insert and delete = delete
 and inorder = inorder and inv = bal
 proof (standard, goal_cases)
   case 2 thus ?case by(simp add: isin_set)
@@ -514,6 +517,6 @@
   case 6 thus ?case by(simp add: bal_insert)
 next
   case 7 thus ?case by(simp add: bal_delete)
-qed simp+
+qed (simp add: empty_def)+
 
 end
--- a/src/HOL/Data_Structures/Tree23_Map.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/Tree23_Map.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -81,7 +81,7 @@
   "sorted1(inorder t) \<Longrightarrow> inorder(tree\<^sub>i(upd x y t)) = upd_list x y (inorder t)"
 by(induction t) (auto simp: upd_list_simps split: up\<^sub>i.splits)
 
-corollary inorder_update:
+corollary inorder_update_23:
   "sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
 by(simp add: update_def inorder_upd)
 
@@ -91,7 +91,7 @@
 by(induction t rule: del.induct)
   (auto simp: del_list_simps inorder_nodes split_minD split!: if_split prod.splits)
 
-corollary inorder_delete: "\<lbrakk> bal t ; sorted1(inorder t) \<rbrakk> \<Longrightarrow>
+corollary inorder_delete_23: "\<lbrakk> bal t ; sorted1(inorder t) \<rbrakk> \<Longrightarrow>
   inorder(delete x t) = del_list x (inorder t)"
 by(simp add: delete_def inorder_del)
 
@@ -120,18 +120,22 @@
 subsection \<open>Overall Correctness\<close>
 
 interpretation Map_by_Ordered
-where empty = Leaf and lookup = lookup and update = update and delete = delete
+where empty = empty and lookup = lookup and update = update and delete = delete
 and inorder = inorder and inv = bal
 proof (standard, goal_cases)
+  case 1 thus ?case by(simp add: empty_def)
+next
   case 2 thus ?case by(simp add: lookup_map_of)
 next
-  case 3 thus ?case by(simp add: inorder_update)
+  case 3 thus ?case by(simp add: inorder_update_23)
 next
-  case 4 thus ?case by(simp add: inorder_delete)
+  case 4 thus ?case by(simp add: inorder_delete_23)
+next
+  case 5 thus ?case by(simp add: empty_def)
 next
   case 6 thus ?case by(simp add: bal_update)
 next
   case 7 thus ?case by(simp add: bal_delete)
-qed simp+
+qed
 
 end
--- a/src/HOL/Data_Structures/Tree23_Set.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/Tree23_Set.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -11,6 +11,9 @@
 
 declare sorted_wrt.simps(2)[simp del]
 
+definition empty :: "'a tree23" where
+"empty = Leaf"
+
 fun isin :: "'a::linorder tree23 \<Rightarrow> 'a \<Rightarrow> bool" where
 "isin Leaf x = False" |
 "isin (Node2 l a r) x =
@@ -377,7 +380,7 @@
 subsection \<open>Overall Correctness\<close>
 
 interpretation Set_by_Ordered
-where empty = Leaf and isin = isin and insert = insert and delete = delete
+where empty = empty and isin = isin and insert = insert and delete = delete
 and inorder = inorder and inv = bal
 proof (standard, goal_cases)
   case 2 thus ?case by(simp add: isin_set)
@@ -389,6 +392,6 @@
   case 6 thus ?case by(simp add: bal_insert)
 next
   case 7 thus ?case by(simp add: bal_delete)
-qed simp+
+qed (simp add: empty_def)+
 
 end
--- a/src/HOL/Data_Structures/Tree_Map.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/Tree_Map.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -34,25 +34,25 @@
   "sorted1(inorder t) \<Longrightarrow> lookup t x = map_of (inorder t) x"
 by (induction t) (auto simp: map_of_simps split: option.split)
 
-lemma inorder_update:
+lemma inorder_update_tree:
   "sorted1(inorder t) \<Longrightarrow> inorder(update a b t) = upd_list a b (inorder t)"
 by(induction t) (auto simp: upd_list_simps)
 
-lemma inorder_delete:
+lemma inorder_delete_tree:
   "sorted1(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
 by(induction t) (auto simp: del_list_simps split_minD split: prod.splits)
 
 interpretation Map_by_Ordered
-where empty = Leaf and lookup = lookup and update = update and delete = delete
+where empty = empty and lookup = lookup and update = update and delete = delete
 and inorder = inorder and inv = "\<lambda>_. True"
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: lookup_map_of)
 next
-  case 3 thus ?case by(simp add: inorder_update)
+  case 3 thus ?case by(simp add: inorder_update_tree)
 next
-  case 4 thus ?case by(simp add: inorder_delete)
+  case 4 thus ?case by(simp add: inorder_delete_tree)
 qed auto
 
 end
--- a/src/HOL/Data_Structures/Tree_Set.thy	Tue Jun 12 07:18:18 2018 +0200
+++ b/src/HOL/Data_Structures/Tree_Set.thy	Tue Jun 12 17:18:40 2018 +0200
@@ -9,6 +9,9 @@
   Set_Specs
 begin
 
+definition empty :: "'a tree" where
+"empty == Leaf"
+
 fun isin :: "'a::linorder tree \<Rightarrow> 'a \<Rightarrow> bool" where
 "isin Leaf x = False" |
 "isin (Node l a r) x =
@@ -60,10 +63,10 @@
 by(induction t) (auto simp: del_list_simps split_minD split: prod.splits)
 
 interpretation Set_by_Ordered
-where empty = Leaf and isin = isin and insert = insert and delete = delete
+where empty = empty and isin = isin and insert = insert and delete = delete
 and inorder = inorder and inv = "\<lambda>_. True"
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: isin_set)
 next