author | nipkow |
Mon, 13 Jan 2025 21:17:40 +0100 | |
changeset 81803 | 40f979c845b7 |
parent 81357 | 21a493abde0f |
permissions | -rw-r--r-- |
69232 | 1 |
(* Author: Tobias Nipkow, with contributions by Thomas Sewell *) |
2 |
||
69145 | 3 |
section "Arrays via Braun Trees" |
4 |
||
5 |
theory Array_Braun |
|
80036 | 6 |
imports |
7 |
Time_Funs |
|
8 |
Array_Specs |
|
9 |
Braun_Tree |
|
69145 | 10 |
begin |
11 |
||
12 |
subsection "Array" |
|
13 |
||
14 |
fun lookup1 :: "'a tree \<Rightarrow> nat \<Rightarrow> 'a" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
15 |
"lookup1 (Node l x r) n = (if n=1 then x else lookup1 (if even n then l else r) (n div 2))" |
69145 | 16 |
|
17 |
fun update1 :: "nat \<Rightarrow> 'a \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
18 |
"update1 n x Leaf = Node Leaf x Leaf" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
19 |
"update1 n x (Node l a r) = |
69145 | 20 |
(if n=1 then Node l x r else |
21 |
if even n then Node (update1 (n div 2) x l) a r |
|
22 |
else Node l a (update1 (n div 2) x r))" |
|
23 |
||
24 |
fun adds :: "'a list \<Rightarrow> nat \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
25 |
"adds [] n t = t" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
26 |
"adds (x#xs) n t = adds xs (n+1) (update1 (n+1) x t)" |
69145 | 27 |
|
28 |
fun list :: "'a tree \<Rightarrow> 'a list" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
29 |
"list Leaf = []" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
30 |
"list (Node l x r) = x # splice (list l) (list r)" |
69145 | 31 |
|
32 |
||
33 |
subsubsection "Functional Correctness" |
|
34 |
||
35 |
lemma size_list: "size(list t) = size t" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
36 |
by(induction t)(auto) |
69145 | 37 |
|
38 |
lemma minus1_div2: "(n - Suc 0) div 2 = (if odd n then n div 2 else n div 2 - 1)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
39 |
by auto arith |
69145 | 40 |
|
41 |
lemma nth_splice: "\<lbrakk> n < size xs + size ys; size ys \<le> size xs; size xs \<le> size ys + 1 \<rbrakk> |
|
42 |
\<Longrightarrow> splice xs ys ! n = (if even n then xs else ys) ! (n div 2)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
43 |
proof(induction xs ys arbitrary: n rule: splice.induct) |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
44 |
qed (auto simp: nth_Cons' minus1_div2) |
69145 | 45 |
|
46 |
lemma div2_in_bounds: |
|
47 |
"\<lbrakk> braun (Node l x r); n \<in> {1..size(Node l x r)}; n > 1 \<rbrakk> \<Longrightarrow> |
|
48 |
(odd n \<longrightarrow> n div 2 \<in> {1..size r}) \<and> (even n \<longrightarrow> n div 2 \<in> {1..size l})" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
49 |
by auto arith |
69145 | 50 |
|
51 |
declare upt_Suc[simp del] |
|
52 |
||
53 |
||
69597 | 54 |
paragraph \<open>\<^const>\<open>lookup1\<close>\<close> |
69145 | 55 |
|
56 |
lemma nth_list_lookup1: "\<lbrakk>braun t; i < size t\<rbrakk> \<Longrightarrow> list t ! i = lookup1 t (i+1)" |
|
57 |
proof(induction t arbitrary: i) |
|
58 |
case Leaf thus ?case by simp |
|
59 |
next |
|
60 |
case Node |
|
61 |
thus ?case using div2_in_bounds[OF Node.prems(1), of "i+1"] |
|
62 |
by (auto simp: nth_splice minus1_div2 size_list) |
|
63 |
qed |
|
64 |
||
65 |
lemma list_eq_map_lookup1: "braun t \<Longrightarrow> list t = map (lookup1 t) [1..<size t + 1]" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
66 |
by(auto simp add: list_eq_iff_nth_eq size_list nth_list_lookup1) |
69145 | 67 |
|
68 |
||
69597 | 69 |
paragraph \<open>\<^const>\<open>update1\<close>\<close> |
69145 | 70 |
|
71 |
lemma size_update1: "\<lbrakk> braun t; n \<in> {1.. size t} \<rbrakk> \<Longrightarrow> size(update1 n x t) = size t" |
|
72 |
proof(induction t arbitrary: n) |
|
73 |
case Leaf thus ?case by simp |
|
74 |
next |
|
75 |
case Node thus ?case using div2_in_bounds[OF Node.prems] by simp |
|
76 |
qed |
|
77 |
||
78 |
lemma braun_update1: "\<lbrakk>braun t; n \<in> {1.. size t} \<rbrakk> \<Longrightarrow> braun(update1 n x t)" |
|
79 |
proof(induction t arbitrary: n) |
|
80 |
case Leaf thus ?case by simp |
|
81 |
next |
|
82 |
case Node thus ?case |
|
83 |
using div2_in_bounds[OF Node.prems] by (simp add: size_update1) |
|
84 |
qed |
|
85 |
||
86 |
lemma lookup1_update1: "\<lbrakk> braun t; n \<in> {1.. size t} \<rbrakk> \<Longrightarrow> |
|
87 |
lookup1 (update1 n x t) m = (if n=m then x else lookup1 t m)" |
|
88 |
proof(induction t arbitrary: m n) |
|
89 |
case Leaf |
|
90 |
then show ?case by simp |
|
91 |
next |
|
92 |
have aux: "\<lbrakk> odd n; odd m \<rbrakk> \<Longrightarrow> n div 2 = (m::nat) div 2 \<longleftrightarrow> m=n" for m n |
|
93 |
using odd_two_times_div_two_succ by fastforce |
|
94 |
case Node |
|
95 |
thus ?case using div2_in_bounds[OF Node.prems] by (auto simp: aux) |
|
96 |
qed |
|
97 |
||
98 |
lemma list_update1: "\<lbrakk> braun t; n \<in> {1.. size t} \<rbrakk> \<Longrightarrow> list(update1 n x t) = (list t)[n-1 := x]" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
99 |
by(auto simp add: list_eq_map_lookup1 list_eq_iff_nth_eq lookup1_update1 size_update1 braun_update1) |
69145 | 100 |
|
101 |
text \<open>A second proof of @{thm list_update1}:\<close> |
|
102 |
||
103 |
lemma diff1_eq_iff: "n > 0 \<Longrightarrow> n - Suc 0 = m \<longleftrightarrow> n = m+1" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
104 |
by arith |
69145 | 105 |
|
106 |
lemma list_update_splice: |
|
107 |
"\<lbrakk> n < size xs + size ys; size ys \<le> size xs; size xs \<le> size ys + 1 \<rbrakk> \<Longrightarrow> |
|
108 |
(splice xs ys) [n := x] = |
|
109 |
(if even n then splice (xs[n div 2 := x]) ys else splice xs (ys[n div 2 := x]))" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
110 |
by(induction xs ys arbitrary: n rule: splice.induct) (auto split: nat.split) |
69145 | 111 |
|
112 |
lemma list_update2: "\<lbrakk> braun t; n \<in> {1.. size t} \<rbrakk> \<Longrightarrow> list(update1 n x t) = (list t)[n-1 := x]" |
|
113 |
proof(induction t arbitrary: n) |
|
114 |
case Leaf thus ?case by simp |
|
115 |
next |
|
116 |
case (Node l a r) thus ?case using div2_in_bounds[OF Node.prems] |
|
117 |
by(auto simp: list_update_splice diff1_eq_iff size_list split: nat.split) |
|
118 |
qed |
|
119 |
||
120 |
||
69597 | 121 |
paragraph \<open>\<^const>\<open>adds\<close>\<close> |
69145 | 122 |
|
123 |
lemma splice_last: shows |
|
124 |
"size ys \<le> size xs \<Longrightarrow> splice (xs @ [x]) ys = splice xs ys @ [x]" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
125 |
and "size ys+1 \<ge> size xs \<Longrightarrow> splice xs (ys @ [y]) = splice xs ys @ [y]" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
126 |
by(induction xs ys arbitrary: x y rule: splice.induct) (auto) |
69145 | 127 |
|
128 |
lemma list_add_hi: "braun t \<Longrightarrow> list(update1 (Suc(size t)) x t) = list t @ [x]" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
129 |
by(induction t)(auto simp: splice_last size_list) |
69145 | 130 |
|
131 |
lemma size_add_hi: "braun t \<Longrightarrow> m = size t \<Longrightarrow> size(update1 (Suc m) x t) = size t + 1" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
132 |
by(induction t arbitrary: m)(auto) |
69145 | 133 |
|
134 |
lemma braun_add_hi: "braun t \<Longrightarrow> braun(update1 (Suc(size t)) x t)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
135 |
by(induction t)(auto simp: size_add_hi) |
69145 | 136 |
|
137 |
lemma size_braun_adds: |
|
138 |
"\<lbrakk> braun t; size t = n \<rbrakk> \<Longrightarrow> size(adds xs n t) = size t + length xs \<and> braun (adds xs n t)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
139 |
by(induction xs arbitrary: t n)(auto simp: braun_add_hi size_add_hi) |
69145 | 140 |
|
141 |
lemma list_adds: "\<lbrakk> braun t; size t = n \<rbrakk> \<Longrightarrow> list(adds xs n t) = list t @ xs" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
142 |
by(induction xs arbitrary: t n)(auto simp: size_braun_adds list_add_hi size_add_hi braun_add_hi) |
69145 | 143 |
|
144 |
||
145 |
subsubsection "Array Implementation" |
|
146 |
||
147 |
interpretation A: Array |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
148 |
where lookup = "\<lambda>(t,l) n. lookup1 t (n+1)" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
149 |
and update = "\<lambda>n x (t,l). (update1 (n+1) x t, l)" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
150 |
and len = "\<lambda>(t,l). l" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
151 |
and array = "\<lambda>xs. (adds xs 0 Leaf, length xs)" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
152 |
and invar = "\<lambda>(t,l). braun t \<and> l = size t" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
153 |
and list = "\<lambda>(t,l). list t" |
69145 | 154 |
proof (standard, goal_cases) |
155 |
case 1 thus ?case by (simp add: nth_list_lookup1 split: prod.splits) |
|
156 |
next |
|
157 |
case 2 thus ?case by (simp add: list_update1 split: prod.splits) |
|
158 |
next |
|
159 |
case 3 thus ?case by (simp add: size_list split: prod.splits) |
|
160 |
next |
|
161 |
case 4 thus ?case by (simp add: list_adds) |
|
162 |
next |
|
163 |
case 5 thus ?case by (simp add: braun_update1 size_update1 split: prod.splits) |
|
164 |
next |
|
165 |
case 6 thus ?case by (simp add: size_braun_adds split: prod.splits) |
|
166 |
qed |
|
167 |
||
168 |
||
169 |
subsection "Flexible Array" |
|
170 |
||
171 |
fun add_lo where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
172 |
"add_lo x Leaf = Node Leaf x Leaf" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
173 |
"add_lo x (Node l a r) = Node (add_lo a r) x l" |
69145 | 174 |
|
175 |
fun merge where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
176 |
"merge Leaf r = r" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
177 |
"merge (Node l a r) rr = Node rr a (merge l r)" |
69145 | 178 |
|
179 |
fun del_lo where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
180 |
"del_lo Leaf = Leaf" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
181 |
"del_lo (Node l a r) = merge l r" |
69145 | 182 |
|
183 |
fun del_hi :: "nat \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
184 |
"del_hi n Leaf = Leaf" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
185 |
"del_hi n (Node l x r) = |
69752 | 186 |
(if n = 1 then Leaf |
187 |
else if even n |
|
188 |
then Node (del_hi (n div 2) l) x r |
|
189 |
else Node l x (del_hi (n div 2) r))" |
|
69145 | 190 |
|
191 |
||
192 |
subsubsection "Functional Correctness" |
|
193 |
||
194 |
||
69597 | 195 |
paragraph \<open>\<^const>\<open>add_lo\<close>\<close> |
69145 | 196 |
|
197 |
lemma list_add_lo: "braun t \<Longrightarrow> list (add_lo a t) = a # list t" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
198 |
by(induction t arbitrary: a) auto |
69145 | 199 |
|
200 |
lemma braun_add_lo: "braun t \<Longrightarrow> braun(add_lo x t)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
201 |
by(induction t arbitrary: x) (auto simp add: list_add_lo simp flip: size_list) |
69145 | 202 |
|
203 |
||
69597 | 204 |
paragraph \<open>\<^const>\<open>del_lo\<close>\<close> |
69145 | 205 |
|
206 |
lemma list_merge: "braun (Node l x r) \<Longrightarrow> list(merge l r) = splice (list l) (list r)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
207 |
by (induction l r rule: merge.induct) auto |
69145 | 208 |
|
209 |
lemma braun_merge: "braun (Node l x r) \<Longrightarrow> braun(merge l r)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
210 |
by (induction l r rule: merge.induct)(auto simp add: list_merge simp flip: size_list) |
69145 | 211 |
|
212 |
lemma list_del_lo: "braun t \<Longrightarrow> list(del_lo t) = tl (list t)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
213 |
by (cases t) (simp_all add: list_merge) |
69145 | 214 |
|
215 |
lemma braun_del_lo: "braun t \<Longrightarrow> braun(del_lo t)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
216 |
by (cases t) (simp_all add: braun_merge) |
69145 | 217 |
|
218 |
||
69597 | 219 |
paragraph \<open>\<^const>\<open>del_hi\<close>\<close> |
69145 | 220 |
|
221 |
lemma list_Nil_iff: "list t = [] \<longleftrightarrow> t = Leaf" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
222 |
by(cases t) simp_all |
69145 | 223 |
|
224 |
lemma butlast_splice: "butlast (splice xs ys) = |
|
225 |
(if size xs > size ys then splice (butlast xs) ys else splice xs (butlast ys))" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
226 |
by(induction xs ys rule: splice.induct) (auto) |
69145 | 227 |
|
228 |
lemma list_del_hi: "braun t \<Longrightarrow> size t = st \<Longrightarrow> list(del_hi st t) = butlast(list t)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
229 |
by (induction t arbitrary: st) (auto simp: list_Nil_iff size_list butlast_splice) |
69145 | 230 |
|
231 |
lemma braun_del_hi: "braun t \<Longrightarrow> size t = st \<Longrightarrow> braun(del_hi st t)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
232 |
by (induction t arbitrary: st) (auto simp: list_del_hi simp flip: size_list) |
69145 | 233 |
|
234 |
||
235 |
subsubsection "Flexible Array Implementation" |
|
236 |
||
237 |
interpretation AF: Array_Flex |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
238 |
where lookup = "\<lambda>(t,l) n. lookup1 t (n+1)" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
239 |
and update = "\<lambda>n x (t,l). (update1 (n+1) x t, l)" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
240 |
and len = "\<lambda>(t,l). l" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
241 |
and array = "\<lambda>xs. (adds xs 0 Leaf, length xs)" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
242 |
and invar = "\<lambda>(t,l). braun t \<and> l = size t" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
243 |
and list = "\<lambda>(t,l). list t" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
244 |
and add_lo = "\<lambda>x (t,l). (add_lo x t, l+1)" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
245 |
and del_lo = "\<lambda>(t,l). (del_lo t, l-1)" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
246 |
and add_hi = "\<lambda>x (t,l). (update1 (Suc l) x t, l+1)" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
247 |
and del_hi = "\<lambda>(t,l). (del_hi l t, l-1)" |
69145 | 248 |
proof (standard, goal_cases) |
249 |
case 1 thus ?case by (simp add: list_add_lo split: prod.splits) |
|
250 |
next |
|
251 |
case 2 thus ?case by (simp add: list_del_lo split: prod.splits) |
|
252 |
next |
|
253 |
case 3 thus ?case by (simp add: list_add_hi braun_add_hi split: prod.splits) |
|
254 |
next |
|
255 |
case 4 thus ?case by (simp add: list_del_hi split: prod.splits) |
|
256 |
next |
|
257 |
case 5 thus ?case by (simp add: braun_add_lo list_add_lo flip: size_list split: prod.splits) |
|
258 |
next |
|
259 |
case 6 thus ?case by (simp add: braun_del_lo list_del_lo flip: size_list split: prod.splits) |
|
260 |
next |
|
261 |
case 7 thus ?case by (simp add: size_add_hi braun_add_hi split: prod.splits) |
|
262 |
next |
|
263 |
case 8 thus ?case by (simp add: braun_del_hi list_del_hi flip: size_list split: prod.splits) |
|
264 |
qed |
|
265 |
||
266 |
||
267 |
subsection "Faster" |
|
268 |
||
69232 | 269 |
|
69941 | 270 |
subsubsection \<open>Size\<close> |
271 |
||
272 |
fun diff :: "'a tree \<Rightarrow> nat \<Rightarrow> nat" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
273 |
"diff Leaf _ = 0" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
274 |
"diff (Node l x r) n = (if n=0 then 1 else if even n then diff r (n div 2 - 1) else diff l (n div 2))" |
69941 | 275 |
|
276 |
fun size_fast :: "'a tree \<Rightarrow> nat" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
277 |
"size_fast Leaf = 0" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
278 |
"size_fast (Node l x r) = (let n = size_fast r in 1 + 2*n + diff l n)" |
69941 | 279 |
|
71846 | 280 |
declare Let_def[simp] |
281 |
||
69943 | 282 |
lemma diff: "braun t \<Longrightarrow> size t : {n, n + 1} \<Longrightarrow> diff t n = size t - n" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
283 |
by (induction t arbitrary: n) auto |
69941 | 284 |
|
285 |
lemma size_fast: "braun t \<Longrightarrow> size_fast t = size t" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
286 |
by (induction t) (auto simp add: diff) |
69941 | 287 |
|
288 |
||
69232 | 289 |
subsubsection \<open>Initialization with 1 element\<close> |
290 |
||
69206 | 291 |
fun braun_of_naive :: "'a \<Rightarrow> nat \<Rightarrow> 'a tree" where |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
292 |
"braun_of_naive x n = (if n=0 then Leaf |
69145 | 293 |
else let m = (n-1) div 2 |
69206 | 294 |
in if odd n then Node (braun_of_naive x m) x (braun_of_naive x m) |
295 |
else Node (braun_of_naive x (m + 1)) x (braun_of_naive x m))" |
|
69145 | 296 |
|
297 |
fun braun2_of :: "'a \<Rightarrow> nat \<Rightarrow> 'a tree * 'a tree" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
298 |
"braun2_of x n = (if n = 0 then (Leaf, Node Leaf x Leaf) |
69145 | 299 |
else let (s,t) = braun2_of x ((n-1) div 2) |
300 |
in if odd n then (Node s x s, Node t x s) else (Node t x s, Node t x t))" |
|
301 |
||
302 |
definition braun_of :: "'a \<Rightarrow> nat \<Rightarrow> 'a tree" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
303 |
"braun_of x n = fst (braun2_of x n)" |
69145 | 304 |
|
305 |
declare braun2_of.simps [simp del] |
|
306 |
||
307 |
lemma braun2_of_size_braun: "braun2_of x n = (s,t) \<Longrightarrow> size s = n \<and> size t = n+1 \<and> braun s \<and> braun t" |
|
308 |
proof(induction x n arbitrary: s t rule: braun2_of.induct) |
|
309 |
case (1 x n) |
|
310 |
then show ?case |
|
311 |
by (auto simp: braun2_of.simps[of x n] split: prod.splits if_splits) presburger+ |
|
312 |
qed |
|
313 |
||
314 |
lemma braun2_of_replicate: |
|
315 |
"braun2_of x n = (s,t) \<Longrightarrow> list s = replicate n x \<and> list t = replicate (n+1) x" |
|
316 |
proof(induction x n arbitrary: s t rule: braun2_of.induct) |
|
317 |
case (1 x n) |
|
318 |
have "x # replicate m x = replicate (m+1) x" for m by simp |
|
319 |
with 1 show ?case |
|
320 |
apply (auto simp: braun2_of.simps[of x n] replicate.simps(2)[of 0 x] |
|
321 |
simp del: replicate.simps(2) split: prod.splits if_splits) |
|
322 |
by presburger+ |
|
323 |
qed |
|
324 |
||
325 |
corollary braun_braun_of: "braun(braun_of x n)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
326 |
unfolding braun_of_def by (metis eq_fst_iff braun2_of_size_braun) |
69145 | 327 |
|
328 |
corollary list_braun_of: "list(braun_of x n) = replicate n x" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
329 |
unfolding braun_of_def by (metis eq_fst_iff braun2_of_replicate) |
69145 | 330 |
|
69232 | 331 |
|
332 |
subsubsection "Proof Infrastructure" |
|
333 |
||
334 |
text \<open>Originally due to Thomas Sewell.\<close> |
|
335 |
||
336 |
paragraph \<open>\<open>take_nths\<close>\<close> |
|
337 |
||
338 |
fun take_nths :: "nat \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
339 |
"take_nths i k [] = []" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
340 |
"take_nths i k (x # xs) = (if i = 0 then x # take_nths (2^k - 1) k xs |
69232 | 341 |
else take_nths (i - 1) k xs)" |
342 |
||
71399 | 343 |
text \<open>This is the more concise definition but seems to complicate the proofs:\<close> |
344 |
||
345 |
lemma take_nths_eq_nths: "take_nths i k xs = nths xs (\<Union>n. {n*2^k + i})" |
|
346 |
proof(induction xs arbitrary: i) |
|
347 |
case Nil |
|
348 |
then show ?case by simp |
|
349 |
next |
|
350 |
case (Cons x xs) |
|
351 |
show ?case |
|
352 |
proof cases |
|
353 |
assume [simp]: "i = 0" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
354 |
have "\<And>x n. Suc x = n * 2 ^ k \<Longrightarrow> \<exists>xa. x = Suc xa * 2 ^ k - Suc 0" |
71399 | 355 |
by (metis diff_Suc_Suc diff_zero mult_eq_0_iff not0_implies_Suc) |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
356 |
then have "(\<Union>n. {(n+1) * 2 ^ k - 1}) = {m. \<exists>n. Suc m = n * 2 ^ k}" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
357 |
by (auto simp del: mult_Suc) |
71399 | 358 |
thus ?thesis by (simp add: Cons.IH ac_simps nths_Cons) |
359 |
next |
|
360 |
assume [arith]: "i \<noteq> 0" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
361 |
have "\<And>x n. Suc x = n * 2 ^ k + i \<Longrightarrow> \<exists>xa. x = xa * 2 ^ k + i - Suc 0" |
71399 | 362 |
by (metis diff_Suc_Suc diff_zero) |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
363 |
then have "(\<Union>n. {n * 2 ^ k + i - 1}) = {m. \<exists>n. Suc m = n * 2 ^ k + i}" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
364 |
by auto |
71399 | 365 |
thus ?thesis by (simp add: Cons.IH nths_Cons) |
366 |
qed |
|
367 |
qed |
|
368 |
||
69232 | 369 |
lemma take_nths_drop: |
370 |
"take_nths i k (drop j xs) = take_nths (i + j) k xs" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
371 |
by (induct xs arbitrary: i j; simp add: drop_Cons split: nat.split) |
69232 | 372 |
|
373 |
lemma take_nths_00: |
|
374 |
"take_nths 0 0 xs = xs" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
375 |
by (induct xs; simp) |
69232 | 376 |
|
377 |
lemma splice_take_nths: |
|
378 |
"splice (take_nths 0 (Suc 0) xs) (take_nths (Suc 0) (Suc 0) xs) = xs" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
379 |
by (induct xs; simp) |
69232 | 380 |
|
381 |
lemma take_nths_take_nths: |
|
382 |
"take_nths i m (take_nths j n xs) = take_nths ((i * 2^n) + j) (m + n) xs" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
383 |
by (induct xs arbitrary: i j; simp add: algebra_simps power_add) |
69232 | 384 |
|
385 |
lemma take_nths_empty: |
|
386 |
"(take_nths i k xs = []) = (length xs \<le> i)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
387 |
by (induction xs arbitrary: i k) auto |
69232 | 388 |
|
389 |
lemma hd_take_nths: |
|
390 |
"i < length xs \<Longrightarrow> hd(take_nths i k xs) = xs ! i" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
391 |
by (induction xs arbitrary: i k) auto |
69232 | 392 |
|
393 |
lemma take_nths_01_splice: |
|
394 |
"\<lbrakk> length xs = length ys \<or> length xs = length ys + 1 \<rbrakk> \<Longrightarrow> |
|
395 |
take_nths 0 (Suc 0) (splice xs ys) = xs \<and> |
|
396 |
take_nths (Suc 0) (Suc 0) (splice xs ys) = ys" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
397 |
by (induct xs arbitrary: ys; case_tac ys; simp) |
69232 | 398 |
|
399 |
lemma length_take_nths_00: |
|
400 |
"length (take_nths 0 (Suc 0) xs) = length (take_nths (Suc 0) (Suc 0) xs) \<or> |
|
401 |
length (take_nths 0 (Suc 0) xs) = length (take_nths (Suc 0) (Suc 0) xs) + 1" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
402 |
by (induct xs) auto |
69232 | 403 |
|
404 |
||
405 |
paragraph \<open>\<open>braun_list\<close>\<close> |
|
406 |
||
407 |
fun braun_list :: "'a tree \<Rightarrow> 'a list \<Rightarrow> bool" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
408 |
"braun_list Leaf xs = (xs = [])" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
409 |
"braun_list (Node l x r) xs = (xs \<noteq> [] \<and> x = hd xs \<and> |
69232 | 410 |
braun_list l (take_nths 1 1 xs) \<and> |
411 |
braun_list r (take_nths 2 1 xs))" |
|
412 |
||
413 |
lemma braun_list_eq: |
|
414 |
"braun_list t xs = (braun t \<and> xs = list t)" |
|
415 |
proof (induct t arbitrary: xs) |
|
416 |
case Leaf |
|
417 |
show ?case by simp |
|
418 |
next |
|
419 |
case Node |
|
420 |
show ?case |
|
421 |
using length_take_nths_00[of xs] splice_take_nths[of xs] |
|
422 |
by (auto simp: neq_Nil_conv Node.hyps size_list[symmetric] take_nths_01_splice) |
|
423 |
qed |
|
424 |
||
425 |
||
426 |
subsubsection \<open>Converting a list of elements into a Braun tree\<close> |
|
427 |
||
428 |
fun nodes :: "'a tree list \<Rightarrow> 'a list \<Rightarrow> 'a tree list \<Rightarrow> 'a tree list" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
429 |
"nodes (l#ls) (x#xs) (r#rs) = Node l x r # nodes ls xs rs" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
430 |
"nodes (l#ls) (x#xs) [] = Node l x Leaf # nodes ls xs []" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
431 |
"nodes [] (x#xs) (r#rs) = Node Leaf x r # nodes [] xs rs" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
432 |
"nodes [] (x#xs) [] = Node Leaf x Leaf # nodes [] xs []" | |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
433 |
"nodes ls [] rs = []" |
69232 | 434 |
|
435 |
fun brauns :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a tree list" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
436 |
"brauns k xs = (if xs = [] then [] else |
69232 | 437 |
let ys = take (2^k) xs; |
438 |
zs = drop (2^k) xs; |
|
439 |
ts = brauns (k+1) zs |
|
440 |
in nodes ts ys (drop (2^k) ts))" |
|
441 |
||
442 |
declare brauns.simps[simp del] |
|
443 |
||
444 |
definition brauns1 :: "'a list \<Rightarrow> 'a tree" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
445 |
"brauns1 xs = (if xs = [] then Leaf else brauns 0 xs ! 0)" |
69232 | 446 |
|
447 |
||
448 |
paragraph "Functional correctness" |
|
449 |
||
450 |
text \<open>The proof is originally due to Thomas Sewell.\<close> |
|
451 |
||
452 |
lemma length_nodes: |
|
453 |
"length (nodes ls xs rs) = length xs" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
454 |
by (induct ls xs rs rule: nodes.induct; simp) |
69232 | 455 |
|
456 |
lemma nth_nodes: |
|
457 |
"i < length xs \<Longrightarrow> nodes ls xs rs ! i = |
|
458 |
Node (if i < length ls then ls ! i else Leaf) (xs ! i) |
|
459 |
(if i < length rs then rs ! i else Leaf)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
460 |
by (induct ls xs rs arbitrary: i rule: nodes.induct; |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
461 |
simp add: nth_Cons split: nat.split) |
69232 | 462 |
|
463 |
theorem length_brauns: |
|
464 |
"length (brauns k xs) = min (length xs) (2 ^ k)" |
|
465 |
proof (induct xs arbitrary: k rule: measure_induct_rule[where f=length]) |
|
71846 | 466 |
case (less xs) thus ?case by (simp add: brauns.simps[of k xs] length_nodes) |
69232 | 467 |
qed |
468 |
||
469 |
theorem brauns_correct: |
|
470 |
"i < min (length xs) (2 ^ k) \<Longrightarrow> braun_list (brauns k xs ! i) (take_nths i k xs)" |
|
471 |
proof (induct xs arbitrary: i k rule: measure_induct_rule[where f=length]) |
|
472 |
case (less xs) |
|
473 |
have "xs \<noteq> []" using less.prems by auto |
|
474 |
let ?zs = "drop (2^k) xs" |
|
475 |
let ?ts = "brauns (Suc k) ?zs" |
|
476 |
from less.hyps[of ?zs _ "Suc k"] |
|
477 |
have IH: "\<lbrakk> j = i + 2 ^ k; i < min (length ?zs) (2 ^ (k+1)) \<rbrakk> \<Longrightarrow> |
|
478 |
braun_list (?ts ! i) (take_nths j (Suc k) xs)" for i j |
|
479 |
using \<open>xs \<noteq> []\<close> by (simp add: take_nths_drop) |
|
480 |
show ?case |
|
71304 | 481 |
using less.prems |
71846 | 482 |
by (auto simp: brauns.simps[of k xs] nth_nodes take_nths_take_nths |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
483 |
IH take_nths_empty hd_take_nths length_brauns) |
69232 | 484 |
qed |
485 |
||
486 |
corollary brauns1_correct: |
|
487 |
"braun (brauns1 xs) \<and> list (brauns1 xs) = xs" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
488 |
using brauns_correct[of 0 xs 0] |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
489 |
by (simp add: brauns1_def braun_list_eq take_nths_00) |
69232 | 490 |
|
491 |
||
492 |
paragraph "Running Time Analysis" |
|
493 |
||
80036 | 494 |
time_fun_0 "(^)" |
495 |
||
496 |
time_fun nodes |
|
497 |
||
498 |
lemma T_nodes: "T_nodes ls xs rs = length xs + 1" |
|
499 |
by(induction ls xs rs rule: T_nodes.induct) auto |
|
500 |
||
501 |
time_fun brauns |
|
502 |
||
503 |
lemma T_brauns_simple: "T_brauns k xs = (if xs = [] then 0 else |
|
504 |
3 * (min (2^k) (length xs) + 1) + (min (2^k) (length xs - 2^k) + 1) + T_brauns (k+1) (drop (2^k) xs)) + 1" |
|
81357
21a493abde0f
uniform name T_f for closed-form lemmas for function T_f
nipkow
parents:
80036
diff
changeset
|
505 |
by(simp add: T_nodes T_take T_drop length_brauns min_def) |
80036 | 506 |
|
507 |
theorem T_brauns_ub: |
|
508 |
"T_brauns k xs \<le> 9 * (length xs + 1)" |
|
69232 | 509 |
proof (induction xs arbitrary: k rule: measure_induct_rule[where f = length]) |
510 |
case (less xs) |
|
511 |
show ?case |
|
512 |
proof cases |
|
513 |
assume "xs = []" |
|
71846 | 514 |
thus ?thesis by(simp) |
69232 | 515 |
next |
516 |
assume "xs \<noteq> []" |
|
80036 | 517 |
let ?n = "length xs" let ?zs = "drop (2^k) xs" |
518 |
have *: "?n - 2^k + 1 \<le> ?n" |
|
519 |
using \<open>xs \<noteq> []\<close> less_eq_Suc_le by fastforce |
|
520 |
have "T_brauns k xs = |
|
521 |
3 * (min (2^k) ?n + 1) + (min (2^k) (?n - 2^k) + 1) + T_brauns (k+1) ?zs + 1" |
|
522 |
unfolding T_brauns_simple[of k xs] using \<open>xs \<noteq> []\<close> by(simp del: T_brauns.simps) |
|
523 |
also have "\<dots> \<le> 4 * min (2^k) ?n + T_brauns (k+1) ?zs + 5" |
|
524 |
by(simp add: min_def) |
|
525 |
also have "\<dots> \<le> 4 * min (2^k) ?n + 9 * (length ?zs + 1) + 5" |
|
69232 | 526 |
using less[of ?zs "k+1"] \<open>xs \<noteq> []\<close> |
80036 | 527 |
by (simp del: T_brauns.simps) |
528 |
also have "\<dots> = 4 * min (2^k) ?n + 9 * (?n - 2^k + 1) + 5" |
|
529 |
by(simp) |
|
530 |
also have "\<dots> = 4 * min (2^k) ?n + 4 * (?n - 2^k) + 5 * (?n - 2^k + 1) + 9" |
|
69232 | 531 |
by(simp) |
80036 | 532 |
also have "\<dots> = 4 * ?n + 5 * (?n - 2^k + 1) + 9" |
533 |
by(simp) |
|
534 |
also have "\<dots> \<le> 4 * ?n + 5 * ?n + 9" |
|
535 |
using * by(simp) |
|
536 |
also have "\<dots> = 9 * (?n + 1)" |
|
537 |
by (simp add: Suc_leI) |
|
538 |
finally show ?thesis . |
|
69232 | 539 |
qed |
540 |
qed |
|
541 |
||
542 |
||
543 |
subsubsection \<open>Converting a Braun Tree into a List of Elements\<close> |
|
544 |
||
545 |
text \<open>The code and the proof are originally due to Thomas Sewell (except running time).\<close> |
|
546 |
||
547 |
function list_fast_rec :: "'a tree list \<Rightarrow> 'a list" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
548 |
"list_fast_rec ts = (let us = filter (\<lambda>t. t \<noteq> Leaf) ts in |
69984
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
549 |
if us = [] then [] else |
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
550 |
map value us @ list_fast_rec (map left us @ map right us))" |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
551 |
by (pat_completeness, auto) |
69232 | 552 |
|
69984
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
553 |
lemma list_fast_rec_term1: "ts \<noteq> [] \<Longrightarrow> Leaf \<notin> set ts \<Longrightarrow> |
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
554 |
sum_list (map (size o left) ts) + sum_list (map (size o right) ts) < sum_list (map size ts)" |
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
555 |
apply (clarsimp simp: sum_list_addf[symmetric] sum_list_map_filter') |
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
556 |
apply (rule sum_list_strict_mono; clarsimp?) |
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
557 |
apply (case_tac x; simp) |
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
558 |
done |
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
559 |
|
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
560 |
lemma list_fast_rec_term: "us \<noteq> [] \<Longrightarrow> us = filter (\<lambda>t. t \<noteq> \<langle>\<rangle>) ts \<Longrightarrow> |
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
561 |
sum_list (map (size o left) us) + sum_list (map (size o right) us) < sum_list (map size ts)" |
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
562 |
apply (rule order_less_le_trans, rule list_fast_rec_term1, simp_all) |
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
563 |
apply (rule sum_list_filter_le_nat) |
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
564 |
done |
69232 | 565 |
|
566 |
termination |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
567 |
by (relation "measure (sum_list o map size)"; simp add: list_fast_rec_term) |
69232 | 568 |
|
569 |
declare list_fast_rec.simps[simp del] |
|
570 |
||
571 |
definition list_fast :: "'a tree \<Rightarrow> 'a list" where |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
572 |
"list_fast t = list_fast_rec [t]" |
69232 | 573 |
|
80036 | 574 |
(* TODO: map and filter are a problem! |
575 |
- The automatically generated T_map is slightly more complicated than needed. |
|
576 |
- We cannot use the manually defined T_map directly because the automatic translation |
|
577 |
assumes that T_map has a more complicated type and generates a "wrong" call. |
|
578 |
Therefore we hide map/filter at the moment. |
|
579 |
*) |
|
580 |
||
581 |
definition "filter_not_Leaf = filter (\<lambda>t. t \<noteq> Leaf)" |
|
582 |
||
583 |
definition "map_left = map left" |
|
584 |
definition "map_right = map right" |
|
585 |
definition "map_value = map value" |
|
586 |
||
587 |
definition "T_filter_not_Leaf ts = length ts + 1" |
|
588 |
definition "T_map_left ts = length ts + 1" |
|
589 |
definition "T_map_right ts = length ts + 1" |
|
590 |
definition "T_map_value ts = length ts + 1" |
|
591 |
(* |
|
592 |
time_fun "tree.value" |
|
593 |
time_fun "left" |
|
594 |
time_fun "right" |
|
595 |
*) |
|
596 |
||
597 |
lemmas defs = filter_not_Leaf_def map_left_def map_right_def map_value_def |
|
598 |
T_filter_not_Leaf_def T_map_value_def T_map_left_def T_map_right_def |
|
599 |
||
600 |
(* A variant w/o explicit map/filter; T_list_fast_rec is generated from it *) |
|
601 |
lemma list_fast_rec_simp: |
|
602 |
"list_fast_rec ts = (let us = filter_not_Leaf ts in |
|
603 |
if us = [] then [] else |
|
604 |
map_value us @ list_fast_rec (map_left us @ map_right us))" |
|
605 |
unfolding defs list_fast_rec.simps[of ts] by(rule refl) |
|
606 |
||
607 |
time_function list_fast_rec equations list_fast_rec_simp |
|
608 |
termination |
|
609 |
by (relation "measure (sum_list o map size)"; simp add: list_fast_rec_term defs) |
|
610 |
||
72550
1d0ae7e578a0
renamed t_ functions to T_ (im Auftrag von T. Nipkow)
blanchet
parents:
71846
diff
changeset
|
611 |
declare T_list_fast_rec.simps[simp del] |
69232 | 612 |
|
80036 | 613 |
|
69232 | 614 |
paragraph "Functional Correctness" |
615 |
||
616 |
lemma list_fast_rec_all_Leaf: |
|
617 |
"\<forall>t \<in> set ts. t = Leaf \<Longrightarrow> list_fast_rec ts = []" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
618 |
by (simp add: filter_empty_conv list_fast_rec.simps) |
69232 | 619 |
|
620 |
lemma take_nths_eq_single: |
|
621 |
"length xs - i < 2^n \<Longrightarrow> take_nths i n xs = take 1 (drop i xs)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
622 |
by (induction xs arbitrary: i n; simp add: drop_Cons') |
69232 | 623 |
|
624 |
lemma braun_list_Nil: |
|
625 |
"braun_list t [] = (t = Leaf)" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
626 |
by (cases t; simp) |
69232 | 627 |
|
628 |
lemma braun_list_not_Nil: "xs \<noteq> [] \<Longrightarrow> |
|
629 |
braun_list t xs = |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
630 |
(\<exists>l x r. t = Node l x r \<and> x = hd xs \<and> |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
631 |
braun_list l (take_nths 1 1 xs) \<and> |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
632 |
braun_list r (take_nths 2 1 xs))" |
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
633 |
by(cases t; simp) |
69232 | 634 |
|
635 |
theorem list_fast_rec_correct: |
|
636 |
"\<lbrakk> length ts = 2 ^ k; \<forall>i < 2 ^ k. braun_list (ts ! i) (take_nths i k xs) \<rbrakk> |
|
637 |
\<Longrightarrow> list_fast_rec ts = xs" |
|
638 |
proof (induct xs arbitrary: k ts rule: measure_induct_rule[where f=length]) |
|
639 |
case (less xs) |
|
640 |
show ?case |
|
641 |
proof (cases "length xs < 2 ^ k") |
|
642 |
case True |
|
643 |
from less.prems True have filter: |
|
644 |
"\<exists>n. ts = map (\<lambda>x. Node Leaf x Leaf) xs @ replicate n Leaf" |
|
645 |
apply (rule_tac x="length ts - length xs" in exI) |
|
646 |
apply (clarsimp simp: list_eq_iff_nth_eq) |
|
647 |
apply(auto simp: nth_append braun_list_not_Nil take_nths_eq_single braun_list_Nil hd_drop_conv_nth) |
|
648 |
done |
|
649 |
thus ?thesis |
|
71846 | 650 |
by (clarsimp simp: list_fast_rec.simps[of ts] o_def list_fast_rec_all_Leaf) |
69232 | 651 |
next |
652 |
case False |
|
653 |
with less.prems(2) have *: |
|
654 |
"\<forall>i < 2 ^ k. ts ! i \<noteq> Leaf |
|
69655 | 655 |
\<and> value (ts ! i) = xs ! i |
69232 | 656 |
\<and> braun_list (left (ts ! i)) (take_nths (i + 2 ^ k) (Suc k) xs) |
657 |
\<and> (\<forall>ys. ys = take_nths (i + 2 * 2 ^ k) (Suc k) xs |
|
658 |
\<longrightarrow> braun_list (right (ts ! i)) ys)" |
|
659 |
by (auto simp: take_nths_empty hd_take_nths braun_list_not_Nil take_nths_take_nths |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
660 |
algebra_simps) |
69655 | 661 |
have 1: "map value ts = take (2 ^ k) xs" |
69232 | 662 |
using less.prems(1) False by (simp add: list_eq_iff_nth_eq *) |
663 |
have 2: "list_fast_rec (map left ts @ map right ts) = drop (2 ^ k) xs" |
|
664 |
using less.prems(1) False |
|
665 |
by (auto intro!: Nat.diff_less less.hyps[where k= "Suc k"] |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
666 |
simp: nth_append * take_nths_drop algebra_simps) |
69232 | 667 |
from less.prems(1) False show ?thesis |
71846 | 668 |
by (auto simp: list_fast_rec.simps[of ts] 1 2 * all_set_conv_all_nth) |
69232 | 669 |
qed |
670 |
qed |
|
671 |
||
672 |
corollary list_fast_correct: |
|
673 |
"braun t \<Longrightarrow> list_fast t = list t" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
674 |
by (simp add: list_fast_def take_nths_00 braun_list_eq list_fast_rec_correct[where k=0]) |
69232 | 675 |
|
80036 | 676 |
|
69232 | 677 |
paragraph "Running Time Analysis" |
678 |
||
679 |
lemma sum_tree_list_children: "\<forall>t \<in> set ts. t \<noteq> Leaf \<Longrightarrow> |
|
680 |
(\<Sum>t\<leftarrow>ts. k * size t) = (\<Sum>t \<leftarrow> map left ts @ map right ts. k * size t) + k * length ts" |
|
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
681 |
by(induction ts)(auto simp add: neq_Leaf_iff algebra_simps) |
69232 | 682 |
|
72550
1d0ae7e578a0
renamed t_ functions to T_ (im Auftrag von T. Nipkow)
blanchet
parents:
71846
diff
changeset
|
683 |
theorem T_list_fast_rec_ub: |
80036 | 684 |
"T_list_fast_rec ts \<le> sum_list (map (\<lambda>t. 14*size t + 1) ts) + 2" |
69984
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
685 |
proof (induction ts rule: measure_induct_rule[where f="sum_list o map size"]) |
69232 | 686 |
case (less ts) |
69984
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
687 |
let ?us = "filter (\<lambda>t. t \<noteq> Leaf) ts" |
69232 | 688 |
show ?case |
689 |
proof cases |
|
69984
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
690 |
assume "?us = []" |
72550
1d0ae7e578a0
renamed t_ functions to T_ (im Auftrag von T. Nipkow)
blanchet
parents:
71846
diff
changeset
|
691 |
thus ?thesis using T_list_fast_rec.simps[of ts] |
80036 | 692 |
by(simp add: defs sum_list_Suc) |
78653
7ed1759fe1bd
tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents:
72550
diff
changeset
|
693 |
next |
69984
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
694 |
assume "?us \<noteq> []" |
69232 | 695 |
let ?children = "map left ?us @ map right ?us" |
80036 | 696 |
have 1: "1 \<le> length ?us" |
697 |
using \<open>?us \<noteq> []\<close> linorder_not_less by auto |
|
698 |
have "T_list_fast_rec ts = T_list_fast_rec ?children + 5 * length ?us + length ts + 7" |
|
699 |
using \<open>?us \<noteq> []\<close> T_list_fast_rec.simps[of ts] by(simp add: defs T_append) |
|
700 |
also have "\<dots> \<le> (\<Sum>t\<leftarrow>?children. 14 * size t + 1) + 5 * length ?us + length ts + 9" |
|
69984
3afa3b25b5e7
Tweak Braun tree list_fast_rec recursion.
Thomas Sewell <sewell@chalmers.se>
parents:
69943
diff
changeset
|
701 |
using less[of "?children"] list_fast_rec_term[of "?us"] \<open>?us \<noteq> []\<close> |
69232 | 702 |
by (simp) |
80036 | 703 |
also have "\<dots> = (\<Sum>t\<leftarrow>?children. 14 * size t) + 7 * length ?us + length ts + 9" |
69232 | 704 |
by(simp add: sum_list_Suc o_def) |
80036 | 705 |
also have "\<dots> \<le> (\<Sum>t\<leftarrow>?children. 14 * size t) + 14 * length ?us + length ts + 2" |
706 |
using 1 by(simp add: sum_list_Suc o_def) |
|
707 |
also have "\<dots> = (\<Sum>t\<leftarrow>?us. 14 * size t) + length ts + 2" |
|
69232 | 708 |
by(simp add: sum_tree_list_children) |
80036 | 709 |
also have "\<dots> \<le> (\<Sum>t\<leftarrow>ts. 14 * size t) + length ts + 2" |
69232 | 710 |
by(simp add: sum_list_filter_le_nat) |
80036 | 711 |
also have "\<dots> = (\<Sum>t\<leftarrow>ts. 14 * size t + 1) + 2" |
69232 | 712 |
by(simp add: sum_list_Suc) |
713 |
finally show ?case . |
|
714 |
qed |
|
715 |
qed |
|
716 |
||
72550
1d0ae7e578a0
renamed t_ functions to T_ (im Auftrag von T. Nipkow)
blanchet
parents:
71846
diff
changeset
|
717 |
end |