"app" -> "join" for uniformity with Join theory; tuned defs
authornipkow
Tue, 12 May 2020 10:24:53 +0200
changeset 71829 6f2663df8374
parent 71828 415c38ef38c0
child 71830 7a997ead54b0
"app" -> "join" for uniformity with Join theory; tuned defs
src/HOL/Data_Structures/Set_Specs.thy
src/HOL/Data_Structures/Tree_Set.thy
--- a/src/HOL/Data_Structures/Set_Specs.thy	Mon May 11 19:41:31 2020 +0200
+++ b/src/HOL/Data_Structures/Set_Specs.thy	Tue May 12 10:24:53 2020 +0200
@@ -51,10 +51,10 @@
 text \<open>It implements the traditional specification:\<close>
 
 definition set :: "'t \<Rightarrow> 'a set" where
-"set == List.set o inorder"
+"set = List.set o inorder"
 
 definition invar :: "'t \<Rightarrow> bool" where
-"invar t == inv t \<and> sorted (inorder t)"
+"invar t = (inv t \<and> sorted (inorder t))"
 
 sublocale Set
   empty insert delete isin set invar
--- a/src/HOL/Data_Structures/Tree_Set.thy	Mon May 11 19:41:31 2020 +0200
+++ b/src/HOL/Data_Structures/Tree_Set.thy	Tue May 12 10:24:53 2020 +0200
@@ -10,7 +10,7 @@
 begin
 
 definition empty :: "'a tree" where
-"empty == Leaf"
+"empty = Leaf"
 
 fun isin :: "'a::linorder tree \<Rightarrow> 'a \<Rightarrow> bool" where
 "isin Leaf x = False" |
@@ -44,13 +44,13 @@
      GT \<Rightarrow>  Node l a (delete x r) |
      EQ \<Rightarrow> if r = Leaf then l else let (a',r') = split_min r in Node l a' r')"
 
-text \<open>Deletion by appending:\<close>
+text \<open>Deletion by joining:\<close>
 
-fun app :: "('a::linorder)tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
-"app t Leaf = t" |
-"app Leaf t = t" |
-"app (Node t1 a t2) (Node t3 b t4) =
-  (case app t2 t3 of
+fun join :: "('a::linorder)tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
+"join t Leaf = t" |
+"join Leaf t = t" |
+"join (Node t1 a t2) (Node t3 b t4) =
+  (case join t2 t3 of
      Leaf \<Rightarrow> Node t1 a (Node Leaf b t4) |
      Node u2 x u3 \<Rightarrow> Node (Node t1 a u2) x (Node u3 b t4))"
 
@@ -60,7 +60,7 @@
   (case cmp x a of
      LT \<Rightarrow>  Node (delete2 x l) a r |
      GT \<Rightarrow>  Node l a (delete2 x r) |
-     EQ \<Rightarrow> app l r)"
+     EQ \<Rightarrow> join l r)"
 
 
 subsection "Functional Correctness Proofs"
@@ -95,13 +95,13 @@
   case 4 thus ?case by(simp add: inorder_delete)
 qed (rule TrueI)+
 
-lemma inorder_app:
-  "inorder(app l r) = inorder l @ inorder r"
-by(induction l r rule: app.induct) (auto split: tree.split)
+lemma inorder_join:
+  "inorder(join l r) = inorder l @ inorder r"
+by(induction l r rule: join.induct) (auto split: tree.split)
 
 lemma inorder_delete2:
   "sorted(inorder t) \<Longrightarrow> inorder(delete2 x t) = del_list x (inorder t)"
-by(induction t) (auto simp: inorder_app del_list_simps)
+by(induction t) (auto simp: inorder_join del_list_simps)
 
 interpretation S2: Set_by_Ordered
 where empty = empty and isin = isin and insert = insert and delete = delete2