src/HOL/ex/Qsort.thy
author nipkow
Fri, 17 May 2002 15:40:59 +0200
changeset 13159 2af7b94892ce
parent 8590 89675b444abe
child 15247 98d3ca56684d
permissions -rw-r--r--
Turned into Isar theories.
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
Quicksort
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     7
*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     8
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
     9
theory Qsort = Sorting:
8590
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    10
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    11
(*Version one: higher-order*)
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    12
consts qsort :: "('a \<Rightarrow> 'a => bool) * 'a list \<Rightarrow> 'a list"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    13
8414
5983668cac15 now uses recdef instead of "rules"
paulson
parents: 1476
diff changeset
    14
recdef qsort "measure (size o snd)"
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    15
"qsort(le, [])   = []"
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    16
"qsort(le, x#xs) = qsort(le, [y:xs . ~ le x y]) @ [x] @
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    17
                   qsort(le, [y:xs . le x y])"
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    18
(hints recdef_simp: length_filter[THEN le_less_trans])
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    19
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    20
lemma qsort_permutes[simp]:
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    21
 "multiset (qsort(le,xs)) x = multiset xs x"
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    22
by (induct le xs rule: qsort.induct, auto)
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    23
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    24
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    25
(*Also provable by induction*)
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    26
lemma set_qsort[simp]: "set (qsort(le,xs)) = set xs";
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    27
by(simp add: set_via_multiset)
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    28
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    29
lemma sorted_qsort:
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    30
 "total(le) ==> transf(le) ==> sorted le (qsort(le,xs))"
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    31
apply (induct le xs rule: qsort.induct)
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    32
 apply simp
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(unfold Sorting.total_def Sorting.transf_def)
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    35
apply blast
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    36
done
8414
5983668cac15 now uses recdef instead of "rules"
paulson
parents: 1476
diff changeset
    37
8590
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    38
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    39
(*Version two: type classes*)
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    40
8590
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    41
consts quickSort :: "('a::linorder) list => 'a list"
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    42
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    43
recdef quickSort "measure size"
13159
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    44
"quickSort []   = []"
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    45
"quickSort (x#l) = quickSort [y:l. ~ x<=y] @ [x] @ quickSort [y:l. x<=y]"
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    46
(hints recdef_simp: length_filter[THEN le_less_trans])
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    47
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    48
lemma quickSort_permutes[simp]:
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    49
 "multiset (quickSort xs) z = multiset xs z"
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    50
by (induct xs rule: quickSort.induct) auto
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    51
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    52
(*Also provable by induction*)
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    53
lemma set_quickSort[simp]: "set (quickSort xs) = set xs"
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    54
by(simp add: set_via_multiset)
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    55
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    56
lemma sorted_quickSort: "sorted (op <=) (quickSort xs)"
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    57
apply (induct xs rule: quickSort.induct)
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    58
 apply simp
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    59
apply simp
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    60
apply(blast intro: linorder_linear[THEN disjE] order_trans)
2af7b94892ce Turned into Isar theories.
nipkow
parents: 8590
diff changeset
    61
done
8590
89675b444abe added an order-sorted version of quickSort
paulson
parents: 8524
diff changeset
    62
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    63
end