src/HOL/ex/Qsort.thy
author nipkow
Wed, 06 Jun 2007 19:12:59 +0200
changeset 23281 e26ec695c9b3
parent 15815 62854cac5410
permissions -rw-r--r--
changed filter syntax from : to <-
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8414
5983668cac15 now uses recdef instead of "rules"
paulson
parents: 1476
diff changeset
     1
(*  Title:      HOL/ex/Qsort.thy
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1376
diff changeset
     3
    Author:     Tobias Nipkow
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1994 TU Muenchen
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     5
*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     6
15815
paulson
parents: 15631
diff changeset
     7
header{*Quicksort*}
8590
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
     8
15815
paulson
parents: 15631
diff changeset
     9
theory Qsort
paulson
parents: 15631
diff changeset
    10
imports Sorting
paulson
parents: 15631
diff changeset
    11
begin
paulson
parents: 15631
diff changeset
    12
paulson
parents: 15631
diff changeset
    13
subsection{*Version 1: higher-order*}
paulson
parents: 15631
diff changeset
    14
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    15
consts qsort :: "('a \<Rightarrow> 'a => bool) * 'a list \<Rightarrow> 'a list"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    16
8414
5983668cac15 now uses recdef instead of "rules"
paulson
parents: 1476
diff changeset
    17
recdef qsort "measure (size o snd)"
15815
paulson
parents: 15631
diff changeset
    18
    "qsort(le, [])   = []"
23281
e26ec695c9b3 changed filter syntax from : to <-
nipkow
parents: 15815
diff changeset
    19
    "qsort(le, x#xs) = qsort(le, [y\<leftarrow>xs . ~ le x y]) @ [x] @
e26ec695c9b3 changed filter syntax from : to <-
nipkow
parents: 15815
diff changeset
    20
		       qsort(le, [y\<leftarrow>xs . le x y])"
15247
nipkow
parents: 13159
diff changeset
    21
(hints recdef_simp: length_filter_le[THEN le_less_trans])
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    22
15815
paulson
parents: 15631
diff changeset
    23
lemma qsort_permutes [simp]:
paulson
parents: 15631
diff changeset
    24
     "multiset_of (qsort(le,xs)) = multiset_of xs"
15631
cbee04ce413b use Library/Multiset instead of own definition
kleing
parents: 15247
diff changeset
    25
by (induct le xs rule: qsort.induct) (auto simp: union_ac)
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    26
15815
paulson
parents: 15631
diff changeset
    27
lemma set_qsort [simp]: "set (qsort(le,xs)) = set xs";
15631
cbee04ce413b use Library/Multiset instead of own definition
kleing
parents: 15247
diff changeset
    28
by(simp add: set_count_greater_0)
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    29
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    30
lemma sorted_qsort:
15815
paulson
parents: 15631
diff changeset
    31
     "total(le) ==> transf(le) ==> sorted le (qsort(le,xs))"
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    32
apply (induct le xs rule: qsort.induct)
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    33
 apply simp
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    34
apply simp
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    35
apply(unfold Sorting.total_def Sorting.transf_def)
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    36
apply blast
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    37
done
8414
5983668cac15 now uses recdef instead of "rules"
paulson
parents: 1476
diff changeset
    38
8590
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    39
15815
paulson
parents: 15631
diff changeset
    40
subsection{*Version 2:type classes*}
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    41
8590
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    42
consts quickSort :: "('a::linorder) list => 'a list"
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    43
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    44
recdef quickSort "measure size"
15815
paulson
parents: 15631
diff changeset
    45
    "quickSort []   = []"
23281
e26ec695c9b3 changed filter syntax from : to <-
nipkow
parents: 15815
diff changeset
    46
    "quickSort (x#l) = quickSort [y\<leftarrow>l. ~ x\<le>y] @ [x] @ quickSort [y\<leftarrow>l. x\<le>y]"
15247
nipkow
parents: 13159
diff changeset
    47
(hints recdef_simp: length_filter_le[THEN le_less_trans])
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    48
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    49
lemma quickSort_permutes[simp]:
15631
cbee04ce413b use Library/Multiset instead of own definition
kleing
parents: 15247
diff changeset
    50
 "multiset_of (quickSort xs) = multiset_of xs"
cbee04ce413b use Library/Multiset instead of own definition
kleing
parents: 15247
diff changeset
    51
by (induct xs rule: quickSort.induct) (auto simp: union_ac)
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    52
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    53
lemma set_quickSort[simp]: "set (quickSort xs) = set xs"
15631
cbee04ce413b use Library/Multiset instead of own definition
kleing
parents: 15247
diff changeset
    54
by(simp add: set_count_greater_0)
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    55
15815
paulson
parents: 15631
diff changeset
    56
theorem sorted_quickSort: "sorted (op \<le>) (quickSort xs)"
paulson
parents: 15631
diff changeset
    57
by (induct xs rule: quickSort.induct, auto)
8590
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    58
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    59
end