tuned
authornipkow
Tue, 22 Dec 2020 23:36:32 +0100
changeset 72979 e734cd65c926
parent 72978 7e7ed27fe625
child 72980 4fc3dc37f406
tuned
src/HOL/Data_Structures/Trie_Fun.thy
src/HOL/Data_Structures/Trie_Map.thy
--- a/src/HOL/Data_Structures/Trie_Fun.thy	Mon Dec 21 23:22:14 2020 +0100
+++ b/src/HOL/Data_Structures/Trie_Fun.thy	Tue Dec 22 23:36:32 2020 +0100
@@ -20,7 +20,7 @@
 fun insert :: "'a 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> empty | Some t \<Rightarrow> t))))"
+   (let s = (case m x of None \<Rightarrow> empty | Some t \<Rightarrow> t) in Nd b (m(x := Some(insert xs s))))"
 
 fun delete :: "'a list \<Rightarrow> 'a trie \<Rightarrow> 'a trie" where
 "delete [] (Nd b m) = Nd False m" |
@@ -29,62 +29,39 @@
       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 (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>Use (a tuned version of) @{const isin} as an abstraction function:\<close>
 
-text \<open>This is the more human-readable version:\<close>
+lemma isin_case: "isin (Nd b m) xs =
+  (case xs of
+   [] \<Rightarrow> b |
+   x # ys \<Rightarrow> (case m x of None \<Rightarrow> False | Some t \<Rightarrow> isin t ys))"
+by(cases xs)auto
 
-lemma set_Nd:
-  "set (Nd b m) =
-     (if b then {[]} else {}) \<union>
-     (\<Union>a. case m a of None \<Rightarrow> {} | Some t \<Rightarrow> (#) a ` set t)"
-by (auto simp: split: option.splits)
+definition set :: "'a trie \<Rightarrow> 'a list set" where
+[simp]: "set t = {xs. isin t xs}"
 
 lemma isin_set: "isin t xs = (xs \<in> set t)"
-apply(induction t xs rule: isin.induct)
-apply (auto split: option.split)
-done
+by simp
 
 lemma set_insert: "set (insert xs t) = set t \<union> {xs}"
-proof(induction xs t rule: insert.induct)
-  case 1 thus ?case by simp
-next
-  case 2
-  thus ?case
-    apply(simp)
-    apply(subst set_eq_iff)
-    apply(auto split!: if_splits option.splits)
-     apply fastforce
-    by (metis imageI option.sel)
-qed
+by (induction xs t rule: insert.induct)
+   (auto simp: isin_case split!: if_splits option.splits list.splits)
 
 lemma set_delete: "set (delete xs t) = set t - {xs}"
-proof(induction xs t rule: delete.induct)
-  case 1 thus ?case by (force split: option.splits)
-next
-  case 2
-  show ?case
-    apply (auto simp add: image_iff 2 split!: if_splits option.splits)
-     apply (metis DiffI empty_iff insert_iff option.inject)
-     apply (metis DiffI empty_iff insert_iff option.inject)
-    done
-qed
+by (induction xs t rule: delete.induct)
+   (auto simp: isin_case split!: if_splits option.splits list.splits)
 
 interpretation S: Set
 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)
+  case 1 show ?case by (simp add: isin_case split: list.split)
 next
-  case 2 thus ?case by(simp add: isin_set)
+  case 2 show ?case by(rule isin_set)
 next
-  case 3 thus ?case by(simp add: set_insert)
+  case 3 show ?case by(rule set_insert)
 next
-  case 4 thus ?case by(simp add: set_delete)
+  case 4 show ?case by(rule set_delete)
 qed (rule TrueI)+
 
 end
--- a/src/HOL/Data_Structures/Trie_Map.thy	Mon Dec 21 23:22:14 2020 +0100
+++ b/src/HOL/Data_Structures/Trie_Map.thy	Tue Dec 22 23:36:32 2020 +0100
@@ -91,13 +91,13 @@
 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)
+  case 1 show ?case by (simp add: isin_case split: list.split)
 next
-  case 2 thus ?case by (simp add: isin_set isin_abs)
+  case 2 thus ?case by (simp add: isin_abs)
 next
-  case 3 thus ?case by (simp add: set_insert abs_insert)
+  case 3 thus ?case by (simp add: set_insert abs_insert del: set_def)
 next
-  case 4 thus ?case by (simp add: set_delete abs_delete)
+  case 4 thus ?case by (simp add: set_delete abs_delete del: set_def)
 next
   case 5 thus ?case by (simp add: M.map_specs RBT_Set.empty_def[symmetric])
 next