1 (* Author: Lukas Bulwahn, TU Muenchen *)
3 theory Imperative_Quicksort
4 imports "~~/src/HOL/Imperative_HOL/Imperative_HOL" Subarray Multiset Efficient_Nat
7 text {* We prove QuickSort correct in the Relational Calculus. *}
9 definition swap :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap"
13 x \<leftarrow> nth arr i;
14 y \<leftarrow> nth arr j;
21 assumes "crel (swap a i j) h h' rs"
22 shows "multiset_of (get_array a h')
23 = multiset_of (get_array a h)"
26 by (auto simp add: Heap.length_def multiset_of_swap dest: sym [of _ "h'"] elim!: crelE crel_nth crel_return crel_upd)
28 function part1 :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap"
30 "part1 a left right p = (
31 if (right \<le> left) then return right
33 v \<leftarrow> nth a left;
34 (if (v \<le> p) then (part1 a (left + 1) right p)
35 else (do swap a left right;
36 part1 a left (right - 1) p done))
38 by pat_completeness auto
41 by (relation "measure (\<lambda>(_,l,r,_). r - l )") auto
43 declare part1.simps[simp del]
46 assumes "crel (part1 a l r p) h h' rs"
47 shows "multiset_of (get_array a h')
48 = multiset_of (get_array a h)"
50 proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
51 case (1 a l r p h h' rs)
53 unfolding part1.simps [of a l r p]
54 by (elim crelE crel_if crel_return crel_nth) (auto simp add: swap_permutes)
57 lemma part_returns_index_in_bounds:
58 assumes "crel (part1 a l r p) h h' rs"
60 shows "l \<le> rs \<and> rs \<le> r"
62 proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
63 case (1 a l r p h h' rs)
64 note cr = `crel (part1 a l r p) h h' rs`
66 proof (cases "r \<le> l")
67 case True (* Terminating case *)
68 with cr `l \<le> r` show ?thesis
69 unfolding part1.simps[of a l r p]
70 by (elim crelE crel_if crel_return crel_nth) auto
72 case False (* recursive case *)
73 note rec_condition = this
74 let ?v = "get_array a h ! l"
76 proof (cases "?v \<le> p")
79 have rec1: "crel (part1 a (l + 1) r p) h h' rs"
80 unfolding part1.simps[of a l r p]
81 by (elim crelE crel_nth crel_if crel_return) auto
82 from rec_condition have "l + 1 \<le> r" by arith
83 from 1(1)[OF rec_condition True rec1 `l + 1 \<le> r`]
88 obtain h1 where swp: "crel (swap a l r) h h1 ()"
89 and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
90 unfolding part1.simps[of a l r p]
91 by (elim crelE crel_nth crel_if crel_return) auto
92 from rec_condition have "l \<le> r - 1" by arith
93 from 1(2) [OF rec_condition False rec2 `l \<le> r - 1`] show ?thesis by fastsimp
98 lemma part_length_remains:
99 assumes "crel (part1 a l r p) h h' rs"
100 shows "Heap.length a h = Heap.length a h'"
102 proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
103 case (1 a l r p h h' rs)
104 note cr = `crel (part1 a l r p) h h' rs`
107 proof (cases "r \<le> l")
108 case True (* Terminating case *)
110 unfolding part1.simps[of a l r p]
111 by (elim crelE crel_if crel_return crel_nth) auto
113 case False (* recursive case *)
114 with cr 1 show ?thesis
115 unfolding part1.simps [of a l r p] swap_def
116 by (auto elim!: crelE crel_if crel_nth crel_return crel_upd) fastsimp
120 lemma part_outer_remains:
121 assumes "crel (part1 a l r p) h h' rs"
122 shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
124 proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
125 case (1 a l r p h h' rs)
126 note cr = `crel (part1 a l r p) h h' rs`
129 proof (cases "r \<le> l")
130 case True (* Terminating case *)
132 unfolding part1.simps[of a l r p]
133 by (elim crelE crel_if crel_return crel_nth) auto
135 case False (* recursive case *)
136 note rec_condition = this
137 let ?v = "get_array a h ! l"
139 proof (cases "?v \<le> p")
142 have rec1: "crel (part1 a (l + 1) r p) h h' rs"
143 unfolding part1.simps[of a l r p]
144 by (elim crelE crel_nth crel_if crel_return) auto
145 from 1(1)[OF rec_condition True rec1]
146 show ?thesis by fastsimp
149 with rec_condition cr
150 obtain h1 where swp: "crel (swap a l r) h h1 ()"
151 and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
152 unfolding part1.simps[of a l r p]
153 by (elim crelE crel_nth crel_if crel_return) auto
154 from swp rec_condition have
155 "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array a h ! i = get_array a h1 ! i"
157 by (elim crelE crel_nth crel_upd crel_return) auto
158 with 1(2) [OF rec_condition False rec2] show ?thesis by fastsimp
164 lemma part_partitions:
165 assumes "crel (part1 a l r p) h h' rs"
166 shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> get_array (a::nat array) h' ! i \<le> p)
167 \<and> (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> get_array a h' ! i \<ge> p)"
169 proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
170 case (1 a l r p h h' rs)
171 note cr = `crel (part1 a l r p) h h' rs`
174 proof (cases "r \<le> l")
175 case True (* Terminating case *)
176 with cr have "rs = r"
177 unfolding part1.simps[of a l r p]
178 by (elim crelE crel_if crel_return crel_nth) auto
182 case False (* recursive case *)
184 let ?v = "get_array a h ! l"
186 proof (cases "?v \<le> p")
189 have rec1: "crel (part1 a (l + 1) r p) h h' rs"
190 unfolding part1.simps[of a l r p]
191 by (elim crelE crel_nth crel_if crel_return) auto
192 from True part_outer_remains[OF rec1] have a_l: "get_array a h' ! l \<le> p"
194 have "\<forall>i. (l \<le> i = (l = i \<or> Suc l \<le> i))" by arith
195 with 1(1)[OF False True rec1] a_l show ?thesis
200 obtain h1 where swp: "crel (swap a l r) h h1 ()"
201 and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
202 unfolding part1.simps[of a l r p]
203 by (elim crelE crel_nth crel_if crel_return) auto
204 from swp False have "get_array a h1 ! r \<ge> p"
206 by (auto simp add: Heap.length_def elim!: crelE crel_nth crel_upd crel_return)
207 with part_outer_remains [OF rec2] lr have a_r: "get_array a h' ! r \<ge> p"
209 have "\<forall>i. (i \<le> r = (i = r \<or> i \<le> r - 1))" by arith
210 with 1(2)[OF lr False rec2] a_r show ?thesis
217 fun partition :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap"
219 "partition a left right = (do
220 pivot \<leftarrow> nth a right;
221 middle \<leftarrow> part1 a left (right - 1) pivot;
222 v \<leftarrow> nth a middle;
223 m \<leftarrow> return (if (v \<le> pivot) then (middle + 1) else middle);
228 declare partition.simps[simp del]
230 lemma partition_permutes:
231 assumes "crel (partition a l r) h h' rs"
232 shows "multiset_of (get_array a h')
233 = multiset_of (get_array a h)"
235 from assms part_permutes swap_permutes show ?thesis
236 unfolding partition.simps
237 by (elim crelE crel_return crel_nth crel_if crel_upd) auto
240 lemma partition_length_remains:
241 assumes "crel (partition a l r) h h' rs"
242 shows "Heap.length a h = Heap.length a h'"
244 from assms part_length_remains show ?thesis
245 unfolding partition.simps swap_def
246 by (elim crelE crel_return crel_nth crel_if crel_upd) auto
249 lemma partition_outer_remains:
250 assumes "crel (partition a l r) h h' rs"
252 shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
254 from assms part_outer_remains part_returns_index_in_bounds show ?thesis
255 unfolding partition.simps swap_def
256 by (elim crelE crel_return crel_nth crel_if crel_upd) fastsimp
259 lemma partition_returns_index_in_bounds:
260 assumes crel: "crel (partition a l r) h h' rs"
262 shows "l \<le> rs \<and> rs \<le> r"
264 from crel obtain middle h'' p where part: "crel (part1 a l (r - 1) p) h h'' middle"
265 and rs_equals: "rs = (if get_array a h'' ! middle \<le> get_array a h ! r then middle + 1
267 unfolding partition.simps
268 by (elim crelE crel_return crel_nth crel_if crel_upd) simp
269 from `l < r` have "l \<le> r - 1" by arith
270 from part_returns_index_in_bounds[OF part this] rs_equals `l < r` show ?thesis by auto
273 lemma partition_partitions:
274 assumes crel: "crel (partition a l r) h h' rs"
276 shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> get_array (a::nat array) h' ! i \<le> get_array a h' ! rs) \<and>
277 (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> get_array a h' ! rs \<le> get_array a h' ! i)"
279 let ?pivot = "get_array a h ! r"
280 from crel obtain middle h1 where part: "crel (part1 a l (r - 1) ?pivot) h h1 middle"
281 and swap: "crel (swap a rs r) h1 h' ()"
282 and rs_equals: "rs = (if get_array a h1 ! middle \<le> ?pivot then middle + 1
284 unfolding partition.simps
285 by (elim crelE crel_return crel_nth crel_if crel_upd) simp
286 from swap have h'_def: "h' = Heap.upd a r (get_array a h1 ! rs)
287 (Heap.upd a rs (get_array a h1 ! r) h1)"
289 by (elim crelE crel_return crel_nth crel_upd) simp
290 from swap have in_bounds: "r < Heap.length a h1 \<and> rs < Heap.length a h1"
292 by (elim crelE crel_return crel_nth crel_upd) simp
293 from swap have swap_length_remains: "Heap.length a h1 = Heap.length a h'"
294 unfolding swap_def by (elim crelE crel_return crel_nth crel_upd) auto
295 from `l < r` have "l \<le> r - 1" by simp
296 note middle_in_bounds = part_returns_index_in_bounds[OF part this]
297 from part_outer_remains[OF part] `l < r`
298 have "get_array a h ! r = get_array a h1 ! r"
301 have right_remains: "get_array a h ! r = get_array a h' ! rs"
303 by (auto simp add: Heap.length_def elim!: crelE crel_return crel_nth crel_upd) (cases "r = rs", auto)
304 from part_partitions [OF part]
306 proof (cases "get_array a h1 ! middle \<le> ?pivot")
308 with rs_equals have rs_equals: "rs = middle + 1" by simp
311 assume i_is_left: "l \<le> i \<and> i < rs"
312 with swap_length_remains in_bounds middle_in_bounds rs_equals `l < r`
313 have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
314 from i_is_left rs_equals have "l \<le> i \<and> i < middle \<or> i = middle" by arith
315 with part_partitions[OF part] right_remains True
316 have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
317 with i_props h'_def in_bounds have "get_array a h' ! i \<le> get_array a h' ! rs"
318 unfolding Heap.upd_def Heap.length_def by simp
323 assume "rs < i \<and> i \<le> r"
325 hence "(rs < i \<and> i \<le> r - 1) \<or> (rs < i \<and> i = r)" by arith
326 hence "get_array a h' ! rs \<le> get_array a h' ! i"
328 assume i_is: "rs < i \<and> i \<le> r - 1"
329 with swap_length_remains in_bounds middle_in_bounds rs_equals
330 have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
331 from part_partitions[OF part] rs_equals right_remains i_is
332 have "get_array a h' ! rs \<le> get_array a h1 ! i"
334 with i_props h'_def show ?thesis by fastsimp
336 assume i_is: "rs < i \<and> i = r"
337 with rs_equals have "Suc middle \<noteq> r" by arith
338 with middle_in_bounds `l < r` have "Suc middle \<le> r - 1" by arith
339 with part_partitions[OF part] right_remains
340 have "get_array a h' ! rs \<le> get_array a h1 ! (Suc middle)"
342 with i_is True rs_equals right_remains h'_def
343 show ?thesis using in_bounds
344 unfolding Heap.upd_def Heap.length_def
348 ultimately show ?thesis by auto
351 with rs_equals have rs_equals: "middle = rs" by simp
354 assume i_is_left: "l \<le> i \<and> i < rs"
355 with swap_length_remains in_bounds middle_in_bounds rs_equals
356 have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
357 from part_partitions[OF part] rs_equals right_remains i_is_left
358 have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
359 with i_props h'_def have "get_array a h' ! i \<le> get_array a h' ! rs"
360 unfolding Heap.upd_def by simp
365 assume "rs < i \<and> i \<le> r"
366 hence "(rs < i \<and> i \<le> r - 1) \<or> i = r" by arith
367 hence "get_array a h' ! rs \<le> get_array a h' ! i"
369 assume i_is: "rs < i \<and> i \<le> r - 1"
370 with swap_length_remains in_bounds middle_in_bounds rs_equals
371 have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
372 from part_partitions[OF part] rs_equals right_remains i_is
373 have "get_array a h' ! rs \<le> get_array a h1 ! i"
375 with i_props h'_def show ?thesis by fastsimp
378 from i_is False rs_equals right_remains h'_def
379 show ?thesis using in_bounds
380 unfolding Heap.upd_def Heap.length_def
390 function quicksort :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap"
392 "quicksort arr left right =
393 (if (right > left) then
395 pivotNewIndex \<leftarrow> partition arr left right;
396 pivotNewIndex \<leftarrow> assert (\<lambda>x. left \<le> x \<and> x \<le> right) pivotNewIndex;
397 quicksort arr left (pivotNewIndex - 1);
398 quicksort arr (pivotNewIndex + 1) right
401 by pat_completeness auto
403 (* For termination, we must show that the pivotNewIndex is between left and right *)
405 by (relation "measure (\<lambda>(a, l, r). (r - l))") auto
407 declare quicksort.simps[simp del]
410 lemma quicksort_permutes:
411 assumes "crel (quicksort a l r) h h' rs"
412 shows "multiset_of (get_array a h')
413 = multiset_of (get_array a h)"
415 proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
416 case (1 a l r h h' rs)
417 with partition_permutes show ?case
418 unfolding quicksort.simps [of a l r]
419 by (elim crel_if crelE crel_assert crel_return) auto
422 lemma length_remains:
423 assumes "crel (quicksort a l r) h h' rs"
424 shows "Heap.length a h = Heap.length a h'"
426 proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
427 case (1 a l r h h' rs)
428 with partition_length_remains show ?case
429 unfolding quicksort.simps [of a l r]
430 by (elim crel_if crelE crel_assert crel_return) auto
433 lemma quicksort_outer_remains:
434 assumes "crel (quicksort a l r) h h' rs"
435 shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
437 proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
438 case (1 a l r h h' rs)
439 note cr = `crel (quicksort a l r) h h' rs`
441 proof (cases "r > l")
443 with cr have "h' = h"
444 unfolding quicksort.simps [of a l r]
445 by (elim crel_if crel_return) auto
450 fix h1 h2 p ret1 ret2 i
451 assume part: "crel (partition a l r) h h1 p"
452 assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ret1"
453 assume qs2: "crel (quicksort a (p + 1) r) h2 h' ret2"
454 assume pivot: "l \<le> p \<and> p \<le> r"
455 assume i_outer: "i < l \<or> r < i"
456 from partition_outer_remains [OF part True] i_outer
457 have "get_array a h !i = get_array a h1 ! i" by fastsimp
459 with 1(1) [OF True pivot qs1] pivot i_outer
460 have "get_array a h1 ! i = get_array a h2 ! i" by auto
462 with qs2 1(2) [of p h2 h' ret2] True pivot i_outer
463 have "get_array a h2 ! i = get_array a h' ! i" by auto
464 ultimately have "get_array a h ! i= get_array a h' ! i" by simp
467 unfolding quicksort.simps [of a l r]
468 by (elim crel_if crelE crel_assert crel_return) auto
472 lemma quicksort_is_skip:
473 assumes "crel (quicksort a l r) h h' rs"
474 shows "r \<le> l \<longrightarrow> h = h'"
476 unfolding quicksort.simps [of a l r]
477 by (elim crel_if crel_return) auto
479 lemma quicksort_sorts:
480 assumes "crel (quicksort a l r) h h' rs"
481 assumes l_r_length: "l < Heap.length a h" "r < Heap.length a h"
482 shows "sorted (subarray l (r + 1) a h')"
484 proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
485 case (1 a l r h h' rs)
486 note cr = `crel (quicksort a l r) h h' rs`
488 proof (cases "r > l")
490 hence "l \<ge> r + 1 \<or> l = r" by arith
491 with length_remains[OF cr] 1(5) show ?thesis
492 by (auto simp add: subarray_Nil subarray_single)
497 assume part: "crel (partition a l r) h h1 p"
498 assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ()"
499 assume qs2: "crel (quicksort a (p + 1) r) h2 h' ()"
500 from partition_returns_index_in_bounds [OF part True]
501 have pivot: "l\<le> p \<and> p \<le> r" .
502 note length_remains = length_remains[OF qs2] length_remains[OF qs1] partition_length_remains[OF part]
503 from quicksort_outer_remains [OF qs2] quicksort_outer_remains [OF qs1] pivot quicksort_is_skip[OF qs1]
504 have pivot_unchanged: "get_array a h1 ! p = get_array a h' ! p" by (cases p, auto)
505 (*-- First of all, by induction hypothesis both sublists are sorted. *)
506 from 1(1)[OF True pivot qs1] length_remains pivot 1(5)
507 have IH1: "sorted (subarray l p a h2)" by (cases p, auto simp add: subarray_Nil)
508 from quicksort_outer_remains [OF qs2] length_remains
509 have left_subarray_remains: "subarray l p a h2 = subarray l p a h'"
510 by (simp add: subarray_eq_samelength_iff)
511 with IH1 have IH1': "sorted (subarray l p a h')" by simp
512 from 1(2)[OF True pivot qs2] pivot 1(5) length_remains
513 have IH2: "sorted (subarray (p + 1) (r + 1) a h')"
514 by (cases "Suc p \<le> r", auto simp add: subarray_Nil)
515 (* -- Secondly, both sublists remain partitioned. *)
516 from partition_partitions[OF part True]
517 have part_conds1: "\<forall>j. j \<in> set (subarray l p a h1) \<longrightarrow> j \<le> get_array a h1 ! p "
518 and part_conds2: "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h1) \<longrightarrow> get_array a h1 ! p \<le> j"
519 by (auto simp add: all_in_set_subarray_conv)
520 from quicksort_outer_remains [OF qs1] quicksort_permutes [OF qs1] True
521 length_remains 1(5) pivot multiset_of_sublist [of l p "get_array a h1" "get_array a h2"]
522 have multiset_partconds1: "multiset_of (subarray l p a h2) = multiset_of (subarray l p a h1)"
523 unfolding Heap.length_def subarray_def by (cases p, auto)
524 with left_subarray_remains part_conds1 pivot_unchanged
525 have part_conds2': "\<forall>j. j \<in> set (subarray l p a h') \<longrightarrow> j \<le> get_array a h' ! p"
526 by (simp, subst set_of_multiset_of[symmetric], simp)
527 (* -- These steps are the analogous for the right sublist \<dots> *)
528 from quicksort_outer_remains [OF qs1] length_remains
529 have right_subarray_remains: "subarray (p + 1) (r + 1) a h1 = subarray (p + 1) (r + 1) a h2"
530 by (auto simp add: subarray_eq_samelength_iff)
531 from quicksort_outer_remains [OF qs2] quicksort_permutes [OF qs2] True
532 length_remains 1(5) pivot multiset_of_sublist [of "p + 1" "r + 1" "get_array a h2" "get_array a h'"]
533 have multiset_partconds2: "multiset_of (subarray (p + 1) (r + 1) a h') = multiset_of (subarray (p + 1) (r + 1) a h2)"
534 unfolding Heap.length_def subarray_def by auto
535 with right_subarray_remains part_conds2 pivot_unchanged
536 have part_conds1': "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h') \<longrightarrow> get_array a h' ! p \<le> j"
537 by (simp, subst set_of_multiset_of[symmetric], simp)
538 (* -- Thirdly and finally, we show that the array is sorted
539 following from the facts above. *)
540 from True pivot 1(5) length_remains have "subarray l (r + 1) a h' = subarray l p a h' @ [get_array a h' ! p] @ subarray (p + 1) (r + 1) a h'"
541 by (simp add: subarray_nth_array_Cons, cases "l < p") (auto simp add: subarray_append subarray_Nil)
542 with IH1' IH2 part_conds1' part_conds2' pivot have ?thesis
543 unfolding subarray_def
544 apply (auto simp add: sorted_append sorted_Cons all_in_set_sublist'_conv)
545 by (auto simp add: set_sublist' dest: le_trans [of _ "get_array a h' ! p"])
547 with True cr show ?thesis
548 unfolding quicksort.simps [of a l r]
549 by (elim crel_if crel_return crelE crel_assert) auto
554 lemma quicksort_is_sort:
555 assumes crel: "crel (quicksort a 0 (Heap.length a h - 1)) h h' rs"
556 shows "get_array a h' = sort (get_array a h)"
557 proof (cases "get_array a h = []")
559 with quicksort_is_skip[OF crel] show ?thesis
560 unfolding Heap.length_def by simp
563 from quicksort_sorts [OF crel] False have "sorted (sublist' 0 (List.length (get_array a h)) (get_array a h'))"
564 unfolding Heap.length_def subarray_def by auto
565 with length_remains[OF crel] have "sorted (get_array a h')"
566 unfolding Heap.length_def by simp
567 with quicksort_permutes [OF crel] properties_for_sort show ?thesis by fastsimp
570 subsection {* No Errors in quicksort *}
571 text {* We have proved that quicksort sorts (if no exceptions occur).
572 We will now show that exceptions do not occur. *}
575 assumes "l < Heap.length a h" "r < Heap.length a h"
576 shows "noError (part1 a l r p) h"
578 proof (induct a l r p arbitrary: h rule: part1.induct)
581 unfolding part1.simps [of a l r] swap_def
582 by (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd elim!: crelE crel_upd crel_nth crel_return)
585 lemma noError_partition:
586 assumes "l < r" "l < Heap.length a h" "r < Heap.length a h"
587 shows "noError (partition a l r) h"
589 unfolding partition.simps swap_def
590 apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_part1 elim!: crelE crel_upd crel_nth crel_return)
591 apply (frule part_length_remains)
592 apply (frule part_returns_index_in_bounds)
594 apply (frule part_length_remains)
595 apply (frule part_returns_index_in_bounds)
597 apply (frule part_length_remains)
601 lemma noError_quicksort:
602 assumes "l < Heap.length a h" "r < Heap.length a h"
603 shows "noError (quicksort a l r) h"
605 proof (induct a l r arbitrary: h rule: quicksort.induct)
608 unfolding quicksort.simps [of a l ri]
609 apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_assert noError_partition)
610 apply (frule partition_returns_index_in_bounds)
612 apply (frule partition_returns_index_in_bounds)
614 apply (auto elim!: crel_assert dest!: partition_length_remains length_remains)
615 apply (subgoal_tac "Suc r \<le> ri \<or> r = ri")
618 unfolding quicksort.simps [of a "Suc ri" ri]
619 apply (auto intro!: noError_if noError_return)
624 subsection {* Example *}
626 definition "qsort a = do
627 k \<leftarrow> length a;
628 quicksort a 0 (k - 1);
632 ML {* @{code qsort} (Array.fromList [42, 2, 3, 5, 0, 1705, 8, 3, 15]) () *}
634 export_code qsort in SML_imp module_name QSort
635 export_code qsort in OCaml module_name QSort file -
636 export_code qsort in OCaml_imp module_name QSort file -
637 export_code qsort in Haskell module_name QSort file -