author | nipkow |
Thu, 18 Jan 2024 14:30:27 +0100 | |
changeset 79494 | c7536609bb9b |
parent 78653 | 7ed1759fe1bd |
permissions | -rw-r--r-- |
70250 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
3 |
section "Binary Tries and Patricia Tries" |
|
4 |
||
5 |
theory Tries_Binary |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
6 |
imports Set_Specs |
70250 | 7 |
begin |
8 |
||
9 |
hide_const (open) insert |
|
10 |
||
11 |
declare Let_def[simp] |
|
12 |
||
13 |
fun sel2 :: "bool \<Rightarrow> 'a * 'a \<Rightarrow> 'a" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
14 |
"sel2 b (a1,a2) = (if b then a2 else a1)" |
70250 | 15 |
|
16 |
fun mod2 :: "('a \<Rightarrow> 'a) \<Rightarrow> bool \<Rightarrow> 'a * 'a \<Rightarrow> 'a * 'a" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
17 |
"mod2 f b (a1,a2) = (if b then (a1,f a2) else (f a1,a2))" |
70250 | 18 |
|
19 |
||
20 |
subsection "Trie" |
|
21 |
||
22 |
datatype trie = Lf | Nd bool "trie * trie" |
|
23 |
||
70267 | 24 |
definition empty :: trie where |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
25 |
[simp]: "empty = Lf" |
70267 | 26 |
|
70250 | 27 |
fun isin :: "trie \<Rightarrow> bool list \<Rightarrow> bool" where |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
28 |
"isin Lf ks = False" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
29 |
"isin (Nd b lr) ks = |
70250 | 30 |
(case ks of |
31 |
[] \<Rightarrow> b | |
|
32 |
k#ks \<Rightarrow> isin (sel2 k lr) ks)" |
|
33 |
||
34 |
fun insert :: "bool list \<Rightarrow> trie \<Rightarrow> trie" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
35 |
"insert [] Lf = Nd True (Lf,Lf)" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
36 |
"insert [] (Nd b lr) = Nd True lr" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
37 |
"insert (k#ks) Lf = Nd False (mod2 (insert ks) k (Lf,Lf))" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
38 |
"insert (k#ks) (Nd b lr) = Nd b (mod2 (insert ks) k lr)" |
70250 | 39 |
|
70267 | 40 |
lemma isin_insert: "isin (insert xs t) ys = (xs = ys \<or> isin t ys)" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
41 |
proof (induction xs t arbitrary: ys rule: insert.induct) |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
42 |
qed (auto split: list.splits if_splits) |
70250 | 43 |
|
44 |
text \<open>A simple implementation of delete; does not shrink the trie!\<close> |
|
45 |
||
46 |
fun delete0 :: "bool list \<Rightarrow> trie \<Rightarrow> trie" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
47 |
"delete0 ks Lf = Lf" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
48 |
"delete0 ks (Nd b lr) = |
70250 | 49 |
(case ks of |
50 |
[] \<Rightarrow> Nd False lr | |
|
51 |
k#ks' \<Rightarrow> Nd b (mod2 (delete0 ks') k lr))" |
|
52 |
||
53 |
lemma isin_delete0: "isin (delete0 as t) bs = (as \<noteq> bs \<and> isin t bs)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
54 |
proof (induction as t arbitrary: bs rule: delete0.induct) |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
55 |
qed (auto split: list.splits if_splits) |
70250 | 56 |
|
57 |
text \<open>Now deletion with shrinking:\<close> |
|
58 |
||
59 |
fun node :: "bool \<Rightarrow> trie * trie \<Rightarrow> trie" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
60 |
"node b lr = (if \<not> b \<and> lr = (Lf,Lf) then Lf else Nd b lr)" |
70250 | 61 |
|
62 |
fun delete :: "bool list \<Rightarrow> trie \<Rightarrow> trie" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
63 |
"delete ks Lf = Lf" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
64 |
"delete ks (Nd b lr) = |
70250 | 65 |
(case ks of |
66 |
[] \<Rightarrow> node False lr | |
|
67 |
k#ks' \<Rightarrow> node b (mod2 (delete ks') k lr))" |
|
68 |
||
70267 | 69 |
lemma isin_delete: "isin (delete xs t) ys = (xs \<noteq> ys \<and> isin t ys)" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
70 |
apply(induction xs t arbitrary: ys rule: delete.induct) |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
71 |
apply (auto split: list.splits if_splits) |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
72 |
apply (metis isin.simps(1))+ |
70250 | 73 |
done |
74 |
||
75 |
definition set_trie :: "trie \<Rightarrow> bool list set" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
76 |
"set_trie t = {xs. isin t xs}" |
70250 | 77 |
|
70267 | 78 |
lemma set_trie_empty: "set_trie empty = {}" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
79 |
by(simp add: set_trie_def) |
70267 | 80 |
|
81 |
lemma set_trie_isin: "isin t xs = (xs \<in> set_trie t)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
82 |
by(simp add: set_trie_def) |
70267 | 83 |
|
70250 | 84 |
lemma set_trie_insert: "set_trie(insert xs t) = set_trie t \<union> {xs}" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
85 |
by(auto simp add: isin_insert set_trie_def) |
70250 | 86 |
|
70267 | 87 |
lemma set_trie_delete: "set_trie(delete xs t) = set_trie t - {xs}" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
88 |
by(auto simp add: isin_delete set_trie_def) |
70267 | 89 |
|
77830 | 90 |
text \<open>Invariant: tries are fully shrunk:\<close> |
91 |
fun invar where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
92 |
"invar Lf = True" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
93 |
"invar (Nd b (l,r)) = (invar l \<and> invar r \<and> (l = Lf \<and> r = Lf \<longrightarrow> b))" |
77830 | 94 |
|
95 |
lemma insert_Lf: "insert xs t \<noteq> Lf" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
96 |
using insert.elims by blast |
77830 | 97 |
|
98 |
lemma invar_insert: "invar t \<Longrightarrow> invar(insert xs t)" |
|
99 |
proof(induction xs t rule: insert.induct) |
|
100 |
case 1 thus ?case by simp |
|
101 |
next |
|
102 |
case (2 b lr) |
|
103 |
thus ?case by(cases lr; simp) |
|
104 |
next |
|
105 |
case (3 k ks) |
|
106 |
thus ?case by(simp; cases ks; auto) |
|
107 |
next |
|
108 |
case (4 k ks b lr) |
|
109 |
then show ?case by(cases lr; auto simp: insert_Lf) |
|
110 |
qed |
|
111 |
||
112 |
lemma invar_delete: "invar t \<Longrightarrow> invar(delete xs t)" |
|
113 |
proof(induction t arbitrary: xs) |
|
114 |
case Lf thus ?case by simp |
|
115 |
next |
|
116 |
case (Nd b lr) |
|
117 |
thus ?case by(cases lr)(auto split: list.split) |
|
118 |
qed |
|
119 |
||
70250 | 120 |
interpretation S: Set |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
121 |
where empty = empty and isin = isin and insert = insert and delete = delete |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
122 |
and set = set_trie and invar = invar |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
123 |
unfolding Set_def |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
124 |
by (smt (verit, best) Tries_Binary.empty_def invar.simps(1) invar_delete invar_insert set_trie_delete set_trie_empty set_trie_insert set_trie_isin) |
70250 | 125 |
|
126 |
||
127 |
subsection "Patricia Trie" |
|
128 |
||
70268 | 129 |
datatype trieP = LfP | NdP "bool list" bool "trieP * trieP" |
70250 | 130 |
|
77830 | 131 |
text \<open>Fully shrunk:\<close> |
132 |
fun invarP where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
133 |
"invarP LfP = True" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
134 |
"invarP (NdP ps b (l,r)) = (invarP l \<and> invarP r \<and> (l = LfP \<or> r = LfP \<longrightarrow> b))" |
77830 | 135 |
|
70268 | 136 |
fun isinP :: "trieP \<Rightarrow> bool list \<Rightarrow> bool" where |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
137 |
"isinP LfP ks = False" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
138 |
"isinP (NdP ps b lr) ks = |
70250 | 139 |
(let n = length ps in |
140 |
if ps = take n ks |
|
141 |
then case drop n ks of [] \<Rightarrow> b | k#ks' \<Rightarrow> isinP (sel2 k lr) ks' |
|
142 |
else False)" |
|
143 |
||
70268 | 144 |
definition emptyP :: trieP where |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
145 |
[simp]: "emptyP = LfP" |
70268 | 146 |
|
77830 | 147 |
fun lcp :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<times> 'a list \<times> 'a list" where |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
148 |
"lcp [] ys = ([],[],ys)" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
149 |
"lcp xs [] = ([],xs,[])" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
150 |
"lcp (x#xs) (y#ys) = |
70250 | 151 |
(if x\<noteq>y then ([],x#xs,y#ys) |
77830 | 152 |
else let (ps,xs',ys') = lcp xs ys in (x#ps,xs',ys'))" |
70250 | 153 |
|
154 |
||
155 |
lemma mod2_cong[fundef_cong]: |
|
156 |
"\<lbrakk> lr = lr'; k = k'; \<And>a b. lr'=(a,b) \<Longrightarrow> f (a) = f' (a) ; \<And>a b. lr'=(a,b) \<Longrightarrow> f (b) = f' (b) \<rbrakk> |
|
157 |
\<Longrightarrow> mod2 f k lr= mod2 f' k' lr'" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
158 |
by(cases lr, cases lr', auto) |
70250 | 159 |
|
70268 | 160 |
|
161 |
fun insertP :: "bool list \<Rightarrow> trieP \<Rightarrow> trieP" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
162 |
"insertP ks LfP = NdP ks True (LfP,LfP)" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
163 |
"insertP ks (NdP ps b lr) = |
77830 | 164 |
(case lcp ks ps of |
77767 | 165 |
(qs, k#ks', p#ps') \<Rightarrow> |
70250 | 166 |
let tp = NdP ps' b lr; tk = NdP ks' True (LfP,LfP) in |
167 |
NdP qs False (if k then (tp,tk) else (tk,tp)) | |
|
77767 | 168 |
(qs, k#ks', []) \<Rightarrow> |
70250 | 169 |
NdP ps b (mod2 (insertP ks') k lr) | |
77767 | 170 |
(qs, [], p#ps') \<Rightarrow> |
70250 | 171 |
let t = NdP ps' b lr in |
172 |
NdP qs True (if p then (LfP,t) else (t,LfP)) | |
|
173 |
(qs,[],[]) \<Rightarrow> NdP ps True lr)" |
|
174 |
||
175 |
||
77830 | 176 |
text \<open>Smart constructor that shrinks:\<close> |
177 |
definition nodeP :: "bool list \<Rightarrow> bool \<Rightarrow> trieP * trieP \<Rightarrow> trieP" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
178 |
"nodeP ps b lr = |
77830 | 179 |
(if b then NdP ps b lr |
180 |
else case lr of |
|
181 |
(LfP,LfP) \<Rightarrow> LfP | |
|
182 |
(LfP, NdP ks b lr) \<Rightarrow> NdP (ps @ True # ks) b lr | |
|
183 |
(NdP ks b lr, LfP) \<Rightarrow> NdP (ps @ False # ks) b lr | |
|
184 |
_ \<Rightarrow> NdP ps b lr)" |
|
70250 | 185 |
|
70268 | 186 |
fun deleteP :: "bool list \<Rightarrow> trieP \<Rightarrow> trieP" where |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
187 |
"deleteP ks LfP = LfP" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
188 |
"deleteP ks (NdP ps b lr) = |
77830 | 189 |
(case lcp ks ps of |
77767 | 190 |
(_, _, _#_) \<Rightarrow> NdP ps b lr | |
191 |
(_, k#ks', []) \<Rightarrow> nodeP ps b (mod2 (deleteP ks') k lr) | |
|
192 |
(_, [], []) \<Rightarrow> nodeP ps False lr)" |
|
70250 | 193 |
|
194 |
||
77830 | 195 |
|
70250 | 196 |
subsubsection \<open>Functional Correctness\<close> |
197 |
||
70268 | 198 |
text \<open>First step: @{typ trieP} implements @{typ trie} via the abstraction function \<open>abs_trieP\<close>:\<close> |
70250 | 199 |
|
200 |
fun prefix_trie :: "bool list \<Rightarrow> trie \<Rightarrow> trie" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
201 |
"prefix_trie [] t = t" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
202 |
"prefix_trie (k#ks) t = |
70250 | 203 |
(let t' = prefix_trie ks t in Nd False (if k then (Lf,t') else (t',Lf)))" |
204 |
||
70268 | 205 |
fun abs_trieP :: "trieP \<Rightarrow> trie" where |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
206 |
"abs_trieP LfP = Lf" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
207 |
"abs_trieP (NdP ps b (l,r)) = prefix_trie ps (Nd b (abs_trieP l, abs_trieP r))" |
70250 | 208 |
|
209 |
||
210 |
text \<open>Correctness of @{const isinP}:\<close> |
|
211 |
||
212 |
lemma isin_prefix_trie: |
|
213 |
"isin (prefix_trie ps t) ks |
|
214 |
= (ps = take (length ps) ks \<and> isin t (drop (length ps) ks))" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
215 |
by (induction ps arbitrary: ks) (auto split: list.split) |
70250 | 216 |
|
70269 | 217 |
lemma abs_trieP_isinP: |
70268 | 218 |
"isinP t ks = isin (abs_trieP t) ks" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
219 |
proof (induction t arbitrary: ks rule: abs_trieP.induct) |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
220 |
qed (auto simp: isin_prefix_trie split: list.split) |
70250 | 221 |
|
222 |
||
223 |
text \<open>Correctness of @{const insertP}:\<close> |
|
224 |
||
225 |
lemma prefix_trie_Lfs: "prefix_trie ks (Nd True (Lf,Lf)) = insert ks Lf" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
226 |
by (induction ks) auto |
70250 | 227 |
|
228 |
lemma insert_prefix_trie_same: |
|
229 |
"insert ps (prefix_trie ps (Nd b lr)) = prefix_trie ps (Nd True lr)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
230 |
by (induction ps) auto |
70250 | 231 |
|
232 |
lemma insert_append: "insert (ks @ ks') (prefix_trie ks t) = prefix_trie ks (insert ks' t)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
233 |
by (induction ks) auto |
70250 | 234 |
|
235 |
lemma prefix_trie_append: "prefix_trie (ps @ qs) t = prefix_trie ps (prefix_trie qs t)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
236 |
by (induction ps) auto |
70250 | 237 |
|
77830 | 238 |
lemma lcp_if: "lcp ks ps = (qs, ks', ps') \<Longrightarrow> |
70250 | 239 |
ks = qs @ ks' \<and> ps = qs @ ps' \<and> (ks' \<noteq> [] \<and> ps' \<noteq> [] \<longrightarrow> hd ks' \<noteq> hd ps')" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
240 |
proof (induction ks ps arbitrary: qs ks' ps' rule: lcp.induct) |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
241 |
qed (auto split: prod.splits if_splits) |
70250 | 242 |
|
70268 | 243 |
lemma abs_trieP_insertP: |
244 |
"abs_trieP (insertP ks t) = insert ks (abs_trieP t)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
245 |
proof (induction t arbitrary: ks) |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
246 |
qed (auto simp: prefix_trie_Lfs insert_prefix_trie_same insert_append prefix_trie_append |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
247 |
dest!: lcp_if split: list.split prod.split if_splits) |
70250 | 248 |
|
249 |
||
250 |
text \<open>Correctness of @{const deleteP}:\<close> |
|
251 |
||
252 |
lemma prefix_trie_Lf: "prefix_trie xs t = Lf \<longleftrightarrow> xs = [] \<and> t = Lf" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
253 |
by(cases xs)(auto) |
70250 | 254 |
|
70268 | 255 |
lemma abs_trieP_Lf: "abs_trieP t = Lf \<longleftrightarrow> t = LfP" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
256 |
by(cases t) (auto simp: prefix_trie_Lf) |
70250 | 257 |
|
258 |
lemma delete_prefix_trie: |
|
259 |
"delete xs (prefix_trie xs (Nd b (l,r))) |
|
260 |
= (if (l,r) = (Lf,Lf) then Lf else prefix_trie xs (Nd False (l,r)))" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
261 |
by(induction xs)(auto simp: prefix_trie_Lf) |
70250 | 262 |
|
263 |
lemma delete_append_prefix_trie: |
|
264 |
"delete (xs @ ys) (prefix_trie xs t) |
|
265 |
= (if delete ys t = Lf then Lf else prefix_trie xs (delete ys t))" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
266 |
by(induction xs)(auto simp: prefix_trie_Lf) |
70250 | 267 |
|
77830 | 268 |
lemma nodeP_LfP2: "nodeP xs False (LfP, LfP) = LfP" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
269 |
by(simp add: nodeP_def) |
77830 | 270 |
|
271 |
text \<open>Some non-inductive aux. lemmas:\<close> |
|
272 |
||
273 |
lemma abs_trieP_nodeP: "a\<noteq>LfP \<or> b \<noteq> LfP \<Longrightarrow> |
|
274 |
abs_trieP (nodeP xs f (a, b)) = prefix_trie xs (Nd f (abs_trieP a, abs_trieP b))" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
275 |
by(auto simp add: nodeP_def prefix_trie_append split: trieP.split) |
77830 | 276 |
|
277 |
lemma nodeP_True: "nodeP ps True lr = NdP ps True lr" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
278 |
by(simp add: nodeP_def) |
77830 | 279 |
|
70268 | 280 |
lemma delete_abs_trieP: |
281 |
"delete ks (abs_trieP t) = abs_trieP (deleteP ks t)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
282 |
proof (induction t arbitrary: ks) |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
283 |
qed (auto simp: delete_prefix_trie delete_append_prefix_trie |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
284 |
prefix_trie_append prefix_trie_Lf abs_trieP_Lf nodeP_LfP2 abs_trieP_nodeP nodeP_True |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
285 |
dest!: lcp_if split: if_splits list.split prod.split) |
70250 | 286 |
|
77830 | 287 |
text \<open>Invariant preservation:\<close> |
288 |
||
289 |
lemma insertP_LfP: "insertP xs t \<noteq> LfP" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
290 |
by(cases t)(auto split: prod.split list.split) |
77830 | 291 |
|
292 |
lemma invarP_insertP: "invarP t \<Longrightarrow> invarP(insertP xs t)" |
|
293 |
proof(induction t arbitrary: xs) |
|
294 |
case LfP thus ?case by simp |
|
295 |
next |
|
296 |
case (NdP bs b lr) |
|
297 |
then show ?case |
|
298 |
by(cases lr)(auto simp: insertP_LfP split: prod.split list.split) |
|
299 |
qed |
|
300 |
||
301 |
(* Inlining this proof leads to nontermination *) |
|
302 |
lemma invarP_nodeP: "\<lbrakk> invarP t1; invarP t2\<rbrakk> \<Longrightarrow> invarP (nodeP xs b (t1, t2))" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
303 |
by (auto simp add: nodeP_def split: trieP.split) |
77830 | 304 |
|
305 |
lemma invarP_deleteP: "invarP t \<Longrightarrow> invarP(deleteP xs t)" |
|
306 |
proof(induction t arbitrary: xs) |
|
307 |
case LfP thus ?case by simp |
|
308 |
next |
|
309 |
case (NdP ks b lr) |
|
310 |
thus ?case by(cases lr)(auto simp: invarP_nodeP split: prod.split list.split) |
|
311 |
qed |
|
312 |
||
70250 | 313 |
|
314 |
text \<open>The overall correctness proof. Simply composes correctness lemmas.\<close> |
|
315 |
||
70268 | 316 |
definition set_trieP :: "trieP \<Rightarrow> bool list set" where |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
317 |
"set_trieP = set_trie o abs_trieP" |
70250 | 318 |
|
70269 | 319 |
lemma isinP_set_trieP: "isinP t xs = (xs \<in> set_trieP t)" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
320 |
by(simp add: abs_trieP_isinP set_trie_isin set_trieP_def) |
70269 | 321 |
|
70268 | 322 |
lemma set_trieP_insertP: "set_trieP (insertP xs t) = set_trieP t \<union> {xs}" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
323 |
by(simp add: abs_trieP_insertP set_trie_insert set_trieP_def) |
70250 | 324 |
|
70269 | 325 |
lemma set_trieP_deleteP: "set_trieP (deleteP xs t) = set_trieP t - {xs}" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
326 |
by(auto simp: set_trie_delete set_trieP_def simp flip: delete_abs_trieP) |
70269 | 327 |
|
70250 | 328 |
interpretation SP: Set |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
329 |
where empty = emptyP and isin = isinP and insert = insertP and delete = deleteP |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
77830
diff
changeset
|
330 |
and set = set_trieP and invar = invarP |
70250 | 331 |
proof (standard, goal_cases) |
70268 | 332 |
case 1 show ?case by (simp add: set_trieP_def set_trie_def) |
70250 | 333 |
next |
70269 | 334 |
case 2 show ?case by(rule isinP_set_trieP) |
70250 | 335 |
next |
70268 | 336 |
case 3 thus ?case by (auto simp: set_trieP_insertP) |
70250 | 337 |
next |
70269 | 338 |
case 4 thus ?case by(auto simp: set_trieP_deleteP) |
77830 | 339 |
next |
340 |
case 5 thus ?case by(simp) |
|
341 |
next |
|
342 |
case 6 thus ?case by(rule invarP_insertP) |
|
343 |
next |
|
344 |
case 7 thus ?case by(rule invarP_deleteP) |
|
345 |
qed |
|
70250 | 346 |
|
347 |
end |