src/HOL/Data_Structures/Sorting.thy
author nipkow
Sun, 13 May 2018 13:15:50 +0200
changeset 68159 620ca44d8b7d
parent 68158 b00f0f990bc5
child 68160 efce008331f6
permissions -rw-r--r--
removed unused lemma
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
66543
a90dbf19f573 new file
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
a90dbf19f573 new file
nipkow
parents:
diff changeset
     2
a90dbf19f573 new file
nipkow
parents:
diff changeset
     3
(* FIXME adjust List.sorted to work like below; [code] is different! *)
a90dbf19f573 new file
nipkow
parents:
diff changeset
     4
a90dbf19f573 new file
nipkow
parents:
diff changeset
     5
theory Sorting
a90dbf19f573 new file
nipkow
parents:
diff changeset
     6
imports
a90dbf19f573 new file
nipkow
parents:
diff changeset
     7
  Complex_Main
a90dbf19f573 new file
nipkow
parents:
diff changeset
     8
  "HOL-Library.Multiset"
a90dbf19f573 new file
nipkow
parents:
diff changeset
     9
begin
a90dbf19f573 new file
nipkow
parents:
diff changeset
    10
a90dbf19f573 new file
nipkow
parents:
diff changeset
    11
hide_const List.sorted List.insort
a90dbf19f573 new file
nipkow
parents:
diff changeset
    12
a90dbf19f573 new file
nipkow
parents:
diff changeset
    13
declare Let_def [simp]
a90dbf19f573 new file
nipkow
parents:
diff changeset
    14
a90dbf19f573 new file
nipkow
parents:
diff changeset
    15
fun sorted :: "'a::linorder list \<Rightarrow> bool" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
    16
"sorted [] = True" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
    17
"sorted (x # xs) = ((\<forall>y\<in>set xs. x \<le> y) & sorted xs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    18
a90dbf19f573 new file
nipkow
parents:
diff changeset
    19
lemma sorted_append:
a90dbf19f573 new file
nipkow
parents:
diff changeset
    20
  "sorted (xs@ys) = (sorted xs & sorted ys & (\<forall>x \<in> set xs. \<forall>y \<in> set ys. x\<le>y))"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    21
by (induct xs) (auto)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    22
67983
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
    23
lemma sorted01: "length xs \<le> 1 \<Longrightarrow> sorted xs"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
    24
by(auto simp: le_Suc_eq length_Suc_conv)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
    25
66543
a90dbf19f573 new file
nipkow
parents:
diff changeset
    26
a90dbf19f573 new file
nipkow
parents:
diff changeset
    27
subsection "Insertion Sort"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    28
a90dbf19f573 new file
nipkow
parents:
diff changeset
    29
fun insort :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
    30
"insort x [] = [x]" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
    31
