src/HOL/Data_Structures/Trie_Map.thy
author nipkow
Thu, 18 Jan 2024 14:30:27 +0100
changeset 79494 c7536609bb9b
parent 77642 a28ee8058ea3
permissions -rw-r--r--
translation to time functions now with canonical let.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     1
section "Tries via Search Trees"
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     2
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     3
theory Trie_Map
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     4
imports
77642
a28ee8058ea3 use tree (simpler) instead of rbt (exercise)
nipkow
parents: 77266
diff changeset
     5
  Tree_Map
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     6
  Trie_Fun
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     7
begin
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     8
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
     9
text \<open>An implementation of tries for an arbitrary alphabet \<open>'a\<close> where
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    10
the mapping from an element of type \<open>'a\<close> to the sub-trie is implemented by a binary search tree.
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    11
Although this implementation uses maps implemented by red-black trees it works for any
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    12
implementation of maps.
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    13
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    14
This is an implementation of the ``ternary search trees'' by Bentley and Sedgewick
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    15
[SODA 1997, Dr. Dobbs 1998]. The name derives from the fact that a node in the BST can now
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    16
be drawn to have 3 children, where the middle child is the sub-trie that the node maps
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    17
its key to. Hence the name \<open>trie3\<close>.
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    18
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    19
Example from @{url "https://en.wikipedia.org/wiki/Ternary_search_tree#Description"}:
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    20
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    21
          c
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    22
        / | \
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    23
       a  u  h
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    24
       |  |  | \
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    25
       t. t  e. u
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    26
     /  / |   / |
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    27
    s. p. e. i. s.
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    28
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    29
Characters with a dot are final.
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    30
Thus the tree represents the set of strings "cute","cup","at","as","he","us" and "i".
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    31
\<close>
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    32
77642
a28ee8058ea3 use tree (simpler) instead of rbt (exercise)
nipkow
parents: 77266
diff changeset
    33
datatype 'a trie3 = Nd3 bool "('a * 'a trie3) tree"
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    34
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    35
text \<open>In principle one should be able to given an implementation of tries
77642
a28ee8058ea3 use tree (simpler) instead of rbt (exercise)
nipkow
parents: 77266
diff changeset
    36
once and for all for any map implementation and not just for a specific one (unbalanced trees) as done here.
a28ee8058ea3 use tree (simpler) instead of rbt (exercise)
nipkow
parents: 77266
diff changeset
    37
But because the map (@{type tree}) is used in a datatype, the HOL type system does not support this.
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    38
77642
a28ee8058ea3 use tree (simpler) instead of rbt (exercise)
nipkow
parents: 77266
diff changeset
    39
However, the development below works verbatim for any map implementation, eg \<open>RBT_Map\<close>,
a28ee8058ea3 use tree (simpler) instead of rbt (exercise)
nipkow
parents: 77266
diff changeset
    40
and not just \<open>Tree_Map\<close>, except for the termination lemma \<open>lookup_size\<close>.\<close>
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    41
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70266
diff changeset
    42
term size_tree
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    43
lemma lookup_size[termination_simp]:
77642
a28ee8058ea3 use tree (simpler) instead of rbt (exercise)
nipkow
parents: 77266
diff changeset
    44
  fixes t :: "('a::linorder * 'a trie3) tree"
