src/HOL/Library/Quicksort.thy
author haftmann
Wed, 20 May 2009 15:35:12 +0200
changeset 31212 a94aea0cef76
parent 30738 0842e906300c
child 37075 a680ce27aa56
permissions -rw-r--r--
eliminated case input syntax on bits
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30738
0842e906300c normalized imports
haftmann
parents: 28041
diff changeset
     1
(*  Author:     Tobias Nipkow
24615
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     2
    Copyright   1994 TU Muenchen
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     3
*)
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     4
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     5
header{*Quicksort*}
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     6
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     7
theory Quicksort
30738
0842e906300c normalized imports
haftmann
parents: 28041
diff changeset
     8
imports Main Multiset
24615
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
     9
begin
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    10
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    11
context linorder
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    12
begin
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    13
28041
f496e9f343b7 quicksort: function -> fun
krauss
parents: 27682
diff changeset
    14
fun quicksort :: "'a list \<Rightarrow> 'a list" where
24615
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    15
"quicksort []     = []" |
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24615
diff changeset
    16
"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
    17
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    18
lemma quicksort_permutes [simp]:
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    19
  "multiset_of (quicksort xs) = multiset_of xs"
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    20
by (induct xs rule: quicksort.induct) (auto simp: union_ac)
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    21
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    22
lemma set_quicksort [simp]: "set (quicksort xs) = set xs"
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    23
by(simp add: set_count_greater_0)
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    24
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    25
lemma sorted_quicksort: "sorted(quicksort xs)"
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    26
apply (induct xs rule: quicksort.induct)
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    27
 apply simp
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    28
apply (simp add:sorted_Cons sorted_append not_le less_imp_le)
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    29
apply (metis leD le_cases le_less_trans)
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    30
done
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    31
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    32
end
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    33
17dbd993293d Added function package to PreList
nipkow
parents:
diff changeset
    34
end