# HG changeset patch # User nipkow # Date 1504001100 -7200 # Node ID a90dbf19f573a2084875d000e39c2a297f341b99 # Parent 075bbb78d33c555b7a38f516304ee2be3a740e35 new file diff -r 075bbb78d33c -r a90dbf19f573 src/HOL/Data_Structures/Sorting.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Data_Structures/Sorting.thy Tue Aug 29 12:05:00 2017 +0200 @@ -0,0 +1,207 @@ +(* Author: Tobias Nipkow *) + +(* FIXME adjust List.sorted to work like below; [code] is different! *) + +theory Sorting +imports + Complex_Main + "HOL-Library.Multiset" +begin + +hide_const List.sorted List.insort + +declare Let_def [simp] + +fun sorted :: "'a::linorder list \ bool" where +"sorted [] = True" | +"sorted (x # xs) = ((\y\set xs. x \ y) & sorted xs)" + +lemma sorted_append: + "sorted (xs@ys) = (sorted xs & sorted ys & (\x \ set xs. \y \ set ys. x\y))" +by (induct xs) (auto) + + +subsection "Insertion Sort" + +fun insort :: "'a::linorder \ 'a list \ 'a list" where +"insort x [] = [x]" | +"insort x (y#ys) = + (if x \ y then x#y#ys else y#(insort x ys))" + +fun isort :: "'a::linorder list \ 'a list" where +"isort [] = []" | +"isort (x#xs) = insort x (isort xs)" + +subsubsection "Functional Correctness" + +lemma mset_insort: "mset (insort x xs) = add_mset x (mset xs)" +apply(induction xs) +apply auto +done + +lemma mset_isort: "mset (isort xs) = mset xs" +apply(induction xs) +apply simp +apply (simp add: mset_insort) +done + +lemma "sorted (insort a xs) = sorted xs" +apply(induction xs) +apply (auto) +oops + +lemma set_insort: "set (insort x xs) = insert x (set xs)" +by (metis mset_insort set_mset_add_mset_insert set_mset_mset) + +lemma set_isort: "set (isort xs) = set xs" +by (metis mset_isort set_mset_mset) + +lemma sorted_insort: "sorted (insort a xs) = sorted xs" +apply(induction xs) +apply(auto simp add: set_insort) +done + +lemma "sorted (isort xs)" +apply(induction xs) +apply(auto simp: sorted_insort) +done + +subsection "Time Complexity" + +text \We count the number of function calls.\ + +text\ +\insort x [] = [x]\ +\insort x (y#ys) = + (if x \ y then x#y#ys else y#(insort x ys))\ +\ +fun t_insort :: "'a::linorder \ 'a list \ nat" where +"t_insort x [] = 1" | +"t_insort x (y#ys) = + (if x \ y then 0 else t_insort x ys) + 1" + +text\ +\isort [] = []\ +\isort (x#xs) = insort x (isort xs)\ +\ +fun t_isort :: "'a::linorder list \ nat" where +"t_isort [] = 1" | +"t_isort (x#xs) = t_isort xs + t_insort x (isort xs) + 1" + + +lemma t_insort_length: "t_insort x xs \ length xs + 1" +apply(induction xs) +apply auto +done + +lemma length_insort: "length (insort x xs) = length xs + 1" +apply(induction xs) +apply auto +done + +lemma length_isort: "length (isort xs) = length xs" +apply(induction xs) +apply (auto simp: length_insort) +done + +lemma t_isort_length: "t_isort xs \ (length xs + 1) ^ 2" +proof(induction xs) + case Nil show ?case by simp +next + case (Cons x xs) + have "t_isort (x#xs) = t_isort xs + t_insort x (isort xs) + 1" by simp + also have "\ \ (length xs + 1) ^ 2 + t_insort x (isort xs) + 1" + using Cons.IH by simp + also have "\ \ (length xs + 1) ^ 2 + length xs + 1 + 1" + using t_insort_length[of x "isort xs"] by (simp add: length_isort) + also have "\ \ (length(x#xs) + 1) ^ 2" + by (simp add: power2_eq_square) + finally show ?case . +qed + + +subsection "Merge Sort" + +fun merge :: "'a::linorder list \ 'a list \ 'a list" where +"merge [] ys = ys" | +"merge xs [] = xs" | +"merge (x#xs) (y#ys) = (if x \ y then x # merge xs (y#ys) else y # merge (x#xs) ys)" + +fun msort :: "'a::linorder list \ 'a list" where +"msort xs = (let n = length xs in + if n \ 1 then xs + else merge (msort (take (n div 2) xs)) (msort (drop (n div 2) xs)))" + +declare msort.simps [simp del] + +(* We count the number of comparisons between list elements only *) + +fun c_merge :: "'a::linorder list \ 'a list \ nat" where +"c_merge [] ys = 0" | +"c_merge xs [] = 0" | +"c_merge (x#xs) (y#ys) = 1 + (if x \ y then c_merge xs (y#ys) else c_merge (x#xs) ys)" + +lemma c_merge_ub: "c_merge xs ys \ length xs + length ys" +by (induction xs ys rule: c_merge.induct) auto + +fun c_msort :: "'a::linorder list \ nat" where +"c_msort xs = + (let n = length xs; + ys = take (n div 2) xs; + zs = drop (n div 2) xs + in if n \ 1 then 0 + else c_msort ys + c_msort zs + c_merge (msort ys) (msort zs))" + +declare c_msort.simps [simp del] + +lemma length_merge: "length(merge xs ys) = length xs + length ys" +apply (induction xs ys rule: merge.induct) +apply auto +done + +lemma length_msort: "length(msort xs) = length xs" +proof (induction xs rule: msort.induct) + case (1 xs) + thus ?case by (auto simp: msort.simps[of xs] length_merge) +qed +text \Why structured proof? + To have the name "xs" to specialize msort.simps with xs + to ensure that msort.simps cannot be used recursively. +Also works without this precaution, but that is just luck.\ + +lemma c_msort_le: "length xs = 2^k \ c_msort xs \ k * 2^k" +proof(induction k arbitrary: xs) + case 0 thus ?case by (simp add: c_msort.simps) +next + case (Suc k) + let ?n = "length xs" + let ?ys = "take (?n div 2) xs" + let ?zs = "drop (?n div 2) xs" + show ?case + proof (cases "?n \ 1") + case True + thus ?thesis by(simp add: c_msort.simps) + next + case False + have "c_msort(xs) = + c_msort ?ys + c_msort ?zs + c_merge (msort ?ys) (msort ?zs)" + by (simp add: c_msort.simps msort.simps) + also have "\ \ c_msort ?ys + c_msort ?zs + length ?ys + length ?zs" + using c_merge_ub[of "msort ?ys" "msort ?zs"] length_msort[of ?ys] length_msort[of ?zs] + by arith + also have "\ \ k * 2^k + c_msort ?zs + length ?ys + length ?zs" + using Suc.IH[of ?ys] Suc.prems by simp + also have "\ \ k * 2^k + k * 2^k + length ?ys + length ?zs" + using Suc.IH[of ?zs] Suc.prems by simp + also have "\ = 2 * k * 2^k + 2 * 2 ^ k" + using Suc.prems by simp + finally show ?thesis by simp + qed +qed + +(* Beware of conversions: *) +lemma "length xs = 2^k \ c_msort xs \ length xs * log 2 (length xs)" +using c_msort_le[of xs k] apply (simp add: log_nat_power algebra_simps) +by (metis (mono_tags) numeral_power_eq_real_of_nat_cancel_iff of_nat_le_iff of_nat_mult) + +end diff -r 075bbb78d33c -r a90dbf19f573 src/HOL/ROOT --- a/src/HOL/ROOT Tue Aug 29 11:08:42 2017 +0200 +++ b/src/HOL/ROOT Tue Aug 29 12:05:00 2017 +0200 @@ -203,6 +203,7 @@ "HOL-Library.Multiset" "HOL-Number_Theory.Fib" theories + Sorting Balance Tree_Map AVL_Map