--- 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 \<open>A trie where each node maps a key to sub-tries via a function.
Nice abstract model. Not efficient because of the function space.\<close>
-datatype 'a trie = Lf | Nd bool "'a \<Rightarrow> 'a trie option"
+datatype 'a trie = Nd bool "'a \<Rightarrow> 'a trie option"
fun isin :: "'a trie \<Rightarrow> 'a list \<Rightarrow> bool" where
-"isin Lf xs = False" |
"isin (Nd b m) [] = b" |
"isin (Nd b m) (k # xs) = (case m k of None \<Rightarrow> False | Some t \<Rightarrow> isin t xs)"
fun insert :: "('a::linorder) list \<Rightarrow> 'a trie \<Rightarrow> 'a trie" where
-"insert [] Lf = Nd True (\<lambda>x. None)" |
"insert [] (Nd b m) = Nd True m" |
-"insert (x#xs) Lf = Nd False ((\<lambda>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 \<Rightarrow> Lf | Some t \<Rightarrow> t))))"
+ Nd b (m(x := Some(insert xs (case m x of None \<Rightarrow> Nd False (\<lambda>_. None) | Some t \<Rightarrow> t))))"
fun delete :: "('a::linorder) list \<Rightarrow> 'a trie \<Rightarrow> '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 \<Rightarrow> m |
Some t \<Rightarrow> m(x := Some(delete xs t)))"
+text \<open>The actual definition of \<open>set\<close> is a bit cryptic but canonical, to enable
+primrec to prove termination:\<close>
+
primrec set :: "'a trie \<Rightarrow> 'a list set" where
-"set Lf = {}" |
"set (Nd b m) = (if b then {[]} else {}) \<union>
(\<Union>a. case (map_option set o m) a of None \<Rightarrow> {} | Some t \<Rightarrow> (#) a ` t)"
+text \<open>This is the more human-readable version:\<close>
+
lemma set_Nd:
"set (Nd b m) =
(if b then {[]} else {}) \<union>
@@ -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 (\<lambda>_. None)" 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 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 \<open>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 \<Rightarrow> 'a list \<Rightarrow> bool" where
-"isin Lf xs = False" |
"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)"
fun insert :: "('a::linorder) list \<Rightarrow> 'a trie_map \<Rightarrow> '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 \<Rightarrow> Lf | Some t \<Rightarrow> t)) m)"
+ Nd b (update x (insert xs (case lookup m x of None \<Rightarrow> Nd False Leaf | Some t \<Rightarrow> t)) m)"
fun delete :: "('a::linorder) list \<Rightarrow> 'a trie_map \<Rightarrow> '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 \<open>Proof by stepwise refinement. First abstract to type @{typ "'a trie"}.\<close>
fun abs :: "'a::linorder trie_map \<Rightarrow> 'a trie" where
-"abs Lf = Trie_Fun.Lf" |
-"abs (Nd b t) = Trie_Fun.Nd b (\<lambda>a. map_option abs (lookup t a))"
+"abs (Nd b t) = Trie_Fun0.Nd b (\<lambda>a. map_option abs (lookup t a))"
fun invar :: "('a::linorder)trie_map \<Rightarrow> bool" where
-"invar Lf = True" |
"invar (Nd b m) = (M.invar m \<and> (\<forall>a t. lookup m a = Some t \<longrightarrow> 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 \<Longrightarrow> abs(insert xs t) = Trie_Fun.insert xs (abs t)"
+lemma abs_insert: "invar t \<Longrightarrow> 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 \<Longrightarrow> abs(delete xs t) = Trie_Fun.delete xs (abs t)"
+lemma abs_delete: "invar t \<Longrightarrow> 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 \<Longrightarrow> 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 \<Longrightarrow> invar (delete xs t)"
@@ -90,7 +84,7 @@
text \<open>Overall correctness w.r.t. the \<open>Set\<close> ADT:\<close>
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