tuned
authornipkow
Sat, 11 May 2019 22:19:28 +0200
changeset 70266 0b813a1a833f
parent 70265 a8238fd25541
child 70267 9fa2cf7142b7
tuned
src/HOL/Data_Structures/Trie_Fun.thy
src/HOL/Data_Structures/Trie_Map.thy
--- a/src/HOL/Data_Structures/Trie_Fun.thy	Sat May 11 19:08:26 2019 +0200
+++ b/src/HOL/Data_Structures/Trie_Fun.thy	Sat May 11 22:19:28 2019 +0200
@@ -10,6 +10,9 @@
 
 datatype 'a trie = Nd bool "'a \<Rightarrow> 'a trie option"
 
+definition empty :: "'a trie" where
+[simp]: "empty = Nd False (\<lambda>_. None)"
+
 fun isin :: "'a trie \<Rightarrow> 'a list \<Rightarrow> bool" where
 "isin (Nd b m) [] = b" |
 "isin (Nd b m) (k # xs) = (case m k of None \<Rightarrow> False | Some t \<Rightarrow> isin t xs)"
@@ -17,7 +20,7 @@
 fun insert :: "('a::linorder) list \<Rightarrow> 'a trie \<Rightarrow> 'a trie" where
 "insert [] (Nd b m) = Nd True m" |
 "insert (x#xs) (Nd b m) =
-   Nd b (m(x := Some(insert xs (case m x of None \<Rightarrow> Nd False (\<lambda>_. None) | Some t \<Rightarrow> t))))"
+   Nd b (m(x := Some(insert xs (case m x of None \<Rightarrow> empty | Some t \<Rightarrow> t))))"
 
 fun delete :: "('a::linorder) list \<Rightarrow> 'a trie \<Rightarrow> 'a trie" where
 "delete [] (Nd b m) = Nd False m" |
@@ -73,7 +76,7 @@
 qed
 
 interpretation S: Set
-where empty = "Nd False (\<lambda>_. None)" and isin = isin and insert = insert and delete = delete
+where empty = empty and isin = isin and insert = insert and delete = delete
 and set = set and invar = "\<lambda>_. True"
 proof (standard, goal_cases)
   case 1 show ?case by (simp)
--- a/src/HOL/Data_Structures/Trie_Map.thy	Sat May 11 19:08:26 2019 +0200
+++ b/src/HOL/Data_Structures/Trie_Map.thy	Sat May 11 22:19:28 2019 +0200
@@ -29,6 +29,10 @@
 apply(auto split: if_splits)
 done
 
+
+definition empty :: "'a trie_map" where
+[simp]: "empty = Nd False Leaf"
+
 fun isin :: "('a::linorder) trie_map \<Rightarrow> 'a list \<Rightarrow> bool" where
 "isin (Nd b m) [] = b" |
 "isin (Nd b m) (x # xs) = (case lookup m x of None \<Rightarrow> False | Some t \<Rightarrow> isin t xs)"
@@ -36,7 +40,7 @@
 fun insert :: "('a::linorder) list \<Rightarrow> 'a trie_map \<Rightarrow> 'a trie_map" where
 "insert [] (Nd b m) = Nd True m" |
 "insert (x#xs) (Nd b m) =
-  Nd b (update x (insert xs (case lookup m x of None \<Rightarrow> Nd False Leaf | Some t \<Rightarrow> t)) m)"
+  Nd b (update x (insert xs (case lookup m x of None \<Rightarrow> empty | Some t \<Rightarrow> t)) m)"
 
 fun delete :: "('a::linorder) list \<Rightarrow> 'a trie_map \<Rightarrow> 'a trie_map" where
 "delete [] (Nd b m) = Nd False m" |
@@ -84,7 +88,7 @@
 text \<open>Overall correctness w.r.t. the \<open>Set\<close> ADT:\<close>
 
 interpretation S2: Set
-where empty = "Nd False Leaf" and isin = isin and insert = insert and delete = delete
+where empty = empty and isin = isin and insert = insert and delete = delete
 and set = "set o abs" and invar = invar
 proof (standard, goal_cases)
   case 1 show ?case by (simp)