# HG changeset patch # User nipkow # Date 1557605968 -7200 # Node ID 0b813a1a833fdc00691b3be1eea12f72e4556a63 # Parent a8238fd25541bea1871b1b212dc9e0a06c2a9ba3 tuned diff -r a8238fd25541 -r 0b813a1a833f src/HOL/Data_Structures/Trie_Fun.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 \ 'a trie option" +definition empty :: "'a trie" where +[simp]: "empty = Nd False (\_. None)" + fun isin :: "'a trie \ 'a list \ bool" where "isin (Nd b m) [] = b" | "isin (Nd b m) (k # xs) = (case m k of None \ False | Some t \ isin t xs)" @@ -17,7 +20,7 @@ fun insert :: "('a::linorder) list \ 'a trie \ '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 \ Nd False (\_. None) | Some t \ t))))" + Nd b (m(x := Some(insert xs (case m x of None \ empty | Some t \ t))))" fun delete :: "('a::linorder) list \ 'a trie \ 'a trie" where "delete [] (Nd b m) = Nd False m" | @@ -73,7 +76,7 @@ qed interpretation S: Set -where empty = "Nd False (\_. 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 = "\_. True" proof (standard, goal_cases) case 1 show ?case by (simp) diff -r a8238fd25541 -r 0b813a1a833f src/HOL/Data_Structures/Trie_Map.thy --- 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 \ 'a list \ bool" where "isin (Nd b m) [] = b" | "isin (Nd b m) (x # xs) = (case lookup m x of None \ False | Some t \ isin t xs)" @@ -36,7 +40,7 @@ fun insert :: "('a::linorder) list \ 'a trie_map \ '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 \ Nd False Leaf | Some t \ t)) m)" + Nd b (update x (insert xs (case lookup m x of None \ empty | Some t \ t)) m)" fun delete :: "('a::linorder) list \ 'a trie_map \ 'a trie_map" where "delete [] (Nd b m) = Nd False m" | @@ -84,7 +88,7 @@ text \Overall correctness w.r.t. the \Set\ ADT:\ 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)