src/HOL/Data_Structures/Trie_Fun.thy
author nipkow
Tue, 22 Dec 2020 23:36:32 +0100
changeset 72979 e734cd65c926
parent 72912 fa364c21df15
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     1
section \<open>Tries via Functions\<close>
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     2
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     3
theory Trie_Fun
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     4
imports
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     5
  Set_Specs
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     6
begin
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     7
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     8
text \<open>A trie where each node maps a key to sub-tries via a function.
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
     9
Nice abstract model. Not efficient because of the function space.\<close>
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    10
70262
e12779b8f5b6 simplified types
nipkow
parents: 70250
diff changeset
    11
datatype 'a trie = Nd bool "'a \<Rightarrow> 'a trie option"
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    12
70266
nipkow
parents: 70262
diff changeset
    13
definition empty :: "'a trie" where
nipkow
parents: 70262
diff changeset
    14
[simp]: "empty = Nd False (\<lambda>_. None)"
nipkow
parents: 70262
diff changeset
    15
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    16
fun isin :: "'a trie \<Rightarrow> 'a list \<Rightarrow> bool" where
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    17
"isin (Nd b m) [] = b" |
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    18
"isin (Nd b m) (k # xs) = (case m k of None \<Rightarrow> False | Some t \<Rightarrow> isin t xs)"
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    19
72912
nipkow
parents: 71918
diff changeset
    20
fun insert :: "'a list \<Rightarrow> 'a trie \<Rightarrow> 'a trie" where
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    21
"insert [] (Nd b m) = Nd True m" |
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    22
"insert (x#xs) (Nd b m) =
72979
nipkow
parents: 72912
diff changeset
    23
   (let s = (case m x of None \<Rightarrow> empty | Some t \<Rightarrow> t) in Nd b (m(x := Some(insert xs s))))"
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    24
72912
nipkow
parents: 71918
diff changeset
    25
fun delete :: "'a list \<Rightarrow> 'a trie \<Rightarrow> 'a trie" where
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    26
"delete [] (Nd b m) = Nd False m" |
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    27
"delete (x#xs) (Nd b m) = Nd b
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    28
   (case m x of
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    29
      None \<Rightarrow> m |
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    30
      Some t \<Rightarrow> m(x := Some(delete xs t)))"
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    31
72979
nipkow
parents: 72912
diff changeset
    32
text \<open>Use (a tuned version of) @{const isin} as an abstraction function:\<close>
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    33
72979
nipkow
parents: 72912
diff changeset
    34
lemma isin_case: "isin (Nd b m) xs =
nipkow
parents: 72912
diff changeset
    35
  (case xs of
nipkow
parents: 72912
diff changeset
    36
   [] \<Rightarrow> b |
nipkow
parents: 72912
diff changeset
    37
   x # ys \<Rightarrow> (case m x of None \<Rightarrow> False | Some t \<Rightarrow> isin t ys))"
nipkow
parents: 72912
diff changeset
    38
by(cases xs)auto
70262
e12779b8f5b6 simplified types
nipkow
parents: 70250
diff changeset
    39
72979
nipkow
parents: 72912
diff changeset
    40
definition set :: "'a trie \<Rightarrow> 'a list set" where
nipkow
parents: 72912
diff changeset
    41
[simp]: "set t = {xs. isin t xs}"
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    42
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    43
lemma isin_set: "isin t xs = (xs \<in> set t)"
72979
nipkow
parents: 72912
diff changeset
    44
by simp
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    45
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    46
lemma set_insert: "set (insert xs t) = set t \<union> {xs}"
72979
nipkow
parents: 72912
diff changeset
    47
by (induction xs t rule: insert.induct)
nipkow
parents: 72912
diff changeset
    48
   (auto simp: isin_case split!: if_splits option.splits list.splits)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    49
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    50
lemma set_delete: "set (delete xs t) = set t - {xs}"
72979
nipkow
parents: 72912
diff changeset
    51
by (induction xs t rule: delete.induct)
nipkow
parents: 72912
diff changeset
    52
   (auto simp: isin_case split!: if_splits option.splits list.splits)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    53
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    54
interpretation S: Set
70266
nipkow
parents: 70262
diff changeset
    55
where empty = empty and isin = isin and insert = insert and delete = delete
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    56
and set = set and invar = "\<lambda>_. True"
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    57
proof (standard, goal_cases)
72979
nipkow
parents: 72912
diff changeset
    58
  case 1 show ?case by (simp add: isin_case split: list.split)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    59
next
72979
nipkow
parents: 72912
diff changeset
    60
  case 2 show ?case by(rule isin_set)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    61
next
72979
nipkow
parents: 72912
diff changeset
    62
  case 3 show ?case by(rule set_insert)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    63
next
72979
nipkow
parents: 72912
diff changeset
    64
  case 4 show ?case by(rule set_delete)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    65
qed (rule TrueI)+
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    66
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    67
end