src/HOL/Data_Structures/Trie.thy
author nipkow
Tue, 26 Jun 2018 22:39:06 +0200
changeset 68516 b0c4a34ccfef
permissions -rw-r--r--
new theory Trie
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68516
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
     2
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
     3
section \<open>Trie and Patricia Trie Implementations of \mbox{\<open>bool list set\<close>}\<close>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
     4
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
     5
theory Trie
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
     6
imports Set_Specs
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
     7
begin
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
     8
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
     9
hide_const (open) insert
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    10
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    11
declare Let_def[simp]
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    12
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    13
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    14
subsection "Trie"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    15
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    16
datatype trie = Leaf | Node bool "trie * trie"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    17
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    18
text \<open>The pairing allows things like \<open>Node b (if \<dots> then (l,r) else (r,l))\<close>.\<close>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    19
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    20
fun isin :: "trie \<Rightarrow> bool list \<Rightarrow> bool" where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    21
"isin Leaf ks = False" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    22
"isin (Node b (l,r)) ks =
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    23
   (case ks of
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    24
      [] \<Rightarrow> b |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    25
      k#ks \<Rightarrow> isin (if k then r else l) ks)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    26
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    27
fun insert :: "bool list \<Rightarrow> trie \<Rightarrow> trie" where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    28
"insert [] Leaf = Node True (Leaf,Leaf)" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    29
"insert [] (Node b lr) = Node True lr" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    30
"insert (k#ks) Leaf =
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    31
  Node False (if k then (Leaf, insert ks Leaf)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    32
                   else (insert ks Leaf, Leaf))" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    33
"insert (k#ks) (Node b (l,r)) =
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    34
  Node b (if k then (l, insert ks r)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    35
               else (insert ks l, r))"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    36
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    37
fun shrink_Node :: "bool \<Rightarrow> trie * trie \<Rightarrow> trie" where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    38
"shrink_Node b lr = (if \<not> b \<and> lr = (Leaf,Leaf) then Leaf else Node b lr)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    39
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    40
fun delete :: "bool list \<Rightarrow> trie \<Rightarrow> trie" where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    41
"delete ks Leaf = Leaf" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    42
"delete ks (Node b (l,r)) =
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    43
   (case ks of
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    44
      [] \<Rightarrow> shrink_Node False (l,r) |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    45
      k#ks' \<Rightarrow> shrink_Node b (if k then (l, delete ks' r) else (delete ks' l, r)))"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    46
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    47
fun invar :: "trie \<Rightarrow> bool" where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    48
"invar Leaf = True" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    49
"invar (Node b (l,r)) = ((\<not> b \<longrightarrow> l \<noteq> Leaf \<or> r \<noteq> Leaf) \<and> invar l \<and> invar r)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    50
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    51
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    52
subsubsection \<open>Functional Correctness\<close>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    53
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    54
lemma isin_insert: "isin (insert as t) bs = (as = bs \<or> isin t bs)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    55
apply(induction as t arbitrary: bs rule: insert.induct)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    56
apply(auto split: list.splits)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    57
done
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    58
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    59
lemma isin_delete: "isin (delete as t) bs = (as \<noteq> bs \<and> isin t bs)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    60
apply(induction as t arbitrary: bs rule: delete.induct)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    61
 apply simp
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    62
apply (auto split: list.splits; fastforce)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    63
done
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    64
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    65
lemma insert_not_Leaf: "insert ks t \<noteq> Leaf"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    66
by(cases "(ks,t)" rule: insert.cases) auto
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    67
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    68
lemma invar_insert: "invar t \<Longrightarrow> invar (insert ks t)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    69
by(induction ks t rule: insert.induct)(auto simp: insert_not_Leaf)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    70
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    71
lemma invar_delete: "invar t \<Longrightarrow> invar (delete ks t)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    72
by(induction ks t rule: delete.induct)(auto split: list.split)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    73
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    74
interpretation T: Set
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    75
where empty = Leaf and insert = insert and delete = delete and isin = isin
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    76
and set = "\<lambda>t. {x. isin t x}" and invar = invar
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    77
proof (standard, goal_cases)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    78
  case 1 thus ?case by simp
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    79
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    80
  case 2 thus ?case by simp
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    81
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    82
  case 3 thus ?case by (auto simp add: isin_insert)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    83
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    84
  case 4 thus ?case by (auto simp add: isin_delete)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    85
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    86
  case 5 thus ?case by simp
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    87
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    88
  case 6 thus ?case by (auto simp add: invar_insert)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    89
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    90
  case 7 thus ?case by (auto simp add: invar_delete)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    91
qed
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    92
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    93
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    94
subsection "Patricia Trie"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    95
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    96
datatype trieP = LeafP | NodeP "bool list" bool "trieP * trieP"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    97
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    98
fun isinP :: "trieP \<Rightarrow> bool list \<Rightarrow> bool" where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
    99
"isinP LeafP ks = False" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   100
"isinP (NodeP ps b (l,r)) ks =
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   101
  (let n = length ps in
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   102
   if ps = take n ks
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   103
   then case drop n ks of [] \<Rightarrow> b | k#ks' \<Rightarrow> isinP (if k then r else l) ks'
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   104
   else False)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   105
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   106
text \<open>\<open>split xs ys = (zs, xs', ys')\<close> iff
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   107
  \<open>zs\<close> is the longest common prefix of \<open>xs\<close> and \<open>ys\<close> and
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   108
  \<open>xs = zs @ xs'\<close> and \<open>ys = zs @ ys'\<close>\<close>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   109
fun split where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   110
"split [] ys = ([],[],ys)" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   111
"split xs [] = ([],xs,[])" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   112
"split (x#xs) (y#ys) =
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   113
  (if x\<noteq>y then ([],x#xs,y#ys)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   114
   else let (ps,xs',ys') = split xs ys in (x#ps,xs',ys'))"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   115
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   116
fun insertP :: "bool list \<Rightarrow> trieP \<Rightarrow> trieP" where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   117
"insertP ks LeafP  = NodeP ks True (LeafP,LeafP)" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   118
"insertP ks (NodeP ps b (l,r)) =
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   119
  (case split ks ps of
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   120
     (qs,k#ks',p#ps') \<Rightarrow>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   121
       let tp = NodeP ps' b (l,r); tk = NodeP ks' True (LeafP,LeafP)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   122
       in NodeP qs False (if k then (tp,tk) else (tk,tp)) |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   123
     (qs,k#ks',[]) \<Rightarrow>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   124
       NodeP ps b (if k then (l,insertP ks' r) else (insertP ks' l, r)) |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   125
     (qs,[],p#ps') \<Rightarrow>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   126
       let t = NodeP ps' b (l,r)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   127
       in NodeP qs True (if p then (LeafP, t) else (t, LeafP)) |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   128
     (qs,[],[]) \<Rightarrow> NodeP ps True (l,r))"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   129
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   130
fun shrink_NodeP where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   131
"shrink_NodeP ps b lr = (if b then NodeP ps b lr else
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   132
  case lr of
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   133
     (LeafP, LeafP) \<Rightarrow> LeafP |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   134
     (LeafP, NodeP ps' b' lr') \<Rightarrow> NodeP (ps @ True # ps') b' lr' |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   135
     (NodeP ps' b' lr', LeafP) \<Rightarrow> NodeP (ps @ False # ps') b' lr' |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   136
     _ \<Rightarrow> NodeP ps b lr)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   137
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   138
fun deleteP :: "bool list \<Rightarrow> trieP \<Rightarrow> trieP" where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   139
"deleteP ks LeafP  = LeafP" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   140
"deleteP ks (NodeP ps b (l,r)) =
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   141
  (case split ks ps of
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   142
     (qs,_,p#ps') \<Rightarrow> NodeP ps b (l,r) |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   143
     (qs,k#ks',[]) \<Rightarrow>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   144
       shrink_NodeP ps b (if k then (l, deleteP ks' r) else (deleteP ks' l, r)) |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   145
     (qs,[],[]) \<Rightarrow> shrink_NodeP ps False (l,r))"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   146
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   147
fun invarP :: "trieP \<Rightarrow> bool" where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   148
"invarP LeafP = True" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   149
"invarP (NodeP ps b (l,r)) = ((\<not>b \<longrightarrow> l \<noteq> LeafP \<or> r \<noteq> LeafP) \<and> invarP l \<and> invarP r)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   150
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   151
text \<open>The abstraction function(s):\<close>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   152
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   153
fun prefix_trie :: "bool list \<Rightarrow> trie \<Rightarrow> trie" where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   154
"prefix_trie [] t = t" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   155
"prefix_trie (k#ks) t =
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   156
  (let t' = prefix_trie ks t in shrink_Node False (if k then (Leaf,t') else (t',Leaf)))"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   157
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   158
fun abs_trieP :: "trieP \<Rightarrow> trie" where
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   159
"abs_trieP LeafP = Leaf" |
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   160
"abs_trieP (NodeP ps b (l,r)) = prefix_trie ps (Node b (abs_trieP l, abs_trieP r))"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   161
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   162
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   163
subsubsection \<open>Functional Correctness\<close>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   164
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   165
text \<open>IsinP:\<close>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   166
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   167
lemma isin_prefix_trie: "isin (prefix_trie ps t) ks =
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   168
  (length ks \<ge> length ps \<and>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   169
  (let n = length ps in ps = take n ks \<and> isin t (drop n ks)))"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   170
by(induction ps arbitrary: ks)(auto split: list.split)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   171
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   172
lemma isinP: "isinP t ks = isin (abs_trieP t) ks"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   173
apply(induction t arbitrary: ks rule: abs_trieP.induct)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   174
 apply(auto simp: isin_prefix_trie split: list.split)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   175
 using nat_le_linear apply force
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   176
using nat_le_linear apply force
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   177
done
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   178
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   179
text \<open>Insert:\<close>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   180
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   181
lemma prefix_trie_Leaf_iff: "prefix_trie ps t = Leaf \<longleftrightarrow> t = Leaf"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   182
by (induction ps) auto
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   183
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   184
lemma prefix_trie_Leafs: "prefix_trie ks (Node True (Leaf,Leaf)) = insert ks Leaf"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   185
by(induction ks) (auto simp: prefix_trie_Leaf_iff)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   186
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   187
lemma prefix_trie_Leaf: "prefix_trie ps Leaf = Leaf"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   188
by(induction ps) auto
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   189
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   190
lemma insert_append: "insert (ks @ ks') (prefix_trie ks t) = prefix_trie ks (insert ks' t)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   191
by(induction ks) (auto simp: prefix_trie_Leaf_iff insert_not_Leaf prefix_trie_Leaf)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   192
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   193
lemma prefix_trie_append: "prefix_trie (ps @ qs) t = prefix_trie ps (prefix_trie qs t)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   194
by(induction ps) auto
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   195
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   196
lemma split_iff: "split xs ys = (zs, xs', ys') \<longleftrightarrow>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   197
  xs = zs @ xs' \<and> ys = zs @ ys' \<and> (xs' \<noteq> [] \<and> ys' \<noteq> [] \<longrightarrow> hd xs' \<noteq> hd ys')"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   198
proof(induction xs ys arbitrary: zs xs' ys' rule: split.induct)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   199
  case 1 thus ?case by auto
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   200
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   201
  case 2 thus ?case by auto
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   202
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   203
  case 3 thus ?case by(clarsimp simp: Cons_eq_append_conv split: prod.splits if_splits) auto
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   204
qed
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   205
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   206
lemma abs_trieP_insertP:
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   207
  "abs_trieP (insertP ks t) = insert ks (abs_trieP t)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   208
apply(induction t arbitrary: ks)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   209
apply(auto simp: prefix_trie_Leafs insert_append prefix_trie_append
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   210
        prefix_trie_Leaf_iff split_iff split: list.split prod.split)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   211
done
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   212
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   213
corollary isinP_insertP: "isinP (insertP ks t) ks' = (ks=ks' \<or> isinP t ks')"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   214
by (simp add: isin_insert isinP abs_trieP_insertP)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   215
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   216
lemma insertP_not_LeafP: "insertP ks t \<noteq> LeafP"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   217
apply(induction ks t rule: insertP.induct)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   218
apply(auto split: prod.split list.split)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   219
done
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   220
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   221
lemma invarP_insertP: "invarP t \<Longrightarrow> invarP (insertP ks t)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   222
apply(induction ks t rule: insertP.induct)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   223
apply(auto simp: insertP_not_LeafP split: prod.split list.split)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   224
done
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   225
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   226
text \<open>Delete:\<close>
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   227
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   228
lemma invar_shrink_NodeP: "\<lbrakk> invarP l; invarP r \<rbrakk> \<Longrightarrow> invarP (shrink_NodeP ps b (l,r))"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   229
by(auto split: trieP.split)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   230
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   231
lemma invarP_deleteP: "invarP t \<Longrightarrow> invarP (deleteP ks t)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   232
apply(induction t arbitrary: ks)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   233
apply(auto simp: invar_shrink_NodeP split_iff simp del: shrink_NodeP.simps
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   234
           split!: list.splits prod.split if_splits)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   235
done
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   236
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   237
lemma delete_append:
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   238
  "delete (ks @ ks') (prefix_trie ks t) = prefix_trie ks (delete ks' t)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   239
by(induction ks) auto
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   240
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   241
lemma abs_trieP_shrink_NodeP:
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   242
  "abs_trieP (shrink_NodeP ps b (l,r)) = prefix_trie ps (shrink_Node b (abs_trieP l, abs_trieP r))"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   243
apply(induction ps arbitrary: b l r)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   244
apply (auto simp: prefix_trie_Leaf prefix_trie_Leaf_iff prefix_trie_append
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   245
            split!: trieP.splits if_splits)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   246
done
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   247
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   248
lemma abs_trieP_deleteP:
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   249
  "abs_trieP (deleteP ks t) = delete ks (abs_trieP t)"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   250
apply(induction t arbitrary: ks)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   251
apply(auto simp: prefix_trie_Leafs delete_append prefix_trie_Leaf
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   252
                 abs_trieP_shrink_NodeP prefix_trie_append split_iff
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   253
           simp del: shrink_NodeP.simps split!: list.split prod.split if_splits)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   254
done
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   255
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   256
corollary isinP_deleteP: "isinP (deleteP ks t) ks' = (ks\<noteq>ks' \<and> isinP t ks')"
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   257
by (simp add: isin_delete isinP abs_trieP_deleteP)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   258
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   259
interpretation PT: Set
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   260
where empty = LeafP and insert = insertP and delete = deleteP
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   261
and isin = isinP and set = "\<lambda>t. {x. isinP t x}" and invar = invarP
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   262
proof (standard, goal_cases)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   263
  case 1 thus ?case by (simp)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   264
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   265
  case 2 thus ?case by (simp)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   266
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   267
  case 3 thus ?case by (auto simp add: isinP_insertP)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   268
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   269
  case 4 thus ?case by (auto simp add: isinP_deleteP)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   270
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   271
  case 5 thus ?case by (simp)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   272
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   273
  case 6 thus ?case by (simp add: invarP_insertP)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   274
next
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   275
  case 7 thus ?case by (auto simp add: invarP_deleteP)
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   276
qed
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   277
b0c4a34ccfef new theory Trie
nipkow
parents:
diff changeset
   278
end