# HG changeset patch # User bulwahn # Date 1216488433 -7200 # Node ID d4f6e64ee7ccef5895e1b2c4e27cf84c14aa201b # Parent cf0c60e821bb1aaa75b5f9ca310df988ad74c2ef added verification framework for the HeapMonad and quicksort as example for this framework diff -r cf0c60e821bb -r d4f6e64ee7cc src/HOL/Library/Array.thy --- a/src/HOL/Library/Array.thy Sat Jul 19 11:05:18 2008 +0200 +++ b/src/HOL/Library/Array.thy Sat Jul 19 19:27:13 2008 +0200 @@ -93,7 +93,16 @@ mapM (nth a) [0..heap \ 'a) \ 'a array \ 'a array Heap" +where + "map f a = (do + n \ length a; + mapM (\n. map_entry n f a) [0.. bool) \ 'a \ 'a Heap" +where + "assert P x = (if P x then return x else raise (''assert''))" + +lemma assert_cong[fundef_cong]: +assumes "P = P'" +assumes "\x. P' x \ f x = f' x" +shows "(assert P x >>= f) = (assert P' x >>= f')" +using assms +by (auto simp add: assert_def return_bind raise_bind) + +section {* Example : Using Assert for showing termination of functions *} + +function until_zero :: "int \ int Heap" +where + "until_zero a = (if a \ 0 then return a else (do x \ return (a - 1); until_zero x done))" +by (pat_completeness, auto) + +termination +apply (relation "measure (\x. nat x)") +apply simp +apply simp +oops + + +function until_zero' :: "int \ int Heap" +where + "until_zero' a = (if a \ 0 then return a else (do x \ return (a - 1); y \ assert (\x. x < a) x; until_zero' y done))" +by (pat_completeness, auto) + +termination +apply (relation "measure (\x. nat x)") +apply simp +apply simp +done + +hide (open) const until_zero until_zero' + +text {* Also look at lemmas about assert in Relational theory. *} + +end diff -r cf0c60e821bb -r d4f6e64ee7cc src/HOL/Library/Imperative_HOL.thy --- a/src/HOL/Library/Imperative_HOL.thy Sat Jul 19 11:05:18 2008 +0200 +++ b/src/HOL/Library/Imperative_HOL.thy Sat Jul 19 19:27:13 2008 +0200 @@ -6,7 +6,7 @@ header {* Entry point into monadic imperative HOL *} theory Imperative_HOL -imports Array Ref +imports Array Ref Relational begin end diff -r cf0c60e821bb -r d4f6e64ee7cc src/HOL/Library/Relational.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Relational.thy Sat Jul 19 19:27:13 2008 +0200 @@ -0,0 +1,700 @@ +theory Relational +imports Array Ref Assert +begin + +section{* Definition of the Relational framework *} + +text {* The crel predicate states that when a computation c runs with the heap h + will result in return value r and a heap h' (if no exception occurs). *} + +definition crel :: "'a Heap \ heap \ heap \ 'a \ bool" +where + crel_def': "crel c h h' r \ Heap_Monad.execute c h = (Inl r, h')" + +lemma crel_def: -- FIXME + "crel c h h' r \ (Inl r, h') = Heap_Monad.execute c h" + unfolding crel_def' by auto + +lemma crel_deterministic: "\ crel f h h' a; crel f h h'' b \ \ (a = b) \ (h' = h'')" +unfolding crel_def' by auto + +section {* Elimination rules *} + +text {* For all commands, we define simple elimination rules. *} +(* FIXME: consumes 1 necessary ?? *) + +subsection {* Elimination rules for basic monadic commands *} + +lemma crelE[consumes 1]: + assumes "crel (f >>= g) h h'' r'" + obtains h' r where "crel f h h' r" "crel (g r) h' h'' r'" + using assms + by (auto simp add: crel_def bindM_def Let_def prod_case_beta split_def Pair_fst_snd_eq split add: sum.split_asm) + +lemma crelE'[consumes 1]: + assumes "crel (f >> g) h h'' r'" + obtains h' r where "crel f h h' r" "crel g h' h'' r'" + using assms + by (elim crelE) auto + +lemma crel_return[consumes 1]: + assumes "crel (return x) h h' r" + obtains "r = x" "h = h'" + using assms + unfolding crel_def return_def by simp + +lemma crel_raise[consumes 1]: + assumes "crel (raise x) h h' r" + obtains "False" + using assms + unfolding crel_def raise_def by simp + +lemma crel_if: + assumes "crel (if c then t else e) h h' r" + obtains "c" "crel t h h' r" + | "\c" "crel e h h' r" + using assms + unfolding crel_def by auto + +lemma crel_option_case: + assumes "crel (case x of None \ n | Some y \ s y) h h' r" + obtains "x = None" "crel n h h' r" + | y where "x = Some y" "crel (s y) h h' r" + using assms + unfolding crel_def by auto + +lemma crel_mapM: + assumes "crel (mapM f xs) h h' r" + assumes "\h h'. P f [] h h' []" + assumes "\h h1 h' x xs y ys. \ crel (f x) h h1 y; crel (mapM f xs) h1 h' ys; P f xs h1 h' ys \ \ P f (x#xs) h h' (y#ys)" + shows "P f xs h h' r" +using assms(1) +proof (induct xs arbitrary: h h' r) + case Nil with assms(2) show ?case + by (auto elim: crel_return) +next + case (Cons x xs) + from Cons(2) obtain h1 y ys where crel_f: "crel (f x) h h1 y" + and crel_mapM: "crel (mapM f xs) h1 h' ys" + and r_def: "r = y#ys" + unfolding mapM.simps run_drop + by (auto elim!: crelE crel_return) + from Cons(1)[OF crel_mapM] crel_mapM crel_f assms(3) r_def + show ?case by auto +qed + +lemma crel_heap: + assumes "crel (Heap_Monad.heap f) h h' r" + obtains "h' = snd (f h)" "r = fst (f h)" + using assms + unfolding heap_def crel_def apfst_def split_def prod_fun_def by simp_all + +subsection {* Elimination rules for array commands *} + +lemma crel_length: + assumes "crel (length a) h h' r" + obtains "h = h'" "r = Heap.length a h'" + using assms + unfolding length_def + by (elim crel_heap) simp + +(* Strong version of the lemma for this operation is missing *) +lemma crel_new_weak: + assumes "crel (Array.new n v) h h' r" + obtains "get_array r h' = List.replicate n v" + using assms unfolding Array.new_def + by (elim crel_heap) (auto simp:Heap.array_def Let_def split_def) + +lemma crel_nth[consumes 1]: + assumes "crel (nth a i) h h' r" + obtains "r = (get_array a h) ! i" "h = h'" "i < Heap.length a h" + using assms + unfolding nth_def run_drop + by (auto elim!: crelE crel_if crel_raise crel_length crel_heap) + +lemma crel_upd[consumes 1]: + assumes "crel (upd i v a) h h' r" + obtains "r = a" "h' = Heap.upd a i v h" + using assms + unfolding upd_def run_drop + by (elim crelE crel_if crel_return crel_raise + crel_length crel_heap) auto + +(* Strong version of the lemma for this operation is missing *) +lemma crel_of_list_weak: + assumes "crel (Array.of_list xs) h h' r" + obtains "get_array r h' = xs" + using assms + unfolding of_list_def + by (elim crel_heap) (simp add:get_array_init_array_list) + +lemma crel_map_entry: + assumes "crel (Array.map_entry i f a) h h' r" + obtains "r = a" "h' = Heap.upd a i (f (get_array a h ! i)) h" + using assms + unfolding Array.map_entry_def run_drop + by (elim crelE crel_upd crel_nth) auto + +lemma crel_swap: + assumes "crel (Array.swap i x a) h h' r" + obtains "r = get_array a h ! i" "h' = Heap.upd a i x h" + using assms + unfolding Array.swap_def run_drop + by (elim crelE crel_upd crel_nth crel_return) auto + +(* Strong version of the lemma for this operation is missing *) +lemma crel_make_weak: + assumes "crel (Array.make n f) h h' r" + obtains "i < n \ get_array r h' ! i = f i" + using assms + unfolding Array.make_def + by (elim crel_of_list_weak) auto + +lemma upt_conv_Cons': + assumes "Suc a \ b" + shows "[b - Suc a.. Heap.length a h" + shows "h = h' \ xs = drop (Heap.length a h - n) (get_array a h)" +using assms +proof (induct n arbitrary: xs h h') + case 0 thus ?case + by (auto elim!: crel_return simp add: Heap.length_def) +next + case (Suc n) + from Suc(3) have "[Heap.length a h - Suc n..n. map_entry n f a) [Heap.length a h - n..n. map_entry n f a) [Heap.length a h - n..n. map_entry n f a) [Heap.length a ?h1 - n..n. map_entry n f a) [Heap.length a h - n.. Heap.length a h" + assumes "i \ Heap.length a h - n" + assumes "i < Heap.length a h" + shows "get_array a h' ! i = f (get_array a h ! i)" +using assms +proof (induct n arbitrary: h h' r) + case 0 + thus ?case + by (auto elim!: crel_return) +next + case (Suc n) + let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h" + from Suc(3) have "[Heap.length a h - Suc n..n. map_entry n f a) [Heap.length a h - n.. i \ Heap.length a ?h1 - n" by arith + from crel_mapM have crel_mapM': "crel (mapM (\n. map_entry n f a) [Heap.length a ?h1 - n..n. map_entry n f a) [Heap.length a h - n.. Heap.length a h" + shows "Heap.length a h' = Heap.length a h" +using assms +proof (induct n arbitrary: h h' r) + case 0 + thus ?case by (auto elim!: crel_return) +next + case (Suc n) + let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h" + from Suc(3) have "[Heap.length a h - Suc n..n. map_entry n f a) [Heap.length a h - n..n. map_entry n f a) [Heap.length a ?h1 - n..n. map_entry n f a) [0..n. map_entry n f a) [Heap.length a h - Heap.length a h.. \ ref_present x h +extends h' h x \ ref_present x h' +extends h' h x \ ref_present y h \ ref_present y h' +extends h' h x \ ref_present y h \ get_ref y h = get_ref y h' +extends h' h x \ lim h' = Suc (lim h) +*) + +lemma crel_Ref_new: + assumes "crel (Ref.new v) h h' x" + obtains "get_ref x h' = v" + and "\ ref_present x h" + and "ref_present x h'" + and "\y. ref_present y h \ get_ref y h = get_ref y h'" + (* and "lim h' = Suc (lim h)" *) + and "\y. ref_present y h \ ref_present y h'" + using assms + unfolding Ref.new_def + apply (elim crel_heap) + unfolding Heap.ref_def + apply (simp add: Let_def) + unfolding Heap.new_ref_def + apply (simp add: Let_def) + unfolding ref_present_def + apply auto + unfolding get_ref_def set_ref_def + apply auto + done + +lemma crel_lookup: + assumes "crel (!ref) h h' r" + obtains "h = h'" "r = get_ref ref h" +using assms +unfolding Ref.lookup_def +by (auto elim: crel_heap) + +lemma crel_update: + assumes "crel (ref := v) h h' r" + obtains "h' = set_ref ref v h" "r = ()" +using assms +unfolding Ref.update_def +by (auto elim: crel_heap) + +lemma crel_change: + assumes "crel (Ref.change f ref) h h' r" + obtains "h' = set_ref ref (f (get_ref ref h)) h" "r = f (get_ref ref h)" +using assms +unfolding Ref.change_def run_drop Let_def +by (auto elim!: crelE crel_lookup crel_update crel_return) + +subsection {* Elimination rules for the assert command *} + +lemma crel_assert[consumes 1]: + assumes "crel (assert P x) h h' r" + obtains "P x" "r = x" "h = h'" + using assms + unfolding assert_def + by (elim crel_if crel_return crel_raise) auto + +lemma crel_assert_eq: "(\h h' r. crel f h h' r \ P r) \ f \= assert P = f" +unfolding crel_def bindM_def Let_def assert_def + raise_def return_def prod_case_beta +apply (cases f) +apply simp +apply (simp add: expand_fun_eq split_def) +apply auto +apply (case_tac "fst (fun x)") +apply (simp_all add: Pair_fst_snd_eq) +apply (erule_tac x="x" in meta_allE) +apply fastsimp +done + +section {* Introduction rules *} + +subsection {* Introduction rules for basic monadic commands *} + +lemma crelI: + assumes "crel f h h' r" "crel (g r) h' h'' r'" + shows "crel (f >>= g) h h'' r'" + using assms by (simp add: crel_def' bindM_def) + +lemma crelI': + assumes "crel f h h' r" "crel g h' h'' r'" + shows "crel (f >> g) h h'' r'" + using assms by (intro crelI) auto + +lemma crel_returnI: + shows "crel (return x) h h x" + unfolding crel_def return_def by simp + +lemma crel_raiseI: + shows "\ (crel (raise x) h h' r)" + unfolding crel_def raise_def by simp + +lemma crel_ifI: + assumes "c \ crel t h h' r" + "\c \ crel e h h' r" + shows "crel (if c then t else e) h h' r" + using assms + unfolding crel_def by auto + +lemma crel_option_caseI: + assumes "\y. x = Some y \ crel (s y) h h' r" + assumes "x = None \ crel n h h' r" + shows "crel (case x of None \ n | Some y \ s y) h h' r" +using assms +by (auto split: option.split) + +lemma crel_heapI: + shows "crel (Heap_Monad.heap f) h (snd (f h)) (fst (f h))" + by (simp add: crel_def apfst_def split_def prod_fun_def) + +lemma crel_heapI': + assumes "h' = snd (f h)" "r = fst (f h)" + shows "crel (Heap_Monad.heap f) h h' r" + using assms + by (simp add: crel_def split_def apfst_def prod_fun_def) + +lemma crelI2: + assumes "\h' rs'. crel f h h' rs' \ (\h'' rs. crel (g rs') h' h'' rs)" + shows "\h'' rs. crel (f\= g) h h'' rs" + oops + +lemma crel_ifI2: + assumes "c \ \h' r. crel t h h' r" + "\ c \ \h' r. crel e h h' r" + shows "\ h' r. crel (if c then t else e) h h' r" + oops + +subsection {* Introduction rules for array commands *} + +lemma crel_lengthI: + shows "crel (length a) h h (Heap.length a h)" + unfolding length_def + by (rule crel_heapI') auto + +(* thm crel_newI for Array.new is missing *) + +lemma crel_nthI: + assumes "i < Heap.length a h" + shows "crel (nth a i) h h ((get_array a h) ! i)" + using assms + unfolding nth_def run_drop + by (auto intro!: crelI crel_ifI crel_raiseI crel_lengthI crel_heapI') + +lemma crel_updI: + assumes "i < Heap.length a h" + shows "crel (upd i v a) h (Heap.upd a i v h) a" + using assms + unfolding upd_def run_drop + by (auto intro!: crelI crel_ifI crel_returnI crel_raiseI + crel_lengthI crel_heapI') + +(* thm crel_of_listI is missing *) + +(* thm crel_map_entryI is missing *) + +(* thm crel_swapI is missing *) + +(* thm crel_makeI is missing *) + +(* thm crel_freezeI is missing *) + +(* thm crel_mapI is missing *) + +subsection {* Introduction rules for reference commands *} + +lemma crel_lookupI: + shows "crel (!ref) h h (get_ref ref h)" + unfolding lookup_def by (auto intro!: crel_heapI') + +lemma crel_updateI: + shows "crel (ref := v) h (set_ref ref v h) ()" + unfolding update_def by (auto intro!: crel_heapI') + +lemma crel_changeI: + shows "crel (Ref.change f ref) h (set_ref ref (f (get_ref ref h)) h) (f (get_ref ref h))" +unfolding change_def Let_def run_drop by (auto intro!: crelI crel_returnI crel_lookupI crel_updateI) + +subsection {* Introduction rules for the assert command *} + +lemma crel_assertI: + assumes "P x" + shows "crel (assert P x) h h x" + using assms + unfolding assert_def + by (auto intro!: crel_ifI crel_returnI crel_raiseI) + +section {* Defintion of the noError predicate *} + +text {* We add a simple definitional setting for crel intro rules + where we only would like to show that the computation does not result in a exception for heap h, + but we do not care about statements about the resulting heap and return value.*} + +definition noError :: "'a Heap \ heap \ bool" +where + "noError c h \ (\r h'. (Inl r, h') = Heap_Monad.execute c h)" + +lemma noError_def': -- FIXME + "noError c h \ (\r h'. Heap_Monad.execute c h = (Inl r, h'))" + unfolding noError_def apply auto proof - + fix r h' + assume "(Inl r, h') = Heap_Monad.execute c h" + then have "Heap_Monad.execute c h = (Inl r, h')" .. + then show "\r h'. Heap_Monad.execute c h = (Inl r, h')" by blast +qed + +subsection {* Introduction rules for basic monadic commands *} + +lemma noErrorI: + assumes "noError f h" + assumes "\h' r. crel f h h' r \ noError (g r) h'" + shows "noError (f \= g) h" + using assms + by (auto simp add: noError_def' crel_def' bindM_def) + +lemma noErrorI': + assumes "noError f h" + assumes "\h' r. crel f h h' r \ noError g h'" + shows "noError (f \ g) h" + using assms + by (auto simp add: noError_def' crel_def' bindM_def) + +lemma noErrorI2: +"\crel f h h' r ; noError f h; noError (g r) h'\ +\ noError (f \= g) h" +by (auto simp add: noError_def' crel_def' bindM_def) + +lemma noError_return: + shows "noError (return x) h" + unfolding noError_def return_def + by auto + +lemma noError_if: + assumes "c \ noError t h" "\ c \ noError e h" + shows "noError (if c then t else e) h" + using assms + unfolding noError_def + by auto + +lemma noError_option_case: + assumes "\y. x = Some y \ noError (s y) h" + assumes "noError n h" + shows "noError (case x of None \ n | Some y \ s y) h" +using assms +by (auto split: option.split) + +lemma noError_mapM: +assumes "\x \ set xs. noError (f x) h \ crel (f x) h h (r x)" +shows "noError (mapM f xs) h" +using assms +proof (induct xs) + case Nil + thus ?case + unfolding mapM.simps by (intro noError_return) +next + case (Cons x xs) + thus ?case + unfolding mapM.simps run_drop + by (auto intro: noErrorI2[of "f x"] noErrorI noError_return) +qed + +lemma noError_heap: + shows "noError (Heap_Monad.heap f) h" + by (simp add: noError_def' apfst_def prod_fun_def split_def) + +subsection {* Introduction rules for array commands *} + +lemma noError_length: + shows "noError (Array.length a) h" + unfolding length_def + by (intro noError_heap) + +lemma noError_new: + shows "noError (Array.new n v) h" +unfolding Array.new_def by (intro noError_heap) + +lemma noError_upd: + assumes "i < Heap.length a h" + shows "noError (Array.upd i v a) h" + using assms + unfolding upd_def run_drop + by (auto intro!: noErrorI noError_if noError_return noError_length noError_heap) (auto elim: crel_length) + +lemma noError_nth: +assumes "i < Heap.length a h" + shows "noError (Array.nth a i) h" + using assms + unfolding nth_def run_drop + by (auto intro!: noErrorI noError_if noError_return noError_length noError_heap) (auto elim: crel_length) + +lemma noError_of_list: + shows "noError (of_list ls) h" + unfolding of_list_def by (rule noError_heap) + +lemma noError_map_entry: + assumes "i < Heap.length a h" + shows "noError (map_entry i f a) h" + using assms + unfolding map_entry_def run_drop + by (auto elim: crel_nth intro!: noErrorI noError_nth noError_upd) + +lemma noError_swap: + assumes "i < Heap.length a h" + shows "noError (swap i x a) h" + using assms + unfolding swap_def run_drop + by (auto elim: crel_nth intro!: noErrorI noError_return noError_nth noError_upd) + +lemma noError_make: + shows "noError (make n f) h" + unfolding make_def + by (auto intro: noError_of_list) + +(*TODO: move to HeapMonad *) +lemma mapM_append: + "mapM f (xs @ ys) = mapM f xs \= (\xs. mapM f ys \= (\ys. return (xs @ ys)))" + by (induct xs) (simp_all add: monad_simp) + +lemma noError_freeze: + shows "noError (freeze a) h" +unfolding freeze_def run_drop +by (auto intro!: noErrorI noError_length noError_mapM[of _ _ _ "\x. get_array a h ! x"] + noError_nth crel_nthI elim: crel_length) + +lemma noError_mapM_map_entry: + assumes "n \ Heap.length a h" + shows "noError (mapM (\n. map_entry n f a) [Heap.length a h - n.. nat \ ('a::heap) array \ heap \ 'a list" +where + "subarray n m a h \ sublist' n m (get_array a h)" + +lemma subarray_upd: "i \ m \ subarray n m a (Heap.upd a i v h) = subarray n m a h" +apply (simp add: subarray_def Heap.upd_def) +apply (simp add: sublist'_update1) +done + +lemma subarray_upd2: " i < n \ subarray n m a (Heap.upd a i v h) = subarray n m a h" +apply (simp add: subarray_def Heap.upd_def) +apply (subst sublist'_update2) +apply fastsimp +apply simp +done + +lemma subarray_upd3: "\ n \ i; i < m\ \ subarray n m a (Heap.upd a i v h) = subarray n m a h[i - n := v]" +unfolding subarray_def Heap.upd_def +by (simp add: sublist'_update3) + +lemma subarray_Nil: "n \ m \ subarray n m a h = []" +by (simp add: subarray_def sublist'_Nil') + +lemma subarray_single: "\ n < Heap.length a h \ \ subarray n (Suc n) a h = [get_array a h ! n]" +by (simp add: subarray_def Heap.length_def sublist'_single) + +lemma length_subarray: "m \ Heap.length a h \ List.length (subarray n m a h) = m - n" +by (simp add: subarray_def Heap.length_def length_sublist') + +lemma length_subarray_0: "m \ Heap.length a h \ List.length (subarray 0 m a h) = m" +by (simp add: length_subarray) + +lemma subarray_nth_array_Cons: "\ i < Heap.length a h; i < j \ \ (get_array a h ! i) # subarray (Suc i) j a h = subarray i j a h" +unfolding Heap.length_def subarray_def +by (simp add: sublist'_front) + +lemma subarray_nth_array_back: "\ i < j; j \ Heap.length a h\ \ subarray i j a h = subarray i (j - 1) a h @ [get_array a h ! (j - 1)]" +unfolding Heap.length_def subarray_def +by (simp add: sublist'_back) + +lemma subarray_append: "\ i < j; j < k \ \ subarray i j a h @ subarray j k a h = subarray i k a h" +unfolding subarray_def +by (simp add: sublist'_append) + +lemma subarray_all: "subarray 0 (Heap.length a h) a h = get_array a h" +unfolding Heap.length_def subarray_def +by (simp add: sublist'_all) + +lemma nth_subarray: "\ k < j - i; j \ Heap.length a h \ \ subarray i j a h ! k = get_array a h ! (i + k)" +unfolding Heap.length_def subarray_def +by (simp add: nth_sublist') + +lemma subarray_eq_samelength_iff: "Heap.length a h = Heap.length a h' \ (subarray i j a h = subarray i j a h') = (\i'. i \ i' \ i' < j \ get_array a h ! i' = get_array a h' ! i')" +unfolding Heap.length_def subarray_def by (rule sublist'_eq_samelength_iff) + +lemma all_in_set_subarray_conv: "(\j. j \ set (subarray l r a h) \ P j) = (\k. l \ k \ k < r \ k < Heap.length a h \ P (get_array a h ! k))" +unfolding subarray_def Heap.length_def by (rule all_in_set_sublist'_conv) + +lemma ball_in_set_subarray_conv: "(\j \ set (subarray l r a h). P j) = (\k. l \ k \ k < r \ k < Heap.length a h \ P (get_array a h ! k))" +unfolding subarray_def Heap.length_def by (rule ball_in_set_sublist'_conv) + +end \ No newline at end of file diff -r cf0c60e821bb -r d4f6e64ee7cc src/HOL/Library/Sublist.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Sublist.thy Sat Jul 19 19:27:13 2008 +0200 @@ -0,0 +1,507 @@ +(* $Id$ *) + +header {* Slices of lists *} + +theory Sublist +imports Multiset +begin + + +lemma sublist_split: "i \ j \ j \ k \ sublist xs {i.. j - 1 \ j - 1 \ k - 1") +apply simp +apply (subgoal_tac "{ja. Suc ja < j} = {0.. Suc ja \ Suc ja < k} = {j - Suc 0.. Suc ja \ Suc ja < j} = {i - 1 .. Suc ja \ Suc ja < k} = {j - 1.. Suc j \ Suc j < k} = {i - 1.. j - 1 \ j - 1 \ k - 1") +apply simp +apply fastsimp +apply fastsimp +apply fastsimp +apply fastsimp +done + +lemma sublist_update1: "i \ inds \ sublist (xs[i := v]) inds = sublist xs inds" +apply (induct xs arbitrary: i inds) +apply simp +apply (case_tac i) +apply (simp add: sublist_Cons) +apply (simp add: sublist_Cons) +done + +lemma sublist_update2: "i \ inds \ sublist (xs[i := v]) inds = (sublist xs inds)[(card {k \ inds. k < i}):= v]" +proof (induct xs arbitrary: i inds) + case Nil thus ?case by simp +next + case (Cons x xs) + thus ?case + proof (cases i) + case 0 with Cons show ?thesis by (simp add: sublist_Cons) + next + case (Suc i') + with Cons show ?thesis + apply simp + apply (simp add: sublist_Cons) + apply auto + apply (auto simp add: nat.split) + apply (simp add: card_less) + apply (simp add: card_less) + apply (simp add: card_less_Suc[symmetric]) + apply (simp add: card_less_Suc2) + done + qed +qed + +lemma sublist_update: "sublist (xs[i := v]) inds = (if i \ inds then (sublist xs inds)[(card {k \ inds. k < i}) := v] else sublist xs inds)" +by (simp add: sublist_update1 sublist_update2) + +lemma sublist_take: "sublist xs {j. j < m} = take m xs" +apply (induct xs arbitrary: m) +apply simp +apply (case_tac m) +apply simp +apply (simp add: sublist_Cons) +done + +lemma sublist_take': "sublist xs {0.. sublist xs {a} = [xs ! a]" +apply (induct xs arbitrary: a) +apply simp +apply(case_tac aa) +apply simp +apply (simp add: sublist_Cons) +apply simp +apply (simp add: sublist_Cons) +done + +lemma sublist_is_Nil: "\i \ inds. i \ length xs \ sublist xs inds = []" +apply (induct xs arbitrary: inds) +apply simp +apply (simp add: sublist_Cons) +apply auto +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply auto +done + +lemma sublist_Nil': "sublist xs inds = [] \ \i \ inds. i \ length xs" +apply (induct xs arbitrary: inds) +apply simp +apply (simp add: sublist_Cons) +apply (auto split: if_splits) +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply (case_tac x, auto) +done + +lemma sublist_Nil[simp]: "(sublist xs inds = []) = (\i \ inds. i \ length xs)" +apply (induct xs arbitrary: inds) +apply simp +apply (simp add: sublist_Cons) +apply auto +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply (case_tac x, auto) +done + +lemma sublist_eq_subseteq: " \ inds' \ inds; sublist xs inds = sublist ys inds \ \ sublist xs inds' = sublist ys inds'" +apply (induct xs arbitrary: ys inds inds') +apply simp +apply (drule sym, rule sym) +apply (simp add: sublist_Nil, fastsimp) +apply (case_tac ys) +apply (simp add: sublist_Nil, fastsimp) +apply (auto simp add: sublist_Cons) +apply (erule_tac x="list" in meta_allE) +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply (erule_tac x="{j. Suc j \ inds'}" in meta_allE) +apply fastsimp +apply (erule_tac x="list" in meta_allE) +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply (erule_tac x="{j. Suc j \ inds'}" in meta_allE) +apply fastsimp +done + +lemma sublist_eq: "\ \i \ inds. ((i < length xs) \ (i < length ys)) \ ((i \ length xs ) \ (i \ length ys)); \i \ inds. xs ! i = ys ! i \ \ sublist xs inds = sublist ys inds" +apply (induct xs arbitrary: ys inds) +apply simp +apply (rule sym, simp add: sublist_Nil) +apply (case_tac ys) +apply (simp add: sublist_Nil) +apply (auto simp add: sublist_Cons) +apply (erule_tac x="list" in meta_allE) +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply fastsimp +apply (erule_tac x="list" in meta_allE) +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply fastsimp +done + +lemma sublist_eq_samelength: "\ length xs = length ys; \i \ inds. xs ! i = ys ! i \ \ sublist xs inds = sublist ys inds" +by (rule sublist_eq, auto) + +lemma sublist_eq_samelength_iff: "length xs = length ys \ (sublist xs inds = sublist ys inds) = (\i \ inds. xs ! i = ys ! i)" +apply (induct xs arbitrary: ys inds) +apply simp +apply (rule sym, simp add: sublist_Nil) +apply (case_tac ys) +apply (simp add: sublist_Nil) +apply (auto simp add: sublist_Cons) +apply (case_tac i) +apply auto +apply (case_tac i) +apply auto +done + +section {* Another sublist function *} + +function sublist' :: "nat \ nat \ 'a list \ 'a list" +where + "sublist' n m [] = []" +| "sublist' n 0 xs = []" +| "sublist' 0 (Suc m) (x#xs) = (x#sublist' 0 m xs)" +| "sublist' (Suc n) (Suc m) (x#xs) = sublist' n m xs" +by pat_completeness auto +termination by lexicographic_order + +subsection {* Proving equivalence to the other sublist command *} + +lemma sublist'_sublist: "sublist' n m xs = sublist xs {j. n \ j \ j < m}" +apply (induct xs arbitrary: n m) +apply simp +apply (case_tac n) +apply (case_tac m) +apply simp +apply (simp add: sublist_Cons) +apply (case_tac m) +apply simp +apply (simp add: sublist_Cons) +done + + +lemma "sublist' n m xs = sublist xs {n.. (x # sublist' 0 j xs) | Suc i' \ sublist' i' j xs)" +by (cases i) auto + +lemma sublist'_Cons2[simp]: "sublist' i j (x#xs) = (if (j = 0) then [] else ((if (i = 0) then [x] else []) @ sublist' (i - 1) (j - 1) xs))" +apply (cases j) +apply auto +apply (cases i) +apply auto +done + +lemma sublist_n_0: "sublist' n 0 xs = []" +by (induct xs, auto) + +lemma sublist'_Nil': "n \ m \ sublist' n m xs = []" +apply (induct xs arbitrary: n m) +apply simp +apply (case_tac m) +apply simp +apply (case_tac n) +apply simp +apply simp +done + +lemma sublist'_Nil2: "n \ length xs \ sublist' n m xs = []" +apply (induct xs arbitrary: n m) +apply simp +apply (case_tac m) +apply simp +apply (case_tac n) +apply simp +apply simp +done + +lemma sublist'_Nil3: "(sublist' n m xs = []) = ((n \ m) \ (n \ length xs))" +apply (induct xs arbitrary: n m) +apply simp +apply (case_tac m) +apply simp +apply (case_tac n) +apply simp +apply simp +done + +lemma sublist'_notNil: "\ n < length xs; n < m \ \ sublist' n m xs \ []" +apply (induct xs arbitrary: n m) +apply simp +apply (case_tac m) +apply simp +apply (case_tac n) +apply simp +apply simp +done + +lemma sublist'_single: "n < length xs \ sublist' n (Suc n) xs = [xs ! n]" +apply (induct xs arbitrary: n) +apply simp +apply simp +apply (case_tac n) +apply (simp add: sublist_n_0) +apply simp +done + +lemma sublist'_update1: "i \ m \ sublist' n m (xs[i:=v]) = sublist' n m xs" +apply (induct xs arbitrary: n m i) +apply simp +apply simp +apply (case_tac i) +apply simp +apply simp +done + +lemma sublist'_update2: "i < n \ sublist' n m (xs[i:=v]) = sublist' n m xs" +apply (induct xs arbitrary: n m i) +apply simp +apply simp +apply (case_tac i) +apply simp +apply simp +done + +lemma sublist'_update3: "\n \ i; i < m\ \ sublist' n m (xs[i := v]) = (sublist' n m xs)[i - n := v]" +proof (induct xs arbitrary: n m i) + case Nil thus ?case by auto +next + case (Cons x xs) + thus ?case + apply - + apply auto + apply (cases i) + apply auto + apply (cases i) + apply auto + done +qed + +lemma "\ sublist' i j xs = sublist' i j ys; n \ i; m \ j \ \ sublist' n m xs = sublist' n m ys" +proof (induct xs arbitrary: i j ys n m) + case Nil + thus ?case + apply - + apply (rule sym, drule sym) + apply (simp add: sublist'_Nil) + apply (simp add: sublist'_Nil3) + apply arith + done +next + case (Cons x xs i j ys n m) + note c = this + thus ?case + proof (cases m) + case 0 thus ?thesis by (simp add: sublist_n_0) + next + case (Suc m') + note a = this + thus ?thesis + proof (cases n) + case 0 note b = this + show ?thesis + proof (cases ys) + case Nil with a b Cons.prems show ?thesis by (simp add: sublist'_Nil3) + next + case (Cons y ys) + show ?thesis + proof (cases j) + case 0 with a b Cons.prems show ?thesis by simp + next + case (Suc j') with a b Cons.prems Cons show ?thesis + apply - + apply (simp, rule Cons.hyps [of "0" "j'" "ys" "0" "m'"], auto) + done + qed + qed + next + case (Suc n') + show ?thesis + proof (cases ys) + case Nil with Suc a Cons.prems show ?thesis by (auto simp add: sublist'_Nil3) + next + case (Cons y ys) with Suc a Cons.prems show ?thesis + apply - + apply simp + apply (cases j) + apply simp + apply (cases i) + apply simp + apply (rule_tac j="nat" in Cons.hyps [of "0" _ "ys" "n'" "m'"]) + apply simp + apply simp + apply simp + apply simp + apply (rule_tac i="nata" and j="nat" in Cons.hyps [of _ _ "ys" "n'" "m'"]) + apply simp + apply simp + apply simp + done + qed + qed + qed +qed + +lemma length_sublist': "j \ length xs \ length (sublist' i j xs) = j - i" +by (induct xs arbitrary: i j, auto) + +lemma sublist'_front: "\ i < j; i < length xs \ \ sublist' i j xs = xs ! i # sublist' (Suc i) j xs" +apply (induct xs arbitrary: a i j) +apply simp +apply (case_tac j) +apply simp +apply (case_tac i) +apply simp +apply simp +done + +lemma sublist'_back: "\ i < j; j \ length xs \ \ sublist' i j xs = sublist' i (j - 1) xs @ [xs ! (j - 1)]" +apply (induct xs arbitrary: a i j) +apply simp +apply simp +apply (case_tac j) +apply simp +apply auto +apply (case_tac nat) +apply auto +done + +(* suffices that j \ length xs and length ys *) +lemma sublist'_eq_samelength_iff: "length xs = length ys \ (sublist' i j xs = sublist' i j ys) = (\i'. i \ i' \ i' < j \ xs ! i' = ys ! i')" +proof (induct xs arbitrary: ys i j) + case Nil thus ?case by simp +next + case (Cons x xs) + thus ?case + apply - + apply (cases ys) + apply simp + apply simp + apply auto + apply (case_tac i', auto) + apply (erule_tac x="Suc i'" in allE, auto) + apply (erule_tac x="i' - 1" in allE, auto) + apply (case_tac i', auto) + apply (erule_tac x="Suc i'" in allE, auto) + done +qed + +lemma sublist'_all[simp]: "sublist' 0 (length xs) xs = xs" +by (induct xs, auto) + +lemma sublist'_sublist': "sublist' n m (sublist' i j xs) = sublist' (i + n) (min (i + m) j) xs" +by (induct xs arbitrary: i j n m) (auto simp add: min_diff) + +lemma sublist'_append: "\ i \ j; j \ k \ \(sublist' i j xs) @ (sublist' j k xs) = sublist' i k xs" +by (induct xs arbitrary: i j k) auto + +lemma nth_sublist': "\ k < j - i; j \ length xs \ \ (sublist' i j xs) ! k = xs ! (i + k)" +apply (induct xs arbitrary: i j k) +apply auto +apply (case_tac k) +apply auto +apply (case_tac i) +apply auto +done + +lemma set_sublist': "set (sublist' i j xs) = {x. \k. i \ k \ k < j \ k < List.length xs \ x = xs ! k}" +apply (simp add: sublist'_sublist) +apply (simp add: set_sublist) +apply auto +done + +lemma all_in_set_sublist'_conv: "(\j. j \ set (sublist' l r xs) \ P j) = (\k. l \ k \ k < r \ k < List.length xs \ P (xs ! k))" +unfolding set_sublist' by blast + +lemma ball_in_set_sublist'_conv: "(\j \ set (sublist' l r xs). P j) = (\k. l \ k \ k < r \ k < List.length xs \ P (xs ! k))" +unfolding set_sublist' by blast + + +lemma multiset_of_sublist: +assumes l_r: "l \ r \ r \ List.length xs" +assumes left: "\ i. i < l \ (xs::'a list) ! i = ys ! i" +assumes right: "\ i. i \ r \ (xs::'a list) ! i = ys ! i" +assumes multiset: "multiset_of xs = multiset_of ys" + shows "multiset_of (sublist' l r xs) = multiset_of (sublist' l r ys)" +proof - + from l_r have xs_def: "xs = (sublist' 0 l xs) @ (sublist' l r xs) @ (sublist' r (List.length xs) xs)" (is "_ = ?xs_long") + by (simp add: sublist'_append) + from multiset have length_eq: "List.length xs = List.length ys" by (rule multiset_of_eq_length) + with l_r have ys_def: "ys = (sublist' 0 l ys) @ (sublist' l r ys) @ (sublist' r (List.length ys) ys)" (is "_ = ?ys_long") + by (simp add: sublist'_append) + from xs_def ys_def multiset have "multiset_of ?xs_long = multiset_of ?ys_long" by simp + moreover + from left l_r length_eq have "sublist' 0 l xs = sublist' 0 l ys" + by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI) + moreover + from right l_r length_eq have "sublist' r (List.length xs) xs = sublist' r (List.length ys) ys" + by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI) + moreover + ultimately show ?thesis by (simp add: multiset_of_append) +qed + + +end diff -r cf0c60e821bb -r d4f6e64ee7cc src/HOL/SetInterval.thy --- a/src/HOL/SetInterval.thy Sat Jul 19 11:05:18 2008 +0200 +++ b/src/HOL/SetInterval.thy Sat Jul 19 19:27:13 2008 +0200 @@ -558,6 +558,54 @@ lemma card_greaterThanLessThan_int [simp]: "card {l<.. k < (i::nat)}" +proof - + have "{k. P k \ k < i} \ {.. M" +shows "card {k \ M. k < Suc i} \ 0" +proof - + from zero_in_M have "{k \ M. k < Suc i} \ {}" by auto + with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff) +qed + +lemma card_less_Suc2: "0 \ M \ card {k. Suc k \ M \ k < i} = card {k \ M. k < Suc i}" +apply (rule card_bij_eq [of "Suc" _ _ "\x. x - 1"]) +apply simp +apply fastsimp +apply auto +apply (rule inj_on_diff_nat) +apply auto +apply (case_tac x) +apply auto +apply (case_tac xa) +apply auto +apply (case_tac xa) +apply auto +apply (auto simp add: finite_M_bounded_by_nat) +done + +lemma card_less_Suc: + assumes zero_in_M: "0 \ M" + shows "Suc (card {k. Suc k \ M \ k < i}) = card {k \ M. k < Suc i}" +proof - + from assms have a: "0 \ {k \ M. k < Suc i}" by simp + hence c: "{k \ M. k < Suc i} = insert 0 ({k \ M. k < Suc i} - {0})" + by (auto simp only: insert_Diff) + have b: "{k \ M. k < Suc i} - {0} = {k \ M - {0}. k < Suc i}" by auto + from finite_M_bounded_by_nat[of "\x. x \ M" "Suc i"] have "Suc (card {k. Suc k \ M \ k < i}) = card (insert 0 ({k \ M. k < Suc i} - {0}))" + apply (subst card_insert) + apply simp_all + apply (subst b) + apply (subst card_less_Suc2[symmetric]) + apply simp_all + done + with c show ?thesis by simp +qed + subsection {*Lemmas useful with the summation operator setsum*} diff -r cf0c60e821bb -r d4f6e64ee7cc src/HOL/ex/ImperativeQuicksort.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/ImperativeQuicksort.thy Sat Jul 19 19:27:13 2008 +0200 @@ -0,0 +1,622 @@ +theory ImperativeQuickSort +imports Imperative_HOL Subarray Multiset +begin + +text {* We prove QuickSort correct in the Relational Calculus. *} + + +definition swap :: "nat array \ nat \ nat \ unit Heap" +where + "swap arr i j = ( + do + x \ nth arr i; + y \ nth arr j; + upd i y arr; + upd j x arr; + return () + done)" + +lemma swap_permutes: + assumes "crel (swap a i j) h h' rs" + shows "multiset_of (get_array a h') + = multiset_of (get_array a h)" + using assms + unfolding swap_def run_drop + by (auto simp add: Heap.length_def multiset_of_swap dest: sym [of _ "h'"] elim!: crelE crel_nth crel_return crel_upd) + +function part1 :: "nat array \ nat \ nat \ nat \ nat Heap" +where + "part1 a left right p = ( + if (right \ left) then return right + else (do + v \ nth a left; + (if (v \ p) then (part1 a (left + 1) right p) + else (do swap a left right; + part1 a left (right - 1) p done)) + done))" +by pat_completeness auto + +termination +by (relation "measure (\(_,l,r,_). r - l )") auto + +declare part1.simps[simp del] + +lemma part_permutes: + assumes "crel (part1 a l r p) h h' rs" + shows "multiset_of (get_array a h') + = multiset_of (get_array a h)" + using assms +proof (induct a l r p arbitrary: h h' rs rule:part1.induct) + case (1 a l r p h h' rs) + thus ?case + unfolding part1.simps [of a l r p] run_drop + by (elim crelE crel_if crel_return crel_nth) (auto simp add: swap_permutes) +qed + +lemma part_returns_index_in_bounds: + assumes "crel (part1 a l r p) h h' rs" + assumes "l \ r" + shows "l \ rs \ rs \ r" +using assms +proof (induct a l r p arbitrary: h h' rs rule:part1.induct) + case (1 a l r p h h' rs) + note cr = `crel (part1 a l r p) h h' rs` + show ?case + proof (cases "r \ l") + case True (* Terminating case *) + with cr `l \ r` show ?thesis + unfolding part1.simps[of a l r p] run_drop + by (elim crelE crel_if crel_return crel_nth) auto + next + case False (* recursive case *) + note rec_condition = this + let ?v = "get_array a h ! l" + show ?thesis + proof (cases "?v \ p") + case True + with cr False + have rec1: "crel (part1 a (l + 1) r p) h h' rs" + unfolding part1.simps[of a l r p] run_drop + by (elim crelE crel_nth crel_if crel_return) auto + from rec_condition have "l + 1 \ r" by arith + from 1(1)[OF rec_condition True rec1 `l + 1 \ r`] + show ?thesis by simp + next + case False + with rec_condition cr + obtain h1 where swp: "crel (swap a l r) h h1 ()" + and rec2: "crel (part1 a l (r - 1) p) h1 h' rs" + unfolding part1.simps[of a l r p] run_drop + by (elim crelE crel_nth crel_if crel_return) auto + from rec_condition have "l \ r - 1" by arith + from 1(2) [OF rec_condition False rec2 `l \ r - 1`] show ?thesis by fastsimp + qed + qed +qed + +lemma part_length_remains: + assumes "crel (part1 a l r p) h h' rs" + shows "Heap.length a h = Heap.length a h'" +using assms +proof (induct a l r p arbitrary: h h' rs rule:part1.induct) + case (1 a l r p h h' rs) + note cr = `crel (part1 a l r p) h h' rs` + + show ?case + proof (cases "r \ l") + case True (* Terminating case *) + with cr show ?thesis + unfolding part1.simps[of a l r p] run_drop + by (elim crelE crel_if crel_return crel_nth) auto + next + case False (* recursive case *) + with cr 1 show ?thesis + unfolding part1.simps [of a l r p] swap_def run_drop + by (auto elim!: crelE crel_if crel_nth crel_return crel_upd) fastsimp + qed +qed + +lemma part_outer_remains: + assumes "crel (part1 a l r p) h h' rs" + shows "\i. i < l \ r < i \ get_array (a::nat array) h ! i = get_array a h' ! i" + using assms +proof (induct a l r p arbitrary: h h' rs rule:part1.induct) + case (1 a l r p h h' rs) + note cr = `crel (part1 a l r p) h h' rs` + + show ?case + proof (cases "r \ l") + case True (* Terminating case *) + with cr show ?thesis + unfolding part1.simps[of a l r p] run_drop + by (elim crelE crel_if crel_return crel_nth) auto + next + case False (* recursive case *) + note rec_condition = this + let ?v = "get_array a h ! l" + show ?thesis + proof (cases "?v \ p") + case True + with cr False + have rec1: "crel (part1 a (l + 1) r p) h h' rs" + unfolding part1.simps[of a l r p] run_drop + by (elim crelE crel_nth crel_if crel_return) auto + from 1(1)[OF rec_condition True rec1] + show ?thesis by fastsimp + next + case False + with rec_condition cr + obtain h1 where swp: "crel (swap a l r) h h1 ()" + and rec2: "crel (part1 a l (r - 1) p) h1 h' rs" + unfolding part1.simps[of a l r p] run_drop + by (elim crelE crel_nth crel_if crel_return) auto + from swp rec_condition have + "\i. i < l \ r < i \ get_array a h ! i = get_array a h1 ! i" + unfolding swap_def run_drop + by (elim crelE crel_nth crel_upd crel_return) auto + with 1(2) [OF rec_condition False rec2] show ?thesis by fastsimp + qed + qed +qed + + +lemma part_partitions: + assumes "crel (part1 a l r p) h h' rs" + shows "(\i. l \ i \ i < rs \ get_array (a::nat array) h' ! i \ p) + \ (\i. rs < i \ i \ r \ get_array a h' ! i \ p)" + using assms +proof (induct a l r p arbitrary: h h' rs rule:part1.induct) + case (1 a l r p h h' rs) + note cr = `crel (part1 a l r p) h h' rs` + + show ?case + proof (cases "r \ l") + case True (* Terminating case *) + with cr have "rs = r" + unfolding part1.simps[of a l r p] run_drop + by (elim crelE crel_if crel_return crel_nth) auto + with True + show ?thesis by auto + next + case False (* recursive case *) + note lr = this + let ?v = "get_array a h ! l" + show ?thesis + proof (cases "?v \ p") + case True + with lr cr + have rec1: "crel (part1 a (l + 1) r p) h h' rs" + unfolding part1.simps[of a l r p] run_drop + by (elim crelE crel_nth crel_if crel_return) auto + from True part_outer_remains[OF rec1] have a_l: "get_array a h' ! l \ p" + by fastsimp + have "\i. (l \ i = (l = i \ Suc l \ i))" by arith + with 1(1)[OF False True rec1] a_l show ?thesis + by auto + next + case False + with lr cr + obtain h1 where swp: "crel (swap a l r) h h1 ()" + and rec2: "crel (part1 a l (r - 1) p) h1 h' rs" + unfolding part1.simps[of a l r p] run_drop + by (elim crelE crel_nth crel_if crel_return) auto + from swp False have "get_array a h1 ! r \ p" + unfolding swap_def run_drop + by (auto simp add: Heap.length_def elim!: crelE crel_nth crel_upd crel_return) + with part_outer_remains [OF rec2] lr have a_r: "get_array a h' ! r \ p" + by fastsimp + have "\i. (i \ r = (i = r \ i \ r - 1))" by arith + with 1(2)[OF lr False rec2] a_r show ?thesis + by auto + qed + qed +qed + + +fun partition :: "nat array \ nat \ nat \ nat Heap" +where + "partition a left right = (do + pivot \ nth a right; + middle \ part1 a left (right - 1) pivot; + v \ nth a middle; + m \ return (if (v \ pivot) then (middle + 1) else middle); + swap a m right; + return m + done)" + +declare partition.simps[simp del] + +lemma partition_permutes: + assumes "crel (partition a l r) h h' rs" + shows "multiset_of (get_array a h') + = multiset_of (get_array a h)" +proof - + from assms part_permutes swap_permutes show ?thesis + unfolding partition.simps run_drop + by (elim crelE crel_return crel_nth crel_if crel_upd) auto +qed + +lemma partition_length_remains: + assumes "crel (partition a l r) h h' rs" + shows "Heap.length a h = Heap.length a h'" +proof - + from assms part_length_remains show ?thesis + unfolding partition.simps run_drop swap_def + by (elim crelE crel_return crel_nth crel_if crel_upd) auto +qed + +lemma partition_outer_remains: + assumes "crel (partition a l r) h h' rs" + assumes "l < r" + shows "\i. i < l \ r < i \ get_array (a::nat array) h ! i = get_array a h' ! i" +proof - + from assms part_outer_remains part_returns_index_in_bounds show ?thesis + unfolding partition.simps swap_def run_drop + by (elim crelE crel_return crel_nth crel_if crel_upd) fastsimp +qed + +lemma partition_returns_index_in_bounds: + assumes crel: "crel (partition a l r) h h' rs" + assumes "l < r" + shows "l \ rs \ rs \ r" +proof - + from crel obtain middle h'' p where part: "crel (part1 a l (r - 1) p) h h'' middle" + and rs_equals: "rs = (if get_array a h'' ! middle \ get_array a h ! r then middle + 1 + else middle)" + unfolding partition.simps run_drop + by (elim crelE crel_return crel_nth crel_if crel_upd) simp + from `l < r` have "l \ r - 1" by arith + from part_returns_index_in_bounds[OF part this] rs_equals `l < r` show ?thesis by auto +qed + +lemma partition_partitions: + assumes crel: "crel (partition a l r) h h' rs" + assumes "l < r" + shows "(\i. l \ i \ i < rs \ get_array (a::nat array) h' ! i \ get_array a h' ! rs) \ + (\i. rs < i \ i \ r \ get_array a h' ! rs \ get_array a h' ! i)" +proof - + let ?pivot = "get_array a h ! r" + from crel obtain middle h1 where part: "crel (part1 a l (r - 1) ?pivot) h h1 middle" + and swap: "crel (swap a rs r) h1 h' ()" + and rs_equals: "rs = (if get_array a h1 ! middle \ ?pivot then middle + 1 + else middle)" + unfolding partition.simps run_drop + by (elim crelE crel_return crel_nth crel_if crel_upd) simp + from swap have h'_def: "h' = Heap.upd a r (get_array a h1 ! rs) + (Heap.upd a rs (get_array a h1 ! r) h1)" + unfolding swap_def run_drop + by (elim crelE crel_return crel_nth crel_upd) simp + from swap have in_bounds: "r < Heap.length a h1 \ rs < Heap.length a h1" + unfolding swap_def run_drop + by (elim crelE crel_return crel_nth crel_upd) simp + from swap have swap_length_remains: "Heap.length a h1 = Heap.length a h'" + unfolding swap_def run_drop by (elim crelE crel_return crel_nth crel_upd) auto + from `l < r` have "l \ r - 1" by simp + note middle_in_bounds = part_returns_index_in_bounds[OF part this] + from part_outer_remains[OF part] `l < r` + have "get_array a h ! r = get_array a h1 ! r" + by fastsimp + with swap + have right_remains: "get_array a h ! r = get_array a h' ! rs" + unfolding swap_def run_drop + by (auto simp add: Heap.length_def elim!: crelE crel_return crel_nth crel_upd) (cases "r = rs", auto) + from part_partitions [OF part] + show ?thesis + proof (cases "get_array a h1 ! middle \ ?pivot") + case True + with rs_equals have rs_equals: "rs = middle + 1" by simp + { + fix i + assume i_is_left: "l \ i \ i < rs" + with swap_length_remains in_bounds middle_in_bounds rs_equals `l < r` + have i_props: "i < Heap.length a h'" "i \ r" "i \ rs" by auto + from i_is_left rs_equals have "l \ i \ i < middle \ i = middle" by arith + with part_partitions[OF part] right_remains True + have "get_array a h1 ! i \ get_array a h' ! rs" by fastsimp + with i_props h'_def in_bounds have "get_array a h' ! i \ get_array a h' ! rs" + unfolding Heap.upd_def Heap.length_def by simp + } + moreover + { + fix i + assume "rs < i \ i \ r" + + hence "(rs < i \ i \ r - 1) \ (rs < i \ i = r)" by arith + hence "get_array a h' ! rs \ get_array a h' ! i" + proof + assume i_is: "rs < i \ i \ r - 1" + with swap_length_remains in_bounds middle_in_bounds rs_equals + have i_props: "i < Heap.length a h'" "i \ r" "i \ rs" by auto + from part_partitions[OF part] rs_equals right_remains i_is + have "get_array a h' ! rs \ get_array a h1 ! i" + by fastsimp + with i_props h'_def show ?thesis by fastsimp + next + assume i_is: "rs < i \ i = r" + with rs_equals have "Suc middle \ r" by arith + with middle_in_bounds `l < r` have "Suc middle \ r - 1" by arith + with part_partitions[OF part] right_remains + have "get_array a h' ! rs \ get_array a h1 ! (Suc middle)" + by fastsimp + with i_is True rs_equals right_remains h'_def + show ?thesis using in_bounds + unfolding Heap.upd_def Heap.length_def + by auto + qed + } + ultimately show ?thesis by auto + next + case False + with rs_equals have rs_equals: "middle = rs" by simp + { + fix i + assume i_is_left: "l \ i \ i < rs" + with swap_length_remains in_bounds middle_in_bounds rs_equals + have i_props: "i < Heap.length a h'" "i \ r" "i \ rs" by auto + from part_partitions[OF part] rs_equals right_remains i_is_left + have "get_array a h1 ! i \ get_array a h' ! rs" by fastsimp + with i_props h'_def have "get_array a h' ! i \ get_array a h' ! rs" + unfolding Heap.upd_def by simp + } + moreover + { + fix i + assume "rs < i \ i \ r" + hence "(rs < i \ i \ r - 1) \ i = r" by arith + hence "get_array a h' ! rs \ get_array a h' ! i" + proof + assume i_is: "rs < i \ i \ r - 1" + with swap_length_remains in_bounds middle_in_bounds rs_equals + have i_props: "i < Heap.length a h'" "i \ r" "i \ rs" by auto + from part_partitions[OF part] rs_equals right_remains i_is + have "get_array a h' ! rs \ get_array a h1 ! i" + by fastsimp + with i_props h'_def show ?thesis by fastsimp + next + assume i_is: "i = r" + from i_is False rs_equals right_remains h'_def + show ?thesis using in_bounds + unfolding Heap.upd_def Heap.length_def + by auto + qed + } + ultimately + show ?thesis by auto + qed +qed + + +function quicksort :: "nat array \ nat \ nat \ unit Heap" +where + "quicksort arr left right = + (if (right > left) then + do + pivotNewIndex \ partition arr left right; + pivotNewIndex \ assert (\x. left \ x \ x \ right) pivotNewIndex; + quicksort arr left (pivotNewIndex - 1); + quicksort arr (pivotNewIndex + 1) right + done + else return ())" +by pat_completeness auto + +(* For termination, we must show that the pivotNewIndex is between left and right *) +termination +by (relation "measure (\(a, l, r). (r - l))") auto + +declare quicksort.simps[simp del] + + +lemma quicksort_permutes: + assumes "crel (quicksort a l r) h h' rs" + shows "multiset_of (get_array a h') + = multiset_of (get_array a h)" + using assms +proof (induct a l r arbitrary: h h' rs rule: quicksort.induct) + case (1 a l r h h' rs) + with partition_permutes show ?case + unfolding quicksort.simps [of a l r] run_drop + by (elim crel_if crelE crel_assert crel_return) auto +qed + +lemma length_remains: + assumes "crel (quicksort a l r) h h' rs" + shows "Heap.length a h = Heap.length a h'" +using assms +proof (induct a l r arbitrary: h h' rs rule: quicksort.induct) + case (1 a l r h h' rs) + with partition_length_remains show ?case + unfolding quicksort.simps [of a l r] run_drop + by (elim crel_if crelE crel_assert crel_return) auto +qed + +lemma quicksort_outer_remains: + assumes "crel (quicksort a l r) h h' rs" + shows "\i. i < l \ r < i \ get_array (a::nat array) h ! i = get_array a h' ! i" + using assms +proof (induct a l r arbitrary: h h' rs rule: quicksort.induct) + case (1 a l r h h' rs) + note cr = `crel (quicksort a l r) h h' rs` + thus ?case + proof (cases "r > l") + case False + with cr have "h' = h" + unfolding quicksort.simps [of a l r] + by (elim crel_if crel_return) auto + thus ?thesis by simp + next + case True + { + fix h1 h2 p ret1 ret2 i + assume part: "crel (partition a l r) h h1 p" + assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ret1" + assume qs2: "crel (quicksort a (p + 1) r) h2 h' ret2" + assume pivot: "l \ p \ p \ r" + assume i_outer: "i < l \ r < i" + from partition_outer_remains [OF part True] i_outer + have "get_array a h !i = get_array a h1 ! i" by fastsimp + moreover + with 1(1) [OF True pivot qs1] pivot i_outer + have "get_array a h1 ! i = get_array a h2 ! i" by auto + moreover + with qs2 1(2) [of p h2 h' ret2] True pivot i_outer + have "get_array a h2 ! i = get_array a h' ! i" by auto + ultimately have "get_array a h ! i= get_array a h' ! i" by simp + } + with cr show ?thesis + unfolding quicksort.simps [of a l r] run_drop + by (elim crel_if crelE crel_assert crel_return) auto + qed +qed + +lemma quicksort_is_skip: + assumes "crel (quicksort a l r) h h' rs" + shows "r \ l \ h = h'" + using assms + unfolding quicksort.simps [of a l r] run_drop + by (elim crel_if crel_return) auto + +lemma quicksort_sorts: + assumes "crel (quicksort a l r) h h' rs" + assumes l_r_length: "l < Heap.length a h" "r < Heap.length a h" + shows "sorted (subarray l (r + 1) a h')" + using assms +proof (induct a l r arbitrary: h h' rs rule: quicksort.induct) + case (1 a l r h h' rs) + note cr = `crel (quicksort a l r) h h' rs` + thus ?case + proof (cases "r > l") + case False + hence "l \ r + 1 \ l = r" by arith + with length_remains[OF cr] 1(5) show ?thesis + by (auto simp add: subarray_Nil subarray_single) + next + case True + { + fix h1 h2 p + assume part: "crel (partition a l r) h h1 p" + assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ()" + assume qs2: "crel (quicksort a (p + 1) r) h2 h' ()" + from partition_returns_index_in_bounds [OF part True] + have pivot: "l\ p \ p \ r" . + note length_remains = length_remains[OF qs2] length_remains[OF qs1] partition_length_remains[OF part] + from quicksort_outer_remains [OF qs2] quicksort_outer_remains [OF qs1] pivot quicksort_is_skip[OF qs1] + have pivot_unchanged: "get_array a h1 ! p = get_array a h' ! p" by (cases p, auto) + (*-- First of all, by induction hypothesis both sublists are sorted. *) + from 1(1)[OF True pivot qs1] length_remains pivot 1(5) + have IH1: "sorted (subarray l p a h2)" by (cases p, auto simp add: subarray_Nil) + from quicksort_outer_remains [OF qs2] length_remains + have left_subarray_remains: "subarray l p a h2 = subarray l p a h'" + by (simp add: subarray_eq_samelength_iff) + with IH1 have IH1': "sorted (subarray l p a h')" by simp + from 1(2)[OF True pivot qs2] pivot 1(5) length_remains + have IH2: "sorted (subarray (p + 1) (r + 1) a h')" + by (cases "Suc p \ r", auto simp add: subarray_Nil) + (* -- Secondly, both sublists remain partitioned. *) + from partition_partitions[OF part True] + have part_conds1: "\j. j \ set (subarray l p a h1) \ j \ get_array a h1 ! p " + and part_conds2: "\j. j \ set (subarray (p + 1) (r + 1) a h1) \ get_array a h1 ! p \ j" + by (auto simp add: all_in_set_subarray_conv) + from quicksort_outer_remains [OF qs1] quicksort_permutes [OF qs1] True + length_remains 1(5) pivot multiset_of_sublist [of l p "get_array a h1" "get_array a h2"] + have multiset_partconds1: "multiset_of (subarray l p a h2) = multiset_of (subarray l p a h1)" + unfolding Heap.length_def subarray_def by (cases p, auto) + with left_subarray_remains part_conds1 pivot_unchanged + have part_conds2': "\j. j \ set (subarray l p a h') \ j \ get_array a h' ! p" + by (simp, subst set_of_multiset_of[symmetric], simp) + (* -- These steps are the analogous for the right sublist \ *) + from quicksort_outer_remains [OF qs1] length_remains + have right_subarray_remains: "subarray (p + 1) (r + 1) a h1 = subarray (p + 1) (r + 1) a h2" + by (auto simp add: subarray_eq_samelength_iff) + from quicksort_outer_remains [OF qs2] quicksort_permutes [OF qs2] True + length_remains 1(5) pivot multiset_of_sublist [of "p + 1" "r + 1" "get_array a h2" "get_array a h'"] + have multiset_partconds2: "multiset_of (subarray (p + 1) (r + 1) a h') = multiset_of (subarray (p + 1) (r + 1) a h2)" + unfolding Heap.length_def subarray_def by auto + with right_subarray_remains part_conds2 pivot_unchanged + have part_conds1': "\j. j \ set (subarray (p + 1) (r + 1) a h') \ get_array a h' ! p \ j" + by (simp, subst set_of_multiset_of[symmetric], simp) + (* -- Thirdly and finally, we show that the array is sorted + following from the facts above. *) + 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'" + by (simp add: subarray_nth_array_Cons, cases "l < p") (auto simp add: subarray_append subarray_Nil) + with IH1' IH2 part_conds1' part_conds2' pivot have ?thesis + unfolding subarray_def + apply (auto simp add: sorted_append sorted_Cons all_in_set_sublist'_conv) + by (auto simp add: set_sublist' dest: le_trans [of _ "get_array a h' ! p"]) + } + with True cr show ?thesis + unfolding quicksort.simps [of a l r] run_drop + by (elim crel_if crel_return crelE crel_assert) auto + qed +qed + + +lemma quicksort_is_sort: + assumes crel: "crel (quicksort a 0 (Heap.length a h - 1)) h h' rs" + shows "get_array a h' = sort (get_array a h)" +proof (cases "get_array a h = []") + case True + with quicksort_is_skip[OF crel] show ?thesis + unfolding Heap.length_def by simp +next + case False + from quicksort_sorts [OF crel] False have "sorted (sublist' 0 (List.length (get_array a h)) (get_array a h'))" + unfolding Heap.length_def subarray_def by auto + with length_remains[OF crel] have "sorted (get_array a h')" + unfolding Heap.length_def by simp + with quicksort_permutes [OF crel] properties_for_sort show ?thesis by fastsimp +qed + +subsection {* No Errors in quicksort *} +text {* We have proved that quicksort sorts (if no exceptions occur). +We will now show that exceptions do not occur. *} + +lemma noError_part1: + assumes "l < Heap.length a h" "r < Heap.length a h" + shows "noError (part1 a l r p) h" + using assms +proof (induct a l r p arbitrary: h rule: part1.induct) + case (1 a l r p) + thus ?case + unfolding part1.simps [of a l r] swap_def run_drop + by (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd elim!: crelE crel_upd crel_nth crel_return) +qed + +lemma noError_partition: + assumes "l < r" "l < Heap.length a h" "r < Heap.length a h" + shows "noError (partition a l r) h" +using assms +unfolding partition.simps swap_def run_drop +apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_part1 elim!: crelE crel_upd crel_nth crel_return) +apply (frule part_length_remains) +apply (frule part_returns_index_in_bounds) +apply auto +apply (frule part_length_remains) +apply (frule part_returns_index_in_bounds) +apply auto +apply (frule part_length_remains) +apply auto +done + +lemma noError_quicksort: + assumes "l < Heap.length a h" "r < Heap.length a h" + shows "noError (quicksort a l r) h" +using assms +proof (induct a l r arbitrary: h rule: quicksort.induct) + case (1 a l ri h) + thus ?case + unfolding quicksort.simps [of a l ri] run_drop + apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_assert noError_partition) + apply (frule partition_returns_index_in_bounds) + apply auto + apply (frule partition_returns_index_in_bounds) + apply auto + apply (auto elim!: crel_assert dest!: partition_length_remains length_remains) + apply (subgoal_tac "Suc r \ ri \ r = ri") + apply (erule disjE) + apply auto + unfolding quicksort.simps [of a "Suc ri" ri] run_drop + apply (auto intro!: noError_if noError_return) + done +qed + +end \ No newline at end of file