41561
|
1 |
(* Title: HOL/SPARK/Examples/Liseq/Longest_Increasing_Subsequence.thy
|
|
2 |
Author: Stefan Berghofer
|
|
3 |
Copyright: secunet Security Networks AG
|
|
4 |
*)
|
|
5 |
|
|
6 |
theory Longest_Increasing_Subsequence
|
|
7 |
imports SPARK
|
|
8 |
begin
|
|
9 |
|
|
10 |
text {*
|
|
11 |
Set of all increasing subsequences in a prefix of an array
|
|
12 |
*}
|
|
13 |
|
|
14 |
definition iseq :: "(nat \<Rightarrow> 'a::linorder) \<Rightarrow> nat \<Rightarrow> nat set set" where
|
|
15 |
"iseq xs l = {is. (\<forall>i\<in>is. i < l) \<and>
|
|
16 |
(\<forall>i\<in>is. \<forall>j\<in>is. i \<le> j \<longrightarrow> xs i \<le> xs j)}"
|
|
17 |
|
|
18 |
text {*
|
|
19 |
Length of longest increasing subsequence in a prefix of an array
|
|
20 |
*}
|
|
21 |
|
|
22 |
definition liseq :: "(nat \<Rightarrow> 'a::linorder) \<Rightarrow> nat \<Rightarrow> nat" where
|
|
23 |
"liseq xs i = Max (card ` iseq xs i)"
|
|
24 |
|
|
25 |
text {*
|
|
26 |
Length of longest increasing subsequence ending at a particular position
|
|
27 |
*}
|
|
28 |
|
|
29 |
definition liseq' :: "(nat \<Rightarrow> 'a::linorder) \<Rightarrow> nat \<Rightarrow> nat" where
|
|
30 |
"liseq' xs i = Max (card ` (iseq xs (Suc i) \<inter> {is. Max is = i}))"
|
|
31 |
|
|
32 |
lemma iseq_finite: "finite (iseq xs i)"
|
|
33 |
apply (simp add: iseq_def)
|
|
34 |
apply (rule finite_subset [OF _
|
|
35 |
finite_Collect_subsets [of "{j. j < i}"]])
|
|
36 |
apply auto
|
|
37 |
done
|
|
38 |
|
|
39 |
lemma iseq_finite': "is \<in> iseq xs i \<Longrightarrow> finite is"
|
|
40 |
by (auto simp add: iseq_def bounded_nat_set_is_finite)
|
|
41 |
|
|
42 |
lemma iseq_singleton: "i < l \<Longrightarrow> {i} \<in> iseq xs l"
|
|
43 |
by (simp add: iseq_def)
|
|
44 |
|
|
45 |
lemma iseq_trivial: "{} \<in> iseq xs i"
|
|
46 |
by (simp add: iseq_def)
|
|
47 |
|
|
48 |
lemma iseq_nonempty: "iseq xs i \<noteq> {}"
|
|
49 |
by (auto intro: iseq_trivial)
|
|
50 |
|
|
51 |
lemma liseq'_ge1: "1 \<le> liseq' xs x"
|
|
52 |
apply (simp add: liseq'_def)
|
|
53 |
apply (subgoal_tac "iseq xs (Suc x) \<inter> {is. Max is = x} \<noteq> {}")
|
|
54 |
apply (simp add: Max_ge_iff iseq_finite)
|
|
55 |
apply (rule_tac x="{x}" in bexI)
|
|
56 |
apply (auto intro: iseq_singleton)
|
|
57 |
done
|
|
58 |
|
|
59 |
lemma liseq_expand:
|
|
60 |
assumes R: "\<And>is. liseq xs i = card is \<Longrightarrow> is \<in> iseq xs i \<Longrightarrow>
|
|
61 |
(\<And>js. js \<in> iseq xs i \<Longrightarrow> card js \<le> card is) \<Longrightarrow> P"
|
|
62 |
shows "P"
|
|
63 |
proof -
|
|
64 |
have "Max (card ` iseq xs i) \<in> card ` iseq xs i"
|
|
65 |
by (rule Max_in) (simp_all add: iseq_finite iseq_nonempty)
|
|
66 |
then obtain js where js: "liseq xs i = card js" and "js \<in> iseq xs i"
|
|
67 |
by (rule imageE) (simp add: liseq_def)
|
|
68 |
moreover {
|
|
69 |
fix js'
|
|
70 |
assume "js' \<in> iseq xs i"
|
|
71 |
then have "card js' \<le> card js"
|
|
72 |
by (simp add: js [symmetric] liseq_def iseq_finite iseq_trivial)
|
|
73 |
}
|
|
74 |
ultimately show ?thesis by (rule R)
|
|
75 |
qed
|
|
76 |
|
|
77 |
lemma liseq'_expand:
|
|
78 |
assumes R: "\<And>is. liseq' xs i = card is \<Longrightarrow> is \<in> iseq xs (Suc i) \<Longrightarrow>
|
|
79 |
finite is \<Longrightarrow> Max is = i \<Longrightarrow>
|
|
80 |
(\<And>js. js \<in> iseq xs (Suc i) \<Longrightarrow> Max js = i \<Longrightarrow> card js \<le> card is) \<Longrightarrow>
|
|
81 |
is \<noteq> {} \<Longrightarrow> P"
|
|
82 |
shows "P"
|
|
83 |
proof -
|
|
84 |
have "Max (card ` (iseq xs (Suc i) \<inter> {is. Max is = i})) \<in>
|
|
85 |
card ` (iseq xs (Suc i) \<inter> {is. Max is = i})"
|
|
86 |
by (auto simp add: iseq_finite intro!: iseq_singleton Max_in)
|
|
87 |
then obtain js where js: "liseq' xs i = card js" and "js \<in> iseq xs (Suc i)"
|
|
88 |
and "finite js" and "Max js = i"
|
|
89 |
by (auto simp add: liseq'_def intro: iseq_finite')
|
|
90 |
moreover {
|
|
91 |
fix js'
|
|
92 |
assume "js' \<in> iseq xs (Suc i)" "Max js' = i"
|
|
93 |
then have "card js' \<le> card js"
|
|
94 |
by (auto simp add: js [symmetric] liseq'_def iseq_finite intro!: iseq_singleton)
|
|
95 |
}
|
|
96 |
note max = this
|
|
97 |
moreover have "card {i} \<le> card js"
|
|
98 |
by (rule max) (simp_all add: iseq_singleton)
|
|
99 |
then have "js \<noteq> {}" by auto
|
|
100 |
ultimately show ?thesis by (rule R)
|
|
101 |
qed
|
|
102 |
|
|
103 |
lemma liseq'_ge:
|
|
104 |
"j = card js \<Longrightarrow> js \<in> iseq xs (Suc i) \<Longrightarrow> Max js = i \<Longrightarrow>
|
|
105 |
js \<noteq> {} \<Longrightarrow> j \<le> liseq' xs i"
|
|
106 |
by (simp add: liseq'_def iseq_finite)
|
|
107 |
|
|
108 |
lemma liseq'_eq:
|
|
109 |
"j = card js \<Longrightarrow> js \<in> iseq xs (Suc i) \<Longrightarrow> Max js = i \<Longrightarrow>
|
|
110 |
js \<noteq> {} \<Longrightarrow> (\<And>js'. js' \<in> iseq xs (Suc i) \<Longrightarrow> Max js' = i \<Longrightarrow> finite js' \<Longrightarrow>
|
|
111 |
js' \<noteq> {} \<Longrightarrow> card js' \<le> card js) \<Longrightarrow>
|
|
112 |
j = liseq' xs i"
|
|
113 |
by (fastsimp simp add: liseq'_def iseq_finite
|
|
114 |
intro: Max_eqI [symmetric])
|
|
115 |
|
|
116 |
lemma liseq_ge:
|
|
117 |
"j = card js \<Longrightarrow> js \<in> iseq xs i \<Longrightarrow> j \<le> liseq xs i"
|
|
118 |
by (auto simp add: liseq_def iseq_finite)
|
|
119 |
|
|
120 |
lemma liseq_eq:
|
|
121 |
"j = card js \<Longrightarrow> js \<in> iseq xs i \<Longrightarrow>
|
|
122 |
(\<And>js'. js' \<in> iseq xs i \<Longrightarrow> finite js' \<Longrightarrow>
|
|
123 |
js' \<noteq> {} \<Longrightarrow> card js' \<le> card js) \<Longrightarrow>
|
|
124 |
j = liseq xs i"
|
|
125 |
by (fastsimp simp add: liseq_def iseq_finite
|
|
126 |
intro: Max_eqI [symmetric])
|
|
127 |
|
|
128 |
lemma max_notin: "finite xs \<Longrightarrow> Max xs < x \<Longrightarrow> x \<notin> xs"
|
|
129 |
by (cases "xs = {}") auto
|
|
130 |
|
|
131 |
lemma iseq_insert:
|
|
132 |
"xs (Max is) \<le> xs i \<Longrightarrow> is \<in> iseq xs i \<Longrightarrow>
|
|
133 |
is \<union> {i} \<in> iseq xs (Suc i)"
|
|
134 |
apply (frule iseq_finite')
|
|
135 |
apply (cases "is = {}")
|
|
136 |
apply (auto simp add: iseq_def)
|
|
137 |
apply (rule order_trans [of _ "xs (Max is)"])
|
|
138 |
apply auto
|
|
139 |
apply (thin_tac "\<forall>a\<in>is. a < i")
|
|
140 |
apply (drule_tac x=ia in bspec)
|
|
141 |
apply assumption
|
|
142 |
apply (drule_tac x="Max is" in bspec)
|
|
143 |
apply (auto intro: Max_in)
|
|
144 |
done
|
|
145 |
|
|
146 |
lemma iseq_diff: "is \<in> iseq xs (Suc (Max is)) \<Longrightarrow>
|
|
147 |
is - {Max is} \<in> iseq xs (Suc (Max (is - {Max is})))"
|
|
148 |
apply (frule iseq_finite')
|
|
149 |
apply (simp add: iseq_def less_Suc_eq_le)
|
|
150 |
done
|
|
151 |
|
|
152 |
lemma iseq_butlast:
|
|
153 |
assumes "js \<in> iseq xs (Suc i)" and "js \<noteq> {}"
|
|
154 |
and "Max js \<noteq> i"
|
|
155 |
shows "js \<in> iseq xs i"
|
|
156 |
proof -
|
|
157 |
from assms have fin: "finite js"
|
|
158 |
by (simp add: iseq_finite')
|
|
159 |
with assms have "Max js \<in> js"
|
|
160 |
by auto
|
|
161 |
with assms have "Max js < i"
|
|
162 |
by (auto simp add: iseq_def)
|
|
163 |
with fin assms have "\<forall>j\<in>js. j < i"
|
41588
|
164 |
by simp
|
41561
|
165 |
with assms show ?thesis
|
|
166 |
by (simp add: iseq_def)
|
|
167 |
qed
|
|
168 |
|
|
169 |
lemma iseq_mono: "is \<in> iseq xs i \<Longrightarrow> i \<le> j \<Longrightarrow> is \<in> iseq xs j"
|
|
170 |
by (auto simp add: iseq_def)
|
|
171 |
|
|
172 |
lemma diff_nonempty:
|
|
173 |
assumes "1 < card is"
|
|
174 |
shows "is - {i} \<noteq> {}"
|
|
175 |
proof -
|
|
176 |
from assms have fin: "finite is" by (auto intro: card_ge_0_finite)
|
|
177 |
with assms fin have "card is - 1 \<le> card (is - {i})"
|
|
178 |
by (simp add: card_Diff_singleton_if)
|
|
179 |
with assms have "0 < card (is - {i})" by simp
|
|
180 |
then show ?thesis by (simp add: card_gt_0_iff)
|
|
181 |
qed
|
|
182 |
|
|
183 |
lemma Max_diff:
|
|
184 |
assumes "1 < card is"
|
|
185 |
shows "Max (is - {Max is}) < Max is"
|
|
186 |
proof -
|
|
187 |
from assms have "finite is" by (auto intro: card_ge_0_finite)
|
|
188 |
moreover from assms have "is - {Max is} \<noteq> {}"
|
|
189 |
by (rule diff_nonempty)
|
|
190 |
ultimately show ?thesis using assms
|
|
191 |
apply (auto simp add: not_less)
|
|
192 |
apply (subgoal_tac "a \<le> Max is")
|
|
193 |
apply auto
|
|
194 |
done
|
|
195 |
qed
|
|
196 |
|
|
197 |
lemma iseq_nth: "js \<in> iseq xs l \<Longrightarrow> 1 < card js \<Longrightarrow>
|
|
198 |
xs (Max (js - {Max js})) \<le> xs (Max js)"
|
|
199 |
apply (auto simp add: iseq_def)
|
|
200 |
apply (subgoal_tac "Max (js - {Max js}) \<in> js")
|
|
201 |
apply (thin_tac "\<forall>i\<in>js. i < l")
|
|
202 |
apply (drule_tac x="Max (js - {Max js})" in bspec)
|
|
203 |
apply assumption
|
|
204 |
apply (drule_tac x="Max js" in bspec)
|
|
205 |
using card_gt_0_iff [of js]
|
|
206 |
apply simp
|
|
207 |
using Max_diff [of js]
|
|
208 |
apply simp
|
|
209 |
using Max_in [of "js - {Max js}", OF _ diff_nonempty] card_gt_0_iff [of js]
|
|
210 |
apply auto
|
|
211 |
done
|
|
212 |
|
|
213 |
lemma card_leq1_singleton:
|
|
214 |
assumes "finite xs" "xs \<noteq> {}" "card xs \<le> 1"
|
|
215 |
obtains x where "xs = {x}"
|
|
216 |
using assms
|
|
217 |
by induct simp_all
|
|
218 |
|
|
219 |
lemma longest_iseq1:
|
|
220 |
"liseq' xs i =
|
|
221 |
Max ({0} \<union> {liseq' xs j |j. j < i \<and> xs j \<le> xs i}) + 1"
|
|
222 |
proof -
|
|
223 |
have "Max ({0} \<union> {liseq' xs j |j. j < i \<and> xs j \<le> xs i}) = liseq' xs i - 1"
|
|
224 |
proof (rule Max_eqI)
|
|
225 |
fix y
|
|
226 |
assume "y \<in> {0} \<union> {liseq' xs j |j. j < i \<and> xs j \<le> xs i}"
|
|
227 |
then show "y \<le> liseq' xs i - 1"
|
|
228 |
proof
|
|
229 |
assume "y \<in> {liseq' xs j |j. j < i \<and> xs j \<le> xs i}"
|
|
230 |
then obtain j where j: "j < i" "xs j \<le> xs i" "y = liseq' xs j"
|
|
231 |
by auto
|
|
232 |
have "liseq' xs j + 1 \<le> liseq' xs i"
|
|
233 |
proof (rule liseq'_expand)
|
|
234 |
fix "is"
|
|
235 |
assume H: "liseq' xs j = card is" "is \<in> iseq xs (Suc j)"
|
|
236 |
"finite is" "Max is = j" "is \<noteq> {}"
|
|
237 |
from H j have "card is + 1 = card (is \<union> {i})"
|
|
238 |
by (simp add: card_insert max_notin)
|
|
239 |
moreover {
|
|
240 |
from H j have "xs (Max is) \<le> xs i" by simp
|
|
241 |
moreover from `j < i` have "Suc j \<le> i" by simp
|
|
242 |
with `is \<in> iseq xs (Suc j)` have "is \<in> iseq xs i"
|
|
243 |
by (rule iseq_mono)
|
|
244 |
ultimately have "is \<union> {i} \<in> iseq xs (Suc i)"
|
|
245 |
by (rule iseq_insert)
|
|
246 |
} moreover from H j have "Max (is \<union> {i}) = i" by simp
|
|
247 |
moreover have "is \<union> {i} \<noteq> {}" by simp
|
|
248 |
ultimately have "card is + 1 \<le> liseq' xs i"
|
|
249 |
by (rule liseq'_ge)
|
|
250 |
with H show ?thesis by simp
|
|
251 |
qed
|
|
252 |
with j show "y \<le> liseq' xs i - 1"
|
|
253 |
by simp
|
|
254 |
qed simp
|
|
255 |
next
|
|
256 |
have "liseq' xs i \<le> 1 \<or>
|
|
257 |
(\<exists>j. liseq' xs i - 1 = liseq' xs j \<and> j < i \<and> xs j \<le> xs i)"
|
|
258 |
proof (rule liseq'_expand)
|
|
259 |
fix "is"
|
|
260 |
assume H: "liseq' xs i = card is" "is \<in> iseq xs (Suc i)"
|
|
261 |
"finite is" "Max is = i" "is \<noteq> {}"
|
|
262 |
assume R: "\<And>js. js \<in> iseq xs (Suc i) \<Longrightarrow> Max js = i \<Longrightarrow>
|
|
263 |
card js \<le> card is"
|
|
264 |
show ?thesis
|
|
265 |
proof (cases "card is \<le> 1")
|
|
266 |
case True with H show ?thesis by simp
|
|
267 |
next
|
|
268 |
case False
|
|
269 |
then have "1 < card is" by simp
|
|
270 |
then have "Max (is - {Max is}) < Max is"
|
|
271 |
by (rule Max_diff)
|
|
272 |
from `is \<in> iseq xs (Suc i)` `1 < card is`
|
|
273 |
have "xs (Max (is - {Max is})) \<le> xs (Max is)"
|
|
274 |
by (rule iseq_nth)
|
|
275 |
have "card is - 1 = liseq' xs (Max (is - {i}))"
|
|
276 |
proof (rule liseq'_eq)
|
|
277 |
from `Max is = i` [symmetric] `finite is` `is \<noteq> {}`
|
|
278 |
show "card is - 1 = card (is - {i})" by simp
|
|
279 |
next
|
|
280 |
from `is \<in> iseq xs (Suc i)` `Max is = i` [symmetric]
|
|
281 |
show "is - {i} \<in> iseq xs (Suc (Max (is - {i})))"
|
|
282 |
by simp (rule iseq_diff)
|
|
283 |
next
|
|
284 |
from `1 < card is`
|
|
285 |
show "is - {i} \<noteq> {}" by (rule diff_nonempty)
|
|
286 |
next
|
|
287 |
fix js
|
|
288 |
assume "js \<in> iseq xs (Suc (Max (is - {i})))"
|
|
289 |
"Max js = Max (is - {i})" "finite js" "js \<noteq> {}"
|
|
290 |
from `xs (Max (is - {Max is})) \<le> xs (Max is)`
|
|
291 |
`Max js = Max (is - {i})` `Max is = i`
|
|
292 |
have "xs (Max js) \<le> xs i" by simp
|
|
293 |
moreover from `Max is = i` `Max (is - {Max is}) < Max is`
|
|
294 |
have "Suc (Max (is - {i})) \<le> i"
|
|
295 |
by simp
|
|
296 |
with `js \<in> iseq xs (Suc (Max (is - {i})))`
|
|
297 |
have "js \<in> iseq xs i"
|
|
298 |
by (rule iseq_mono)
|
|
299 |
ultimately have "js \<union> {i} \<in> iseq xs (Suc i)"
|
|
300 |
by (rule iseq_insert)
|
|
301 |
moreover from `js \<noteq> {}` `finite js` `Max js = Max (is - {i})`
|
|
302 |
`Max is = i` [symmetric] `Max (is - {Max is}) < Max is`
|
|
303 |
have "Max (js \<union> {i}) = i"
|
|
304 |
by simp
|
|
305 |
ultimately have "card (js \<union> {i}) \<le> card is" by (rule R)
|
|
306 |
moreover from `Max is = i` [symmetric] `finite js`
|
|
307 |
`Max (is - {Max is}) < Max is` `Max js = Max (is - {i})`
|
|
308 |
have "i \<notin> js" by (simp add: max_notin)
|
|
309 |
with `finite js`
|
|
310 |
have "card (js \<union> {i}) = card ((js \<union> {i}) - {i}) + 1"
|
|
311 |
by simp
|
|
312 |
ultimately show "card js \<le> card (is - {i})"
|
|
313 |
using `i \<notin> js` `Max is = i` [symmetric] `is \<noteq> {}` `finite is`
|
|
314 |
by simp
|
|
315 |
qed simp
|
|
316 |
with H `Max (is - {Max is}) < Max is`
|
|
317 |
`xs (Max (is - {Max is})) \<le> xs (Max is)`
|
|
318 |
show ?thesis by auto
|
|
319 |
qed
|
|
320 |
qed
|
|
321 |
then show "liseq' xs i - 1 \<in> {0} \<union>
|
|
322 |
{liseq' xs j |j. j < i \<and> xs j \<le> xs i}" by simp
|
|
323 |
qed simp
|
|
324 |
moreover have "1 \<le> liseq' xs i" by (rule liseq'_ge1)
|
|
325 |
ultimately show ?thesis by simp
|
|
326 |
qed
|
|
327 |
|
|
328 |
lemma longest_iseq2': "liseq xs i < liseq' xs i \<Longrightarrow>
|
|
329 |
liseq xs (Suc i) = liseq' xs i"
|
|
330 |
apply (rule_tac xs=xs and i=i in liseq'_expand)
|
|
331 |
apply simp
|
|
332 |
apply (rule liseq_eq [symmetric])
|
|
333 |
apply (rule refl)
|
|
334 |
apply assumption
|
|
335 |
apply (case_tac "Max js' = i")
|
|
336 |
apply simp
|
|
337 |
apply (drule_tac js=js' in iseq_butlast)
|
|
338 |
apply assumption+
|
|
339 |
apply (drule_tac js=js' in liseq_ge [OF refl])
|
|
340 |
apply simp
|
|
341 |
done
|
|
342 |
|
|
343 |
lemma longest_iseq2: "liseq xs i < liseq' xs i \<Longrightarrow>
|
|
344 |
liseq xs i + 1 = liseq' xs i"
|
|
345 |
apply (rule_tac xs=xs and i=i in liseq'_expand)
|
|
346 |
apply simp
|
|
347 |
apply (rule_tac xs=xs and i=i in liseq_expand)
|
|
348 |
apply (drule_tac s="Max is" in sym)
|
|
349 |
apply simp
|
|
350 |
apply (case_tac "card is \<le> 1")
|
|
351 |
apply simp
|
|
352 |
apply (drule iseq_diff)
|
|
353 |
apply (drule_tac i="Suc (Max (is - {Max is}))" and j="Max is" in iseq_mono)
|
|
354 |
apply (simp add: less_eq_Suc_le [symmetric])
|
|
355 |
apply (rule Max_diff)
|
|
356 |
apply simp
|
|
357 |
apply (drule_tac x="is - {Max is}" in meta_spec,
|
|
358 |
drule meta_mp, assumption)
|
|
359 |
apply simp
|
|
360 |
done
|
|
361 |
|
|
362 |
lemma longest_iseq3:
|
|
363 |
"liseq xs j = liseq' xs i \<Longrightarrow> xs i \<le> xs j \<Longrightarrow> i < j \<Longrightarrow>
|
|
364 |
liseq xs (Suc j) = liseq xs j + 1"
|
|
365 |
apply (rule_tac xs=xs and i=j in liseq_expand)
|
|
366 |
apply simp
|
|
367 |
apply (rule_tac xs=xs and i=i in liseq'_expand)
|
|
368 |
apply simp
|
|
369 |
apply (rule_tac js="isa \<union> {j}" in liseq_eq [symmetric])
|
|
370 |
apply (simp add: card_insert card_Diff_singleton_if max_notin)
|
|
371 |
apply (rule iseq_insert)
|
|
372 |
apply simp
|
|
373 |
apply (erule iseq_mono)
|
|
374 |
apply simp
|
|
375 |
apply (case_tac "j = Max js'")
|
|
376 |
apply simp
|
|
377 |
apply (drule iseq_diff)
|
|
378 |
apply (drule_tac x="js' - {j}" in meta_spec)
|
|
379 |
apply (drule meta_mp)
|
|
380 |
apply simp
|
|
381 |
apply (case_tac "card js' \<le> 1")
|
|
382 |
apply (erule_tac xs=js' in card_leq1_singleton)
|
|
383 |
apply assumption+
|
|
384 |
apply (simp add: iseq_trivial)
|
|
385 |
apply (erule iseq_mono)
|
|
386 |
apply (simp add: less_eq_Suc_le [symmetric])
|
|
387 |
apply (rule Max_diff)
|
|
388 |
apply simp
|
|
389 |
apply (rule le_diff_iff [THEN iffD1, of 1])
|
|
390 |
apply (simp add: card_0_eq [symmetric] del: card_0_eq)
|
|
391 |
apply (simp add: card_insert)
|
|
392 |
apply (subgoal_tac "card (js' - {j}) = card js' - 1")
|
|
393 |
apply (simp add: card_insert card_Diff_singleton_if max_notin)
|
|
394 |
apply (frule_tac A=js' in Max_in)
|
|
395 |
apply assumption
|
|
396 |
apply (simp add: card_Diff_singleton_if)
|
|
397 |
apply (drule_tac js=js' in iseq_butlast)
|
|
398 |
apply assumption
|
|
399 |
apply (erule not_sym)
|
|
400 |
apply (drule_tac x=js' in meta_spec)
|
|
401 |
apply (drule meta_mp)
|
|
402 |
apply assumption
|
|
403 |
apply (simp add: card_insert_disjoint max_notin)
|
|
404 |
done
|
|
405 |
|
|
406 |
lemma longest_iseq4:
|
|
407 |
"liseq xs j = liseq' xs i \<Longrightarrow> xs i \<le> xs j \<Longrightarrow> i < j \<Longrightarrow>
|
|
408 |
liseq' xs j = liseq' xs i + 1"
|
|
409 |
apply (rule_tac xs=xs and i=j in liseq_expand)
|
|
410 |
apply simp
|
|
411 |
apply (rule_tac xs=xs and i=i in liseq'_expand)
|
|
412 |
apply simp
|
|
413 |
apply (rule_tac js="isa \<union> {j}" in liseq'_eq [symmetric])
|
|
414 |
apply (simp add: card_insert card_Diff_singleton_if max_notin)
|
|
415 |
apply (rule iseq_insert)
|
|
416 |
apply simp
|
|
417 |
apply (erule iseq_mono)
|
|
418 |
apply simp
|
|
419 |
apply simp
|
|
420 |
apply simp
|
|
421 |
apply (drule_tac s="Max js'" in sym)
|
|
422 |
apply simp
|
|
423 |
apply (drule iseq_diff)
|
|
424 |
apply (drule_tac x="js' - {j}" in meta_spec)
|
|
425 |
apply (drule meta_mp)
|
|
426 |
apply simp
|
|
427 |
apply (case_tac "card js' \<le> 1")
|
|
428 |
apply (erule_tac xs=js' in card_leq1_singleton)
|
|
429 |
apply assumption+
|
|
430 |
apply (simp add: iseq_trivial)
|
|
431 |
apply (erule iseq_mono)
|
|
432 |
apply (simp add: less_eq_Suc_le [symmetric])
|
|
433 |
apply (rule Max_diff)
|
|
434 |
apply simp
|
|
435 |
apply (rule le_diff_iff [THEN iffD1, of 1])
|
|
436 |
apply (simp add: card_0_eq [symmetric] del: card_0_eq)
|
|
437 |
apply (simp add: card_insert)
|
|
438 |
apply (subgoal_tac "card (js' - {j}) = card js' - 1")
|
|
439 |
apply (simp add: card_insert card_Diff_singleton_if max_notin)
|
|
440 |
apply (frule_tac A=js' in Max_in)
|
|
441 |
apply assumption
|
|
442 |
apply (simp add: card_Diff_singleton_if)
|
|
443 |
done
|
|
444 |
|
|
445 |
lemma longest_iseq5: "liseq' xs i \<le> liseq xs i \<Longrightarrow>
|
|
446 |
liseq xs (Suc i) = liseq xs i"
|
|
447 |
apply (rule_tac i=i and xs=xs in liseq'_expand)
|
|
448 |
apply simp
|
|
449 |
apply (rule_tac xs=xs and i=i in liseq_expand)
|
|
450 |
apply simp
|
|
451 |
apply (rule liseq_eq [symmetric])
|
|
452 |
apply (rule refl)
|
|
453 |
apply (erule iseq_mono)
|
|
454 |
apply simp
|
|
455 |
apply (case_tac "Max js' = i")
|
|
456 |
apply (drule_tac x=js' in meta_spec)
|
|
457 |
apply simp
|
|
458 |
apply (drule iseq_butlast, assumption, assumption)
|
|
459 |
apply simp
|
|
460 |
done
|
|
461 |
|
|
462 |
lemma liseq_empty: "liseq xs 0 = 0"
|
|
463 |
apply (rule_tac js="{}" in liseq_eq [symmetric])
|
|
464 |
apply simp
|
|
465 |
apply (rule iseq_trivial)
|
|
466 |
apply (simp add: iseq_def)
|
|
467 |
done
|
|
468 |
|
|
469 |
lemma liseq'_singleton: "liseq' xs 0 = 1"
|
|
470 |
by (simp add: longest_iseq1 [of _ 0])
|
|
471 |
|
|
472 |
lemma liseq_singleton: "liseq xs (Suc 0) = Suc 0"
|
|
473 |
by (simp add: longest_iseq2' liseq_empty liseq'_singleton)
|
|
474 |
|
|
475 |
lemma liseq'_Suc_unfold:
|
|
476 |
"A j \<le> x \<Longrightarrow>
|
|
477 |
(insert 0 {liseq' A j' |j'. j' < Suc j \<and> A j' \<le> x}) =
|
|
478 |
(insert 0 {liseq' A j' |j'. j' < j \<and> A j' \<le> x}) \<union>
|
|
479 |
{liseq' A j}"
|
|
480 |
by (auto simp add: less_Suc_eq)
|
|
481 |
|
|
482 |
lemma liseq'_Suc_unfold':
|
|
483 |
"\<not> (A j \<le> x) \<Longrightarrow>
|
|
484 |
{liseq' A j' |j'. j' < Suc j \<and> A j' \<le> x} =
|
|
485 |
{liseq' A j' |j'. j' < j \<and> A j' \<le> x}"
|
|
486 |
by (auto simp add: less_Suc_eq)
|
|
487 |
|
|
488 |
lemma iseq_card_limit:
|
|
489 |
assumes "is \<in> iseq A i"
|
|
490 |
shows "card is \<le> i"
|
|
491 |
proof -
|
|
492 |
from assms have "is \<subseteq> {0..<i}"
|
|
493 |
by (auto simp add: iseq_def)
|
|
494 |
with finite_atLeastLessThan have "card is \<le> card {0..<i}"
|
|
495 |
by (rule card_mono)
|
|
496 |
with card_atLeastLessThan show ?thesis by simp
|
|
497 |
qed
|
|
498 |
|
|
499 |
lemma liseq_limit: "liseq A i \<le> i"
|
|
500 |
by (rule_tac xs=A and i=i in liseq_expand)
|
|
501 |
(simp add: iseq_card_limit)
|
|
502 |
|
|
503 |
lemma liseq'_limit: "liseq' A i \<le> i + 1"
|
|
504 |
by (rule_tac xs=A and i=i in liseq'_expand)
|
|
505 |
(simp add: iseq_card_limit)
|
|
506 |
|
|
507 |
definition max_ext :: "(nat \<Rightarrow> 'a::linorder) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
|
|
508 |
"max_ext A i j = Max ({0} \<union> {liseq' A j' |j'. j' < j \<and> A j' \<le> A i})"
|
|
509 |
|
|
510 |
lemma max_ext_limit: "max_ext A i j \<le> j"
|
|
511 |
apply (auto simp add: max_ext_def)
|
|
512 |
apply (drule Suc_leI)
|
|
513 |
apply (cut_tac i=j' and A=A in liseq'_limit)
|
|
514 |
apply simp
|
|
515 |
done
|
|
516 |
|
|
517 |
|
|
518 |
text {* Proof functions *}
|
|
519 |
|
|
520 |
abbreviation (input)
|
|
521 |
"arr_conv a \<equiv> (\<lambda>n. a (int n))"
|
|
522 |
|
|
523 |
lemma idx_conv_suc:
|
|
524 |
"0 \<le> i \<Longrightarrow> nat (i + 1) = nat i + 1"
|
|
525 |
by simp
|
|
526 |
|
|
527 |
abbreviation liseq_ends_at' :: "(int \<Rightarrow> 'a::linorder) \<Rightarrow> int \<Rightarrow> int" where
|
|
528 |
"liseq_ends_at' A i \<equiv> int (liseq' (\<lambda>l. A (int l)) (nat i))"
|
|
529 |
|
|
530 |
abbreviation liseq_prfx' :: "(int \<Rightarrow> 'a::linorder) \<Rightarrow> int \<Rightarrow> int" where
|
|
531 |
"liseq_prfx' A i \<equiv> int (liseq (\<lambda>l. A (int l)) (nat i))"
|
|
532 |
|
|
533 |
abbreviation max_ext' :: "(int \<Rightarrow> 'a::linorder) \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int" where
|
|
534 |
"max_ext' A i j \<equiv> int (max_ext (\<lambda>l. A (int l)) (nat i) (nat j))"
|
|
535 |
|
|
536 |
spark_proof_functions
|
|
537 |
liseq_ends_at = "liseq_ends_at' :: (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> int"
|
|
538 |
liseq_prfx = "liseq_prfx' :: (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> int"
|
|
539 |
max_ext = "max_ext' :: (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int"
|
|
540 |
|
|
541 |
|
|
542 |
text {* The verification conditions *}
|
|
543 |
|
|
544 |
spark_open "liseq/liseq_length.siv"
|
|
545 |
|
|
546 |
spark_vc procedure_liseq_length_5
|
|
547 |
by (simp_all add: liseq_singleton liseq'_singleton)
|
|
548 |
|
|
549 |
spark_vc procedure_liseq_length_6
|
|
550 |
proof -
|
|
551 |
from H1 H2 H3 H4
|
|
552 |
have eq: "liseq (arr_conv a) (nat i) =
|
|
553 |
liseq' (arr_conv a) (nat pmax)"
|
|
554 |
by simp
|
|
555 |
from H14 H3 H4
|
|
556 |
have pmax1: "arr_conv a (nat pmax) \<le> arr_conv a (nat i)"
|
|
557 |
by simp
|
|
558 |
from H3 H4 have pmax2: "nat pmax < nat i"
|
|
559 |
by simp
|
|
560 |
{
|
|
561 |
fix i2
|
|
562 |
assume i2: "0 \<le> i2" "i2 \<le> i"
|
|
563 |
have "(l(i := l pmax + 1)) i2 =
|
|
564 |
int (liseq' (arr_conv a) (nat i2))"
|
|
565 |
proof (cases "i2 = i")
|
|
566 |
case True
|
|
567 |
from eq pmax1 pmax2 have "liseq' (arr_conv a) (nat i) =
|
|
568 |
liseq' (arr_conv a) (nat pmax) + 1"
|
|
569 |
by (rule longest_iseq4)
|
|
570 |
with True H1 H3 H4 show ?thesis
|
|
571 |
by simp
|
|
572 |
next
|
|
573 |
case False
|
|
574 |
with H1 i2 show ?thesis
|
|
575 |
by simp
|
|
576 |
qed
|
|
577 |
}
|
|
578 |
then show ?C1 by simp
|
|
579 |
from eq pmax1 pmax2
|
|
580 |
have "liseq (arr_conv a) (Suc (nat i)) =
|
|
581 |
liseq (arr_conv a) (nat i) + 1"
|
|
582 |
by (rule longest_iseq3)
|
|
583 |
with H2 H3 H4 show ?C2
|
|
584 |
by (simp add: idx_conv_suc)
|
|
585 |
qed
|
|
586 |
|
|
587 |
spark_vc procedure_liseq_length_7
|
|
588 |
proof -
|
|
589 |
from H1 show ?C1
|
|
590 |
by (simp add: max_ext_def longest_iseq1 [of _ "nat i"])
|
|
591 |
from H6
|
|
592 |
have m: "max_ext (arr_conv a) (nat i) (nat i) + 1 =
|
|
593 |
liseq' (arr_conv a) (nat i)"
|
|
594 |
by (simp add: max_ext_def longest_iseq1 [of _ "nat i"])
|
|
595 |
with H2 H18
|
|
596 |
have gt: "liseq (arr_conv a) (nat i) < liseq' (arr_conv a) (nat i)"
|
|
597 |
by simp
|
|
598 |
then have "liseq' (arr_conv a) (nat i) = liseq (arr_conv a) (nat i) + 1"
|
|
599 |
by (rule longest_iseq2 [symmetric])
|
|
600 |
with H2 m show ?C2 by simp
|
|
601 |
from gt have "liseq (arr_conv a) (Suc (nat i)) = liseq' (arr_conv a) (nat i)"
|
|
602 |
by (rule longest_iseq2')
|
|
603 |
with m H6 show ?C3 by (simp add: idx_conv_suc)
|
|
604 |
qed
|
|
605 |
|
|
606 |
spark_vc procedure_liseq_length_8
|
|
607 |
proof -
|
|
608 |
{
|
|
609 |
fix i2
|
|
610 |
assume i2: "0 \<le> i2" "i2 \<le> i"
|
|
611 |
have "(l(i := max_ext' a i i + 1)) i2 =
|
|
612 |
int (liseq' (arr_conv a) (nat i2))"
|
|
613 |
proof (cases "i2 = i")
|
|
614 |
case True
|
|
615 |
with H1 show ?thesis
|
|
616 |
by (simp add: max_ext_def longest_iseq1 [of _ "nat i"])
|
|
617 |
next
|
|
618 |
case False
|
|
619 |
with H1 i2 show ?thesis by simp
|
|
620 |
qed
|
|
621 |
}
|
|
622 |
then show ?C1 by simp
|
|
623 |
from H2 H6 H18
|
|
624 |
have "liseq' (arr_conv a) (nat i) \<le> liseq (arr_conv a) (nat i)"
|
|
625 |
by (simp add: max_ext_def longest_iseq1 [of _ "nat i"])
|
|
626 |
then have "liseq (arr_conv a) (Suc (nat i)) = liseq (arr_conv a) (nat i)"
|
|
627 |
by (rule longest_iseq5)
|
|
628 |
with H2 H6 show ?C2 by (simp add: idx_conv_suc)
|
|
629 |
qed
|
|
630 |
|
|
631 |
spark_vc procedure_liseq_length_12
|
|
632 |
by (simp add: max_ext_def)
|
|
633 |
|
|
634 |
spark_vc procedure_liseq_length_13
|
|
635 |
using H1 H6 H13 H21 H22
|
|
636 |
by (simp add: max_ext_def
|
|
637 |
idx_conv_suc liseq'_Suc_unfold max_def del: Max_less_iff)
|
|
638 |
|
|
639 |
spark_vc procedure_liseq_length_14
|
|
640 |
using H1 H6 H13 H21
|
|
641 |
by (cases "a j \<le> a i")
|
|
642 |
(simp_all add: max_ext_def
|
|
643 |
idx_conv_suc liseq'_Suc_unfold liseq'_Suc_unfold')
|
|
644 |
|
|
645 |
spark_vc procedure_liseq_length_19
|
|
646 |
using H3 H4 H5 H8 H9
|
|
647 |
apply (rule_tac y="int (nat i)" in order_trans)
|
|
648 |
apply (cut_tac A="arr_conv a" and i="nat i" and j="nat i" in max_ext_limit)
|
|
649 |
apply simp_all
|
|
650 |
done
|
|
651 |
|
|
652 |
spark_vc procedure_liseq_length_23
|
|
653 |
using H2 H3 H4 H7 H8 H11
|
|
654 |
apply (rule_tac y="int (nat i)" in order_trans)
|
|
655 |
apply (cut_tac A="arr_conv a" and i="nat i" in liseq_limit)
|
|
656 |
apply simp_all
|
|
657 |
done
|
|
658 |
|
|
659 |
spark_vc procedure_liseq_length_29
|
|
660 |
using H2 H3 H8 H13
|
|
661 |
by (simp add: add1_zle_eq [symmetric])
|
|
662 |
|
|
663 |
spark_end
|
|
664 |
|
|
665 |
end
|