"insort x (y#ys) =
a90dbf19f573 new file
nipkow
parents:
diff changeset
    32
  (if x \<le> y then x#y#ys else y#(insort x ys))"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    33
a90dbf19f573 new file
nipkow
parents:
diff changeset
    34
fun isort :: "'a::linorder list \<Rightarrow> 'a list" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
    35
"isort [] = []" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
    36
"isort (x#xs) = insort x (isort xs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    37
68078
nipkow
parents: 67983
diff changeset
    38
66543
a90dbf19f573 new file
nipkow
parents:
diff changeset
    39
subsubsection "Functional Correctness"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    40
a90dbf19f573 new file
nipkow
parents:
diff changeset
    41
lemma mset_insort: "mset (insort x xs) = add_mset x (mset xs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    42
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    43
apply auto
a90dbf19f573 new file
nipkow
parents:
diff changeset
    44
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
    45
a90dbf19f573 new file
nipkow
parents:
diff changeset
    46
lemma mset_isort: "mset (isort xs) = mset xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    47
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    48
apply simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
    49
apply (simp add: mset_insort)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    50
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
    51
a90dbf19f573 new file
nipkow
parents:
diff changeset
    52
lemma set_insort: "set (insort x xs) = insert x (set xs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    53
by (metis mset_insort set_mset_add_mset_insert set_mset_mset)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    54
a90dbf19f573 new file
nipkow
parents:
diff changeset
    55
lemma sorted_insort: "sorted (insort a xs) = sorted xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    56
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    57
apply(auto simp add: set_insort)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    58
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
    59
a90dbf19f573 new file
nipkow
parents:
diff changeset
    60
lemma "sorted (isort xs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    61
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    62
apply(auto simp: sorted_insort)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    63
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
    64
68078
nipkow
parents: 67983
diff changeset
    65
nipkow
parents: 67983
diff changeset
    66
subsubsection "Time Complexity"
66543
a90dbf19f573 new file
nipkow
parents:
diff changeset
    67
a90dbf19f573 new file
nipkow
parents:
diff changeset
    68
text \<open>We count the number of function calls.\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    69
a90dbf19f573 new file
nipkow
parents:
diff changeset
    70
text\<open>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    71
\<open>insort x [] = [x]\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    72
\<open>insort x (y#ys) =
a90dbf19f573 new file
nipkow
parents:
diff changeset
    73
  (if x \<le> y then x#y#ys else y#(insort x ys))\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    74
\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    75
fun t_insort :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> nat" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
    76
"t_insort x [] = 1" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
    77
"t_insort x (y#ys) =
a90dbf19f573 new file
nipkow
parents:
diff changeset
    78
  (if x \<le> y then 0 else t_insort x ys) + 1"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    79
a90dbf19f573 new file
nipkow
parents:
diff changeset
    80
text\<open>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    81
\<open>isort [] = []\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    82
\<open>isort (x#xs) = insort x (isort xs)\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    83
\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    84
fun t_isort :: "'a::linorder list \<Rightarrow> nat" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
    85
"t_isort [] = 1" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
    86
"t_isort (x#xs) = t_isort xs + t_insort x (isort xs) + 1"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    87
a90dbf19f573 new file
nipkow
parents:
diff changeset
    88
a90dbf19f573 new file
nipkow
parents:
diff changeset
    89
lemma t_insort_length: "t_insort x xs \<le> length xs + 1"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    90
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    91
apply auto
a90dbf19f573 new file
nipkow
parents:
diff changeset
    92
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
    93
a90dbf19f573 new file
nipkow
parents:
diff changeset
    94
lemma length_insort: "length (insort x xs) = length xs + 1"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    95
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    96
apply auto
a90dbf19f573 new file
nipkow
parents:
diff changeset
    97
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
    98
a90dbf19f573 new file
nipkow
parents:
diff changeset
    99
lemma length_isort: "length (isort xs) = length xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   100
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   101
apply (auto simp: length_insort)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   102
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
   103
a90dbf19f573 new file
nipkow
parents:
diff changeset
   104
lemma t_isort_length: "t_isort xs \<le> (length xs + 1) ^ 2"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   105
proof(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   106
  case Nil show ?case by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   107
next
a90dbf19f573 new file
nipkow
parents:
diff changeset
   108
  case (Cons x xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   109
  have "t_isort (x#xs) = t_isort xs + t_insort x (isort xs) + 1" by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   110
  also have "\<dots> \<le> (length xs + 1) ^ 2 + t_insort x (isort xs) + 1"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   111
    using Cons.IH by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   112
  also have "\<dots> \<le> (length xs + 1) ^ 2 + length xs + 1 + 1"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   113
    using t_insort_length[of x "isort xs"] by (simp add: length_isort)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   114
  also have "\<dots> \<le> (length(x#xs) + 1) ^ 2"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   115
    by (simp add: power2_eq_square)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   116
  finally show ?case .
a90dbf19f573 new file
nipkow
parents:
diff changeset
   117
qed
a90dbf19f573 new file
nipkow
parents:
diff changeset
   118
a90dbf19f573 new file
nipkow
parents:
diff changeset
   119
a90dbf19f573 new file
nipkow
parents:
diff changeset
   120
subsection "Merge Sort"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   121
a90dbf19f573 new file
nipkow
parents:
diff changeset
   122
fun merge :: "'a::linorder list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
   123
"merge [] ys = ys" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
   124
"merge xs [] = xs" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
   125
"merge (x#xs) (y#ys) = (if x \<le> y then x # merge xs (y#ys) else y # merge (x#xs) ys)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   126
a90dbf19f573 new file
nipkow
parents:
diff changeset
   127
fun msort :: "'a::linorder list \<Rightarrow> 'a list" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
   128
"msort xs = (let n = length xs in
a90dbf19f573 new file
nipkow
parents:
diff changeset
   129
  if n \<le> 1 then xs
a90dbf19f573 new file
nipkow
parents:
diff changeset
   130
  else merge (msort (take (n div 2) xs)) (msort (drop (n div 2) xs)))"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   131
a90dbf19f573 new file
nipkow
parents:
diff changeset
   132
declare msort.simps [simp del]
a90dbf19f573 new file
nipkow
parents:
diff changeset
   133
68078
nipkow
parents: 67983
diff changeset
   134
67983
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   135
subsubsection "Functional Correctness"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   136
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   137
lemma mset_merge: "mset(merge xs ys) = mset xs + mset ys"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   138
by(induction xs ys rule: merge.induct) auto
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   139
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   140
lemma "mset (msort xs) = mset xs"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   141
proof(induction xs rule: msort.induct)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   142
  case (1 xs)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   143
  let ?n = "length xs"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   144
  let ?xs1 = "take (?n div 2) xs"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   145
  let ?xs2 = "drop (?n div 2) xs"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   146
  show ?case
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   147
  proof cases
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   148
    assume "?n \<le> 1"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   149
    thus ?thesis by(simp add: msort.simps[of xs])
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   150
  next
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   151
    assume "\<not> ?n \<le> 1"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   152
    hence "mset (msort xs) = mset (msort ?xs1) + mset (msort ?xs2)"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   153
      by(simp add: msort.simps[of xs] mset_merge)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   154
    also have "\<dots> = mset ?xs1 + mset ?xs2"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   155
      using \<open>\<not> ?n \<le> 1\<close> by(simp add: "1.IH")
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   156
    also have "\<dots> = mset (?xs1 @ ?xs2)" by (simp del: append_take_drop_id)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   157
    also have "\<dots> = mset xs" by simp
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   158
    finally show ?thesis .
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   159
  qed
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   160
qed
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   161
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   162
lemma set_merge: "set(merge xs ys) = set xs \<union> set ys"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   163
by(induction xs ys rule: merge.induct) (auto)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   164
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   165
lemma sorted_merge: "sorted (merge xs ys) \<longleftrightarrow> (sorted xs \<and> sorted ys)"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   166
by(induction xs ys rule: merge.induct) (auto simp: set_merge)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   167
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   168
lemma "sorted (msort xs)"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   169
proof(induction xs rule: msort.induct)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   170
  case (1 xs)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   171
  let ?n = "length xs"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   172
  show ?case
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   173
  proof cases
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   174
    assume "?n \<le> 1"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   175
    thus ?thesis by(simp add: msort.simps[of xs] sorted01)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   176
  next
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   177
    assume "\<not> ?n \<le> 1"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   178
    thus ?thesis using "1.IH"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   179
      by(simp add: sorted_merge  msort.simps[of xs] mset_merge)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   180
  qed
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   181
qed
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   182
68078
nipkow
parents: 67983
diff changeset
   183
nipkow
parents: 67983
diff changeset
   184
subsubsection "Time Complexity"
67983
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   185
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   186
text \<open>We only count the number of comparisons between list elements.\<close>
66543
a90dbf19f573 new file
nipkow
parents:
diff changeset
   187
a90dbf19f573 new file
nipkow
parents:
diff changeset
   188
fun c_merge :: "'a::linorder list \<Rightarrow> 'a list \<Rightarrow> nat" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
   189
"c_merge [] ys = 0" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
   190
"c_merge xs [] = 0" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
   191
"c_merge (x#xs) (y#ys) = 1 + (if x \<le> y then c_merge xs (y#ys) else c_merge (x#xs) ys)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   192
a90dbf19f573 new file
nipkow
parents:
diff changeset
   193
lemma c_merge_ub: "c_merge xs ys \<le> length xs + length ys"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   194
by (induction xs ys rule: c_merge.induct) auto
a90dbf19f573 new file
nipkow
parents:
diff changeset
   195
a90dbf19f573 new file
nipkow
parents:
diff changeset
   196
fun c_msort :: "'a::linorder list \<Rightarrow> nat" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
   197
"c_msort xs =
a90dbf19f573 new file
nipkow
parents:
diff changeset
   198
  (let n = length xs;
a90dbf19f573 new file
nipkow
parents:
diff changeset
   199
       ys = take (n div 2) xs;
a90dbf19f573 new file
nipkow
parents:
diff changeset
   200
       zs = drop (n div 2) xs
a90dbf19f573 new file
nipkow
parents:
diff changeset
   201
   in if n \<le> 1 then 0
a90dbf19f573 new file
nipkow
parents:
diff changeset
   202
      else c_msort ys + c_msort zs + c_merge (msort ys) (msort zs))"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   203
a90dbf19f573 new file
nipkow
parents:
diff changeset
   204
declare c_msort.simps [simp del]
a90dbf19f573 new file
nipkow
parents:
diff changeset
   205
a90dbf19f573 new file
nipkow
parents:
diff changeset
   206
lemma length_merge: "length(merge xs ys) = length xs + length ys"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   207
apply (induction xs ys rule: merge.induct)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   208
apply auto
a90dbf19f573 new file
nipkow
parents:
diff changeset
   209
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
   210
a90dbf19f573 new file
nipkow
parents:
diff changeset
   211
lemma length_msort: "length(msort xs) = length xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   212
proof (induction xs rule: msort.induct)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   213
  case (1 xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   214
  thus ?case by (auto simp: msort.simps[of xs] length_merge)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   215
qed
a90dbf19f573 new file
nipkow
parents:
diff changeset
   216
text \<open>Why structured proof?
a90dbf19f573 new file
nipkow
parents:
diff changeset
   217
   To have the name "xs" to specialize msort.simps with xs
a90dbf19f573 new file
nipkow
parents:
diff changeset
   218
   to ensure that msort.simps cannot be used recursively.
a90dbf19f573 new file
nipkow
parents:
diff changeset
   219
Also works without this precaution, but that is just luck.\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
   220
a90dbf19f573 new file
nipkow
parents:
diff changeset
   221
lemma c_msort_le: "length xs = 2^k \<Longrightarrow> c_msort xs \<le> k * 2^k"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   222
proof(induction k arbitrary: xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   223
  case 0 thus ?case by (simp add: c_msort.simps)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   224
next
a90dbf19f573 new file
nipkow
parents:
diff changeset
   225
  case (Suc k)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   226
  let ?n = "length xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   227
  let ?ys = "take (?n div 2) xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   228
  let ?zs = "drop (?n div 2) xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   229
  show ?case
a90dbf19f573 new file
nipkow
parents:
diff changeset
   230
  proof (cases "?n \<le> 1")
a90dbf19f573 new file
nipkow
parents:
diff changeset
   231
    case True
a90dbf19f573 new file
nipkow
parents:
diff changeset
   232
    thus ?thesis by(simp add: c_msort.simps)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   233
  next
a90dbf19f573 new file
nipkow
parents:
diff changeset
   234
    case False
a90dbf19f573 new file
nipkow
parents:
diff changeset
   235
    have "c_msort(xs) =
a90dbf19f573 new file
nipkow
parents:
diff changeset
   236
      c_msort ?ys + c_msort ?zs + c_merge (msort ?ys) (msort ?zs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   237
      by (simp add: c_msort.simps msort.simps)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   238
    also have "\<dots> \<le> c_msort ?ys + c_msort ?zs + length ?ys + length ?zs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   239
      using c_merge_ub[of "msort ?ys" "msort ?zs"] length_msort[of ?ys] length_msort[of ?zs]
a90dbf19f573 new file
nipkow
parents:
diff changeset
   240
      by arith
a90dbf19f573 new file
nipkow
parents:
diff changeset
   241
    also have "\<dots> \<le> k * 2^k + c_msort ?zs + length ?ys + length ?zs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   242
      using Suc.IH[of ?ys] Suc.prems by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   243
    also have "\<dots> \<le> k * 2^k + k * 2^k + length ?ys + length ?zs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   244
      using Suc.IH[of ?zs] Suc.prems by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   245
    also have "\<dots> = 2 * k * 2^k + 2 * 2 ^ k"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   246
      using Suc.prems by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   247
    finally show ?thesis by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   248
  qed
a90dbf19f573 new file
nipkow
parents:
diff changeset
   249
qed
a90dbf19f573 new file
nipkow
parents:
diff changeset
   250
a90dbf19f573 new file
nipkow
parents:
diff changeset
   251
(* Beware of conversions: *)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   252
lemma "length xs = 2^k \<Longrightarrow> c_msort xs \<le> length xs * log 2 (length xs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   253
using c_msort_le[of xs k] apply (simp add: log_nat_power algebra_simps)
66912
a99a7cbf0fb5 generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents: 66543
diff changeset
   254
by (metis (mono_tags) numeral_power_eq_of_nat_cancel_iff of_nat_le_iff of_nat_mult)
66543
a90dbf19f573 new file
nipkow
parents:
diff changeset
   255
67983
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   256
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   257
subsection "Bottom-Up Merge Sort"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   258
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   259
(* Exercise: make tail recursive *)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   260
fun merge_adj :: "('a::linorder) list list \<Rightarrow> 'a list list" where
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   261
"merge_adj [] = []" |
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   262
"merge_adj [xs] = [xs]" |
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   263
"merge_adj (xs # ys # zss) = merge xs ys # merge_adj zss"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   264
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   265
text \<open>For the termination proof of \<open>merge_all\<close> below.\<close>
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   266
lemma length_merge_adjacent[simp]: "length (merge_adj xs) = (length xs + 1) div 2"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   267
by (induction xs rule: merge_adj.induct) auto
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   268
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   269
fun merge_all :: "('a::linorder) list list \<Rightarrow> 'a list" where
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   270
"merge_all [] = undefined" |
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   271
"merge_all [xs] = xs" |
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   272
"merge_all xss = merge_all (merge_adj xss)"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   273
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   274
definition msort_bu :: "('a::linorder) list \<Rightarrow> 'a list" where
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   275
"msort_bu xs = (if xs = [] then [] else merge_all (map (\<lambda>x. [x]) xs))"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   276
68078
nipkow
parents: 67983
diff changeset
   277
67983
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   278
subsubsection "Functional Correctness"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   279
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   280
lemma mset_merge_adj:
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   281
  "\<Union># image_mset mset (mset (merge_adj xss)) = \<Union># image_mset mset (mset xss)"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   282
by(induction xss rule: merge_adj.induct) (auto simp: mset_merge)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   283
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   284
lemma msec_merge_all:
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   285
  "xss \<noteq> [] \<Longrightarrow> mset (merge_all xss) = (\<Union># (mset (map mset xss)))"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   286
by(induction xss rule: merge_all.induct) (auto simp: mset_merge mset_merge_adj)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   287
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   288
lemma sorted_merge_adj:
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   289
  "\<forall>xs \<in> set xss. sorted xs \<Longrightarrow> \<forall>xs \<in> set (merge_adj xss). sorted xs"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   290
by(induction xss rule: merge_adj.induct) (auto simp: sorted_merge)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   291
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   292
lemma sorted_merge_all:
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   293
  "\<forall>xs \<in> set xss. sorted xs \<Longrightarrow> xss \<noteq> [] \<Longrightarrow> sorted (merge_all xss)"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   294
apply(induction xss rule: merge_all.induct)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   295
using [[simp_depth_limit=3]] by (auto simp add: sorted_merge_adj)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   296
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   297
lemma sorted_msort_bu: "sorted (msort_bu xs)"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   298
by(simp add: msort_bu_def sorted_merge_all)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   299
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   300
lemma mset_msort: "mset (msort_bu xs) = mset xs"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   301
by(simp add: msort_bu_def msec_merge_all comp_def)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   302
68078
nipkow
parents: 67983
diff changeset
   303
nipkow
parents: 67983
diff changeset
   304
subsubsection "Time Complexity"
67983
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   305
68139
cba8eaa2174f simpler types
nipkow
parents: 68079
diff changeset
   306
fun c_merge_adj :: "('a::linorder) list list \<Rightarrow> nat" where
67983
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   307
"c_merge_adj [] = 0" |
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   308
"c_merge_adj [x] = 0" |
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   309
"c_merge_adj (x # y # zs) = c_merge x y + c_merge_adj zs"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   310
68139
cba8eaa2174f simpler types
nipkow
parents: 68079
diff changeset
   311
fun c_merge_all :: "('a::linorder) list list \<Rightarrow> nat" where
68079
nipkow
parents: 68078
diff changeset
   312
"c_merge_all [] = undefined" |
67983
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   313
"c_merge_all [x] = 0" |
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   314
"c_merge_all xs = c_merge_adj xs + c_merge_all (merge_adj xs)"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   315
68139
cba8eaa2174f simpler types
nipkow
parents: 68079
diff changeset
   316
definition c_msort_bu :: "('a::linorder) list \<Rightarrow> nat" where
67983
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   317
"c_msort_bu xs = (if xs = [] then 0 else c_merge_all (map (\<lambda>x. [x]) xs))"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   318
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   319
lemma length_merge_adj:
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   320
  "\<lbrakk> even(length xs); \<forall>x \<in> set xs. length x = m \<rbrakk> \<Longrightarrow> \<forall>x \<in> set (merge_adj xs). length x = 2*m"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   321
by(induction xs rule: merge_adj.induct) (auto simp: length_merge)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   322
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   323
lemma c_merge_adj: "\<forall>x \<in> set xs. length x = m \<Longrightarrow> c_merge_adj xs \<le> m * length xs"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   324
proof(induction xs rule: c_merge_adj.induct)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   325
  case 1 thus ?case by simp
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   326
next
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   327
  case 2 thus ?case by simp
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   328
next
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   329
  case (3 x y) thus ?case using c_merge_ub[of x y] by (simp add: algebra_simps)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   330
qed
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   331
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   332
lemma c_merge_all: "\<lbrakk> \<forall>x \<in> set xs. length x = m; length xs = 2^k \<rbrakk>
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   333
  \<Longrightarrow> c_merge_all xs \<le> m * k * 2^k"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   334
proof (induction xs arbitrary: k m rule: c_merge_all.induct)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   335
  case 1 thus ?case by simp
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   336
next
68158
nipkow
parents: 68139
diff changeset
   337
  case 2 thus ?case by simp
67983
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   338
next
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   339
  case (3 x y xs)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   340
  let ?xs = "x # y # xs"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   341
  let ?xs2 = "merge_adj ?xs"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   342
  obtain k' where k': "k = Suc k'" using "3.prems"(2)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   343
    by (metis length_Cons nat.inject nat_power_eq_Suc_0_iff nat.exhaust)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   344
  have "even (length xs)" using "3.prems"(2) even_Suc_Suc_iff by fastforce
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   345
  from "3.prems"(1) length_merge_adj[OF this]
68079
nipkow
parents: 68078
diff changeset
   346
  have *: "\<forall>x \<in> set(merge_adj ?xs). length x = 2*m" by(auto simp: length_merge)
nipkow
parents: 68078
diff changeset
   347
  have **: "length ?xs2 = 2 ^ k'" using "3.prems"(2) k' by auto
67983
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   348
  have "c_merge_all ?xs = c_merge_adj ?xs + c_merge_all ?xs2" by simp
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   349
  also have "\<dots> \<le> m * 2^k + c_merge_all ?xs2"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   350
    using "3.prems"(2) c_merge_adj[OF "3.prems"(1)] by (auto simp: algebra_simps)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   351
  also have "\<dots> \<le> m * 2^k + (2*m) * k' * 2^k'"
68079
nipkow
parents: 68078
diff changeset
   352
    using "3.IH"[OF * **] by simp
67983
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   353
  also have "\<dots> = m * k * 2^k"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   354
    using k' by (simp add: algebra_simps)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   355
  finally show ?case .
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   356
qed
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   357
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   358
corollary c_msort_bu: "length xs = 2 ^ k \<Longrightarrow> c_msort_bu xs \<le> k * 2 ^ k"
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   359
using c_merge_all[of "map (\<lambda>x. [x]) xs" 1] by (simp add: c_msort_bu_def)
487685540a51 added bottom-up merge sort
nipkow
parents: 66912
diff changeset
   360
66543
a90dbf19f573 new file
nipkow
parents:
diff changeset
   361
end