src/HOL/Library/Quicksort.thy
author huffman
Wed, 18 Feb 2009 19:32:26 -0800
changeset 29985 57975b45ab70
parent 28041 f496e9f343b7
child 30738 0842e906300c
permissions -rw-r--r--
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24615
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     1
(*  ID:         $Id$
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     2
    Author:     Tobias Nipkow
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     3
    Copyright   1994 TU Muenchen
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     4
*)
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     5
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     6
header{*Quicksort*}
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     7
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     8
theory Quicksort
27368
9f90ac19e32b established Plain theory and image
haftmann
parents: 25062
diff changeset
     9
imports Plain Multiset
24615
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    10
begin
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    11
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    12
context linorder
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    13
begin
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    14
28041
f496e9f343b7 quicksort: function -> fun
krauss
parents: 27682
diff changeset
    15
fun quicksort :: "'a list \<Rightarrow> 'a list" where
24615
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    16
"quicksort []     = []" |
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24615
diff changeset
    17
"quicksort (x#xs) = quicksort([y\<leftarrow>xs. ~ x\<le>y]) @ [x] @ quicksort([y\<leftarrow>xs. x\<le>y])"
24615
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    18
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    19
lemma quicksort_permutes [simp]:
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    20
  "multiset_of (quicksort xs) = multiset_of xs"
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    21
by (induct xs rule: quicksort.induct) (auto simp: union_ac)
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    22
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    23
lemma set_quicksort [simp]: "set (quicksort xs) = set xs"
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    24
by(simp add: set_count_greater_0)
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    25
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    26
lemma sorted_quicksort: "sorted(quicksort xs)"
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    27
apply (induct xs rule: quicksort.induct)
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    28
 apply simp
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    29
apply (simp add:sorted_Cons sorted_append not_le less_imp_le)
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    30
apply (metis leD le_cases le_less_trans)
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    31
done
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    32
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    33
end
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    34
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    35
end