a28ee8058ea3 use tree (simpler) instead of rbt (exercise)
nipkow
parents: 77266
diff changeset
    45
  shows "lookup t a = Some b \<Longrightarrow> size b < Suc (size_tree (\<lambda>ab. Suc (size (snd( ab)))) t)"
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    46
apply(induction t a rule: lookup.induct)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    47
apply(auto split: if_splits)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    48
done
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    49
70266
nipkow
parents: 70263
diff changeset
    50
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    51
definition empty3 :: "'a trie3" where
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    52
[simp]: "empty3 = Nd3 False Leaf"
70266
nipkow
parents: 70263
diff changeset
    53
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    54
fun isin3 :: "('a::linorder) trie3 \<Rightarrow> 'a list \<Rightarrow> bool" where
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    55
"isin3 (Nd3 b m) [] = b" |
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    56
"isin3 (Nd3 b m) (x # xs) = (case lookup m x of None \<Rightarrow> False | Some t \<Rightarrow> isin3 t xs)"
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    57
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    58
fun insert3 :: "('a::linorder) list \<Rightarrow> 'a trie3 \<Rightarrow> 'a trie3" where
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    59
"insert3 [] (Nd3 b m) = Nd3 True m" |
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    60
"insert3 (x#xs) (Nd3 b m) =
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    61
  Nd3 b (update x (insert3 xs (case lookup m x of None \<Rightarrow> empty3 | Some t \<Rightarrow> t)) m)"
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    62
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    63
fun delete3 :: "('a::linorder) list \<Rightarrow> 'a trie3 \<Rightarrow> 'a trie3" where
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    64
"delete3 [] (Nd3 b m) = Nd3 False m" |
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    65
"delete3 (x#xs) (Nd3 b m) = Nd3 b
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    66
   (case lookup m x of
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    67
      None \<Rightarrow> m |
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    68
      Some t \<Rightarrow> update x (delete3 xs t) m)"
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    69
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    70
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    71
subsection "Correctness"
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    72
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    73
text \<open>Proof by stepwise refinement. First abs3tract to type @{typ "'a trie"}.\<close>
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    74
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    75
fun abs3 :: "'a::linorder trie3 \<Rightarrow> 'a trie" where
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    76
"abs3 (Nd3 b t) = Nd b (\<lambda>a. map_option abs3 (lookup t a))"
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    77
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    78
fun invar3 :: "('a::linorder)trie3 \<Rightarrow> bool" where
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    79
"invar3 (Nd3 b m) = (M.invar m \<and> (\<forall>a t. lookup m a = Some t \<longrightarrow> invar3 t))"
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    80
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    81
lemma isin_abs3: "isin3 t xs = isin (abs3 t) xs"
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    82
apply(induction t xs rule: isin3.induct)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    83
apply(auto split: option.split)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    84
done
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    85
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    86
lemma abs3_insert3: "invar3 t \<Longrightarrow> abs3(insert3 xs t) = insert xs (abs3 t)"
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    87
apply(induction xs t rule: insert3.induct)
77642
a28ee8058ea3 use tree (simpler) instead of rbt (exercise)
nipkow
parents: 77266
diff changeset
    88
apply(auto simp: M.map_specs Tree_Set.empty_def[symmetric] split: option.split)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    89
done
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    90
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    91
lemma abs3_delete3: "invar3 t \<Longrightarrow> abs3(delete3 xs t) = delete xs (abs3 t)"
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    92
apply(induction xs t rule: delete3.induct)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    93
apply(auto simp: M.map_specs split: option.split)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    94
done
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    95
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    96
lemma invar3_insert3: "invar3 t \<Longrightarrow> invar3 (insert3 xs t)"
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
    97
apply(induction xs t rule: insert3.induct)
77642
a28ee8058ea3 use tree (simpler) instead of rbt (exercise)
nipkow
parents: 77266
diff changeset
    98
apply(auto simp: M.map_specs Tree_Set.empty_def[symmetric] split: option.split)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    99
done
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   100
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
   101
lemma invar3_delete3: "invar3 t \<Longrightarrow> invar3 (delete3 xs t)"
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
   102
apply(induction xs t rule: delete3.induct)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   103
apply(auto simp: M.map_specs split: option.split)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   104
done
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   105
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   106
text \<open>Overall correctness w.r.t. the \<open>Set\<close> ADT:\<close>
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   107
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   108
interpretation S2: Set
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
   109
where empty = empty3 and isin = isin3 and insert = insert3 and delete = delete3
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
   110
and set = "set o abs3" and invar = invar3
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   111
proof (standard, goal_cases)
72979
nipkow
parents: 70755
diff changeset
   112
  case 1 show ?case by (simp add: isin_case split: list.split)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   113
next
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
   114
  case 2 thus ?case by (simp add: isin_abs3)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   115
next
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
   116
  case 3 thus ?case by (simp add: set_insert abs3_insert3 del: set_def)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   117
next
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
   118
  case 4 thus ?case by (simp add: set_delete abs3_delete3 del: set_def)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   119
next
77642
a28ee8058ea3 use tree (simpler) instead of rbt (exercise)
nipkow
parents: 77266
diff changeset
   120
  case 5 thus ?case by (simp add: M.map_specs Tree_Set.empty_def[symmetric])
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   121
next
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
   122
  case 6 thus ?case by (simp add: invar3_insert3)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   123
next
77266
334015f9098e removed Map from docu
nipkow
parents: 72979
diff changeset
   124
  case 7 thus ?case by (simp add: invar3_delete3)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   125
qed
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   126
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
   127
end