# HG changeset patch # User nipkow # Date 1557581231 -7200 # Node ID e12779b8f5b65df011a887742bedaa24c3ec051e # Parent efbdfcaa6258356fc03e8a075f96c3a25d842e98 simplified types diff -r efbdfcaa6258 -r e12779b8f5b6 src/HOL/Data_Structures/Trie_Fun.thy --- a/src/HOL/Data_Structures/Trie_Fun.thy Fri May 10 11:20:02 2019 +0200 +++ b/src/HOL/Data_Structures/Trie_Fun.thy Sat May 11 15:27:11 2019 +0200 @@ -8,33 +8,33 @@ text \A trie where each node maps a key to sub-tries via a function. Nice abstract model. Not efficient because of the function space.\ -datatype 'a trie = Lf | Nd bool "'a \ 'a trie option" +datatype 'a trie = Nd bool "'a \ 'a trie option" fun isin :: "'a trie \ 'a list \ bool" where -"isin Lf xs = False" | "isin (Nd b m) [] = b" | "isin (Nd b m) (k # xs) = (case m k of None \ False | Some t \ isin t xs)" fun insert :: "('a::linorder) list \ 'a trie \ 'a trie" where -"insert [] Lf = Nd True (\x. None)" | "insert [] (Nd b m) = Nd True m" | -"insert (x#xs) Lf = Nd False ((\x. None)(x := Some(insert xs Lf)))" | "insert (x#xs) (Nd b m) = - Nd b (m(x := Some(insert xs (case m x of None \ Lf | Some t \ t))))" + Nd b (m(x := Some(insert xs (case m x of None \ Nd False (\_. None) | Some t \ t))))" fun delete :: "('a::linorder) list \ 'a trie \ 'a trie" where -"delete xs Lf = Lf" | "delete [] (Nd b m) = Nd False m" | "delete (x#xs) (Nd b m) = Nd b (case m x of None \ m | Some t \ m(x := Some(delete xs t)))" +text \The actual definition of \set\ is a bit cryptic but canonical, to enable +primrec to prove termination:\ + primrec set :: "'a trie \ 'a list set" where -"set Lf = {}" | "set (Nd b m) = (if b then {[]} else {}) \ (\a. case (map_option set o m) a of None \ {} | Some t \ (#) a ` t)" +text \This is the more human-readable version:\ + lemma set_Nd: "set (Nd b m) = (if b then {[]} else {}) \ @@ -50,11 +50,7 @@ proof(induction xs t rule: insert.induct) case 1 thus ?case by simp next - case 2 thus ?case by simp -next - case 3 thus ?case by simp (subst set_eq_iff, simp) -next - case 4 + case 2 thus ?case apply(simp) apply(subst set_eq_iff) @@ -65,11 +61,9 @@ lemma set_delete: "set (delete xs t) = set t - {xs}" proof(induction xs t rule: delete.induct) - case 1 thus ?case by simp + case 1 thus ?case by (force split: option.splits) next - case 2 thus ?case by (force split: option.splits) -next - case 3 + case 2 thus ?case apply (auto simp add: image_iff split!: if_splits option.splits) apply blast @@ -79,7 +73,7 @@ qed interpretation S: Set -where empty = Lf and isin = isin and insert = insert and delete = delete +where empty = "Nd False (\_. None)" 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 efbdfcaa6258 -r e12779b8f5b6 src/HOL/Data_Structures/Trie_Map.thy --- a/src/HOL/Data_Structures/Trie_Map.thy Fri May 10 11:20:02 2019 +0200 +++ b/src/HOL/Data_Structures/Trie_Map.thy Sat May 11 15:27:11 2019 +0200 @@ -13,7 +13,7 @@ type_synonym 'a mapi = "'a rbt" -datatype 'a trie_map = Lf | Nd bool "('a * 'a trie_map) mapi" +datatype 'a trie_map = Nd bool "('a * 'a trie_map) mapi" text \In principle one should be able to given an implementation of tries once and for all for any map implementation and not just for a specific one (RBT) as done here. @@ -30,19 +30,15 @@ done fun isin :: "('a::linorder) trie_map \ 'a list \ bool" where -"isin Lf xs = False" | "isin (Nd b m) [] = b" | "isin (Nd b m) (x # xs) = (case lookup m x of None \ False | Some t \ isin t xs)" fun insert :: "('a::linorder) list \ 'a trie_map \ 'a trie_map" where -"insert [] Lf = Nd True empty" | "insert [] (Nd b m) = Nd True m" | -"insert (x#xs) Lf = Nd False (update x (insert xs Lf) empty)" | "insert (x#xs) (Nd b m) = - Nd b (update x (insert xs (case lookup m x of None \ Lf | Some t \ t)) m)" + Nd b (update x (insert xs (case lookup m x of None \ Nd False Leaf | Some t \ t)) m)" fun delete :: "('a::linorder) list \ 'a trie_map \ 'a trie_map" where -"delete xs Lf = Lf" | "delete [] (Nd b m) = Nd False m" | "delete (x#xs) (Nd b m) = Nd b (case lookup m x of @@ -55,31 +51,29 @@ text \Proof by stepwise refinement. First abstract to type @{typ "'a trie"}.\ fun abs :: "'a::linorder trie_map \ 'a trie" where -"abs Lf = Trie_Fun.Lf" | -"abs (Nd b t) = Trie_Fun.Nd b (\a. map_option abs (lookup t a))" +"abs (Nd b t) = Trie_Fun0.Nd b (\a. map_option abs (lookup t a))" fun invar :: "('a::linorder)trie_map \ bool" where -"invar Lf = True" | "invar (Nd b m) = (M.invar m \ (\a t. lookup m a = Some t \ invar t))" -lemma isin_abs: "isin t xs = Trie_Fun.isin (abs t) xs" +lemma isin_abs: "isin t xs = Trie_Fun0.isin (abs t) xs" apply(induction t xs rule: isin.induct) apply(auto split: option.split) done -lemma abs_insert: "invar t \ abs(insert xs t) = Trie_Fun.insert xs (abs t)" +lemma abs_insert: "invar t \ abs(insert xs t) = Trie_Fun0.insert xs (abs t)" apply(induction xs t rule: insert.induct) -apply(auto simp: M.map_specs split: option.split) +apply(auto simp: M.map_specs RBT_Set.empty_def[symmetric] split: option.split) done -lemma abs_delete: "invar t \ abs(delete xs t) = Trie_Fun.delete xs (abs t)" +lemma abs_delete: "invar t \ abs(delete xs t) = Trie_Fun0.delete xs (abs t)" apply(induction xs t rule: delete.induct) apply(auto simp: M.map_specs split: option.split) done lemma invar_insert: "invar t \ invar (insert xs t)" apply(induction xs t rule: insert.induct) -apply(auto simp: M.map_specs split: option.split) +apply(auto simp: M.map_specs RBT_Set.empty_def[symmetric] split: option.split) done lemma invar_delete: "invar t \ invar (delete xs t)" @@ -90,7 +84,7 @@ text \Overall correctness w.r.t. the \Set\ ADT:\ interpretation S2: Set -where empty = Lf and isin = isin and insert = insert and delete = delete +where empty = "Nd False Leaf" 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) @@ -101,7 +95,7 @@ next case 4 thus ?case by (simp add: set_delete abs_delete) next - case 5 thus ?case by (simp) + case 5 thus ?case by (simp add: M.map_specs RBT_Set.empty_def[symmetric]) next case 6 thus ?case by (simp add: invar_insert) next