src/HOL/Data_Structures/Trie_Fun.thy
changeset 70262 e12779b8f5b6
parent 70250 20d819b0a29d
child 70266 0b813a1a833f
--- 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)