src/HOL/Data_Structures/Sorting.thy
author nipkow
Tue, 29 Aug 2017 12:05:00 +0200
changeset 66543 a90dbf19f573
child 66912 a99a7cbf0fb5
permissions -rw-r--r--
new file
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
a90dbf19f573 new file
nipkow
parents:
diff changeset
    23
a90dbf19f573 new file
nipkow
parents:
diff changeset
    24
subsection "Insertion Sort"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    25
a90dbf19f573 new file
nipkow
parents:
diff changeset
    26
fun insort :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
    27
"insort x [] = [x]" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
    28
"insort x (y#ys) =
a90dbf19f573 new file
nipkow
parents:
diff changeset
    29
  (if x \<le> y then x#y#ys else y#(insort x ys))"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    30
a90dbf19f573 new file
nipkow
parents:
diff changeset
    31
fun isort :: "'a::linorder list \<Rightarrow> 'a list" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
    32
"isort [] = []" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
    33
"isort (x#xs) = insort x (isort xs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    34
a90dbf19f573 new file
nipkow
parents:
diff changeset
    35
subsubsection "Functional Correctness"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    36
a90dbf19f573 new file
nipkow
parents:
diff changeset
    37
lemma mset_insort: "mset (insort x xs) = add_mset x (mset xs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    38
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    39
apply auto
a90dbf19f573 new file
nipkow
parents:
diff changeset
    40
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
    41
a90dbf19f573 new file
nipkow
parents:
diff changeset
    42
lemma mset_isort: "mset (isort xs) = mset xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    43
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    44
apply simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
    45
apply (simp add: mset_insort)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    46
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
    47
a90dbf19f573 new file
nipkow
parents:
diff changeset
    48
lemma "sorted (insort a xs) = sorted xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    49
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    50
apply (auto)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    51
oops
a90dbf19f573 new file
nipkow
parents:
diff changeset
    52
a90dbf19f573 new file
nipkow
parents:
diff changeset
    53
lemma set_insort: "set (insort x xs) = insert x (set xs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    54
by (metis mset_insort set_mset_add_mset_insert set_mset_mset)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    55
a90dbf19f573 new file
nipkow
parents:
diff changeset
    56
lemma set_isort: "set (isort xs) = set xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    57
by (metis mset_isort set_mset_mset)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    58
a90dbf19f573 new file
nipkow
parents:
diff changeset
    59
lemma sorted_insort: "sorted (insort a xs) = sorted xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    60
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    61
apply(auto simp add: set_insort)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    62
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
    63
a90dbf19f573 new file
nipkow
parents:
diff changeset
    64
lemma "sorted (isort xs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    65
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    66
apply(auto simp: sorted_insort)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    67
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
    68
a90dbf19f573 new file
nipkow
parents:
diff changeset
    69
subsection "Time Complexity"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    70
a90dbf19f573 new file
nipkow
parents:
diff changeset
    71
text \<open>We count the number of function calls.\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    72
a90dbf19f573 new file
nipkow
parents:
diff changeset
    73
text\<open>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    74
\<open>insort x [] = [x]\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    75
\<open>insort x (y#ys) =
a90dbf19f573 new file
nipkow
parents:
diff changeset
    76
  (if x \<le> y then x#y#ys else y#(insort x ys))\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    77
\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    78
fun t_insort :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> nat" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
    79
"t_insort x [] = 1" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
    80
"t_insort x (y#ys) =
a90dbf19f573 new file
nipkow
parents:
diff changeset
    81
  (if x \<le> y then 0 else t_insort x ys) + 1"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    82
a90dbf19f573 new file
nipkow
parents:
diff changeset
    83
text\<open>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    84
\<open>isort [] = []\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    85
\<open>isort (x#xs) = insort x (isort xs)\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    86
\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
    87
fun t_isort :: "'a::linorder list \<Rightarrow> nat" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
    88
"t_isort [] = 1" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
    89
"t_isort (x#xs) = t_isort xs + t_insort x (isort xs) + 1"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    90
a90dbf19f573 new file
nipkow
parents:
diff changeset
    91
a90dbf19f573 new file
nipkow
parents:
diff changeset
    92
lemma t_insort_length: "t_insort x xs \<le> length xs + 1"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    93
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    94
apply auto
a90dbf19f573 new file
nipkow
parents:
diff changeset
    95
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
    96
a90dbf19f573 new file
nipkow
parents:
diff changeset
    97
lemma length_insort: "length (insort x xs) = length xs + 1"
a90dbf19f573 new file
nipkow
parents:
diff changeset
    98
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
    99
apply auto
a90dbf19f573 new file
nipkow
parents:
diff changeset
   100
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
   101
a90dbf19f573 new file
nipkow
parents:
diff changeset
   102
lemma length_isort: "length (isort xs) = length xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   103
apply(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   104
apply (auto simp: length_insort)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   105
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
   106
a90dbf19f573 new file
nipkow
parents:
diff changeset
   107
lemma t_isort_length: "t_isort xs \<le> (length xs + 1) ^ 2"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   108
proof(induction xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   109
  case Nil show ?case by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   110
next
a90dbf19f573 new file
nipkow
parents:
diff changeset
   111
  case (Cons x xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   112
  have "t_isort (x#xs) = t_isort xs + t_insort x (isort xs) + 1" by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   113
  also have "\<dots> \<le> (length xs + 1) ^ 2 + t_insort x (isort xs) + 1"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   114
    using Cons.IH by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   115
  also have "\<dots> \<le> (length xs + 1) ^ 2 + length xs + 1 + 1"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   116
    using t_insort_length[of x "isort xs"] by (simp add: length_isort)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   117
  also have "\<dots> \<le> (length(x#xs) + 1) ^ 2"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   118
    by (simp add: power2_eq_square)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   119
  finally show ?case .
a90dbf19f573 new file
nipkow
parents:
diff changeset
   120
qed
a90dbf19f573 new file
nipkow
parents:
diff changeset
   121
a90dbf19f573 new file
nipkow
parents:
diff changeset
   122
a90dbf19f573 new file
nipkow
parents:
diff changeset
   123
subsection "Merge Sort"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   124
a90dbf19f573 new file
nipkow
parents:
diff changeset
   125
fun merge :: "'a::linorder list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
   126
"merge [] ys = ys" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
   127
"merge xs [] = xs" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
   128
"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
   129
a90dbf19f573 new file
nipkow
parents:
diff changeset
   130
fun msort :: "'a::linorder list \<Rightarrow> 'a list" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
   131
"msort xs = (let n = length xs in
a90dbf19f573 new file
nipkow
parents:
diff changeset
   132
  if n \<le> 1 then xs
a90dbf19f573 new file
nipkow
parents:
diff changeset
   133
  else merge (msort (take (n div 2) xs)) (msort (drop (n div 2) xs)))"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   134
a90dbf19f573 new file
nipkow
parents:
diff changeset
   135
declare msort.simps [simp del]
a90dbf19f573 new file
nipkow
parents:
diff changeset
   136
a90dbf19f573 new file
nipkow
parents:
diff changeset
   137
(* We count the number of comparisons between list elements only *)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   138
a90dbf19f573 new file
nipkow
parents:
diff changeset
   139
fun c_merge :: "'a::linorder list \<Rightarrow> 'a list \<Rightarrow> nat" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
   140
"c_merge [] ys = 0" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
   141
"c_merge xs [] = 0" |
a90dbf19f573 new file
nipkow
parents:
diff changeset
   142
"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
   143
a90dbf19f573 new file
nipkow
parents:
diff changeset
   144
lemma c_merge_ub: "c_merge xs ys \<le> length xs + length ys"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   145
by (induction xs ys rule: c_merge.induct) auto
a90dbf19f573 new file
nipkow
parents:
diff changeset
   146
a90dbf19f573 new file
nipkow
parents:
diff changeset
   147
fun c_msort :: "'a::linorder list \<Rightarrow> nat" where
a90dbf19f573 new file
nipkow
parents:
diff changeset
   148
"c_msort xs =
a90dbf19f573 new file
nipkow
parents:
diff changeset
   149
  (let n = length xs;
a90dbf19f573 new file
nipkow
parents:
diff changeset
   150
       ys = take (n div 2) xs;
a90dbf19f573 new file
nipkow
parents:
diff changeset
   151
       zs = drop (n div 2) xs
a90dbf19f573 new file
nipkow
parents:
diff changeset
   152
   in if n \<le> 1 then 0
a90dbf19f573 new file
nipkow
parents:
diff changeset
   153
      else c_msort ys + c_msort zs + c_merge (msort ys) (msort zs))"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   154
a90dbf19f573 new file
nipkow
parents:
diff changeset
   155
declare c_msort.simps [simp del]
a90dbf19f573 new file
nipkow
parents:
diff changeset
   156
a90dbf19f573 new file
nipkow
parents:
diff changeset
   157
lemma length_merge: "length(merge xs ys) = length xs + length ys"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   158
apply (induction xs ys rule: merge.induct)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   159
apply auto
a90dbf19f573 new file
nipkow
parents:
diff changeset
   160
done
a90dbf19f573 new file
nipkow
parents:
diff changeset
   161
a90dbf19f573 new file
nipkow
parents:
diff changeset
   162
lemma length_msort: "length(msort xs) = length xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   163
proof (induction xs rule: msort.induct)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   164
  case (1 xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   165
  thus ?case by (auto simp: msort.simps[of xs] length_merge)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   166
qed
a90dbf19f573 new file
nipkow
parents:
diff changeset
   167
text \<open>Why structured proof?
a90dbf19f573 new file
nipkow
parents:
diff changeset
   168
   To have the name "xs" to specialize msort.simps with xs
a90dbf19f573 new file
nipkow
parents:
diff changeset
   169
   to ensure that msort.simps cannot be used recursively.
a90dbf19f573 new file
nipkow
parents:
diff changeset
   170
Also works without this precaution, but that is just luck.\<close>
a90dbf19f573 new file
nipkow
parents:
diff changeset
   171
a90dbf19f573 new file
nipkow
parents:
diff changeset
   172
lemma c_msort_le: "length xs = 2^k \<Longrightarrow> c_msort xs \<le> k * 2^k"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   173
proof(induction k arbitrary: xs)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   174
  case 0 thus ?case by (simp add: c_msort.simps)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   175
next
a90dbf19f573 new file
nipkow
parents:
diff changeset
   176
  case (Suc k)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   177
  let ?n = "length xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   178
  let ?ys = "take (?n div 2) xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   179
  let ?zs = "drop (?n div 2) xs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   180
  show ?case
a90dbf19f573 new file
nipkow
parents:
diff changeset
   181
  proof (cases "?n \<le> 1")
a90dbf19f573 new file
nipkow
parents:
diff changeset
   182
    case True
a90dbf19f573 new file
nipkow
parents:
diff changeset
   183
    thus ?thesis by(simp add: c_msort.simps)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   184
  next
a90dbf19f573 new file
nipkow
parents:
diff changeset
   185
    case False
a90dbf19f573 new file
nipkow
parents:
diff changeset
   186
    have "c_msort(xs) =
a90dbf19f573 new file
nipkow
parents:
diff changeset
   187
      c_msort ?ys + c_msort ?zs + c_merge (msort ?ys) (msort ?zs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   188
      by (simp add: c_msort.simps msort.simps)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   189
    also have "\<dots> \<le> c_msort ?ys + c_msort ?zs + length ?ys + length ?zs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   190
      using c_merge_ub[of "msort ?ys" "msort ?zs"] length_msort[of ?ys] length_msort[of ?zs]
a90dbf19f573 new file
nipkow
parents:
diff changeset
   191
      by arith
a90dbf19f573 new file
nipkow
parents:
diff changeset
   192
    also have "\<dots> \<le> k * 2^k + c_msort ?zs + length ?ys + length ?zs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   193
      using Suc.IH[of ?ys] Suc.prems by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   194
    also have "\<dots> \<le> k * 2^k + k * 2^k + length ?ys + length ?zs"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   195
      using Suc.IH[of ?zs] Suc.prems by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   196
    also have "\<dots> = 2 * k * 2^k + 2 * 2 ^ k"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   197
      using Suc.prems by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   198
    finally show ?thesis by simp
a90dbf19f573 new file
nipkow
parents:
diff changeset
   199
  qed
a90dbf19f573 new file
nipkow
parents:
diff changeset
   200
qed
a90dbf19f573 new file
nipkow
parents:
diff changeset
   201
a90dbf19f573 new file
nipkow
parents:
diff changeset
   202
(* Beware of conversions: *)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   203
lemma "length xs = 2^k \<Longrightarrow> c_msort xs \<le> length xs * log 2 (length xs)"
a90dbf19f573 new file
nipkow
parents:
diff changeset
   204
using c_msort_le[of xs k] apply (simp add: log_nat_power algebra_simps)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   205
by (metis (mono_tags) numeral_power_eq_real_of_nat_cancel_iff of_nat_le_iff of_nat_mult)
a90dbf19f573 new file
nipkow
parents:
diff changeset
   206
a90dbf19f573 new file
nipkow
parents:
diff changeset
   207
end