src/HOL/Data_Structures/Trie_Fun.thy
author wenzelm
Mon, 13 Apr 2020 22:08:14 +0200
changeset 71751 abf3e80bd815
parent 70266 0b813a1a833f
child 71918 4e0a58818edc
permissions -rw-r--r--
tuned NEWS;
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
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    20
fun insert :: "('a::linorder) list \<Rightarrow> 'a trie \<Rightarrow> 'a trie" where
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) =
70266
nipkow
parents: 70262
diff changeset
    23
   Nd b (m(x := Some(insert xs (case m x of None \<Rightarrow> empty | Some t \<Rightarrow> t))))"
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    24
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    25
fun delete :: "('a::linorder) list \<Rightarrow> 'a trie \<Rightarrow> 'a trie" where
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
70262
e12779b8f5b6 simplified types
nipkow
parents: 70250
diff changeset
    32
text \<open>The actual definition of \<open>set\<close> is a bit cryptic but canonical, to enable
e12779b8f5b6 simplified types
nipkow
parents: 70250
diff changeset
    33
primrec to prove termination:\<close>
e12779b8f5b6 simplified types
nipkow
parents: 70250
diff changeset
    34
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    35
primrec set :: "'a trie \<Rightarrow> 'a list set" where
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    36
"set (Nd b m) = (if b then {[]} else {}) \<union>
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    37
    (\<Union>a. case (map_option set o m) a of None \<Rightarrow> {} | Some t \<Rightarrow> (#) a ` t)"
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    38
70262
e12779b8f5b6 simplified types
nipkow
parents: 70250
diff changeset
    39
text \<open>This is the more human-readable version:\<close>
e12779b8f5b6 simplified types
nipkow
parents: 70250
diff changeset
    40
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    41
lemma set_Nd:
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    42
  "set (Nd b m) =
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    43
     (if b then {[]} else {}) \<union>
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    44
     (\<Union>a. case m a of None \<Rightarrow> {} | Some t \<Rightarrow> (#) a ` set t)"
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    45
by (auto simp: split: option.splits)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    46
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    47
lemma isin_set: "isin t xs = (xs \<in> set t)"
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    48
apply(induction t xs rule: isin.induct)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    49
apply (auto split: option.split)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    50
done
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    51
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    52
lemma set_insert: "set (insert xs t) = set t \<union> {xs}"
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    53
proof(induction xs t rule: insert.induct)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    54
  case 1 thus ?case by simp
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    55
next
70262
e12779b8f5b6 simplified types
nipkow
parents: 70250
diff changeset
    56
  case 2
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    57
  thus ?case
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    58
    apply(simp)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    59
    apply(subst set_eq_iff)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    60
    apply(auto split!: if_splits option.splits)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    61
     apply fastforce
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    62
    by (metis imageI option.sel)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    63
qed
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    64
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    65
lemma set_delete: "set (delete xs t) = set t - {xs}"
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    66
proof(induction xs t rule: delete.induct)
70262
e12779b8f5b6 simplified types
nipkow
parents: 70250
diff changeset
    67
  case 1 thus ?case by (force split: option.splits)
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    68
next
70262
e12779b8f5b6 simplified types
nipkow
parents: 70250
diff changeset
    69
  case 2
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    70
  thus ?case
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    71
    apply (auto simp add: image_iff split!: if_splits option.splits)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    72
       apply blast
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    73
      apply (metis insertE insertI2 insert_Diff_single option.inject)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    74
     apply blast
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    75
    by (metis insertE insertI2 insert_Diff_single option.inject)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    76
qed
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    77
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    78
interpretation S: Set
70266
nipkow
parents: 70262
diff changeset
    79
where empty = empty and isin = isin and insert = insert and delete = delete
70250
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    80
and set = set and invar = "\<lambda>_. True"
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    81
proof (standard, goal_cases)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    82
  case 1 show ?case by (simp)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    83
next
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    84
  case 2 thus ?case by(simp add: isin_set)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    85
next
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    86
  case 3 thus ?case by(simp add: set_insert)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    87
next
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    88
  case 4 thus ?case by(simp add: set_delete)
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    89
qed (rule TrueI)+
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    90
20d819b0a29d New version of tries
nipkow
parents:
diff changeset
    91
end