| 1476 |      1 | (*  Title:      HOL/ex/sorting.thy
 | 
| 969 |      2 |     ID:         $Id$
 | 
| 1476 |      3 |     Author:     Tobias Nipkow
 | 
| 969 |      4 |     Copyright   1994 TU Muenchen
 | 
|  |      5 | 
 | 
|  |      6 | Specification of sorting
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
| 13159 |      9 | theory Sorting = Main:
 | 
|  |     10 | 
 | 
| 969 |     11 | consts
 | 
| 13159 |     12 |   sorted1:: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool"
 | 
|  |     13 |   sorted :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool"
 | 
|  |     14 |   multiset   :: "'a list => ('a => nat)"
 | 
| 969 |     15 | 
 | 
| 5184 |     16 | primrec
 | 
| 2517 |     17 |   "sorted1 le [] = True"
 | 
|  |     18 |   "sorted1 le (x#xs) = ((case xs of [] => True | y#ys => le x y) &
 | 
|  |     19 |                         sorted1 le xs)"
 | 
| 969 |     20 | 
 | 
| 5184 |     21 | primrec
 | 
| 2517 |     22 |   "sorted le [] = True"
 | 
| 3465 |     23 |   "sorted le (x#xs) = ((!y:set xs. le x y) & sorted le xs)"
 | 
| 969 |     24 | 
 | 
| 5184 |     25 | primrec
 | 
| 8415 |     26 |   "multiset [] y = 0"
 | 
|  |     27 |   "multiset (x#xs) y = (if x=y then Suc(multiset xs y) else multiset xs y)"
 | 
| 969 |     28 | 
 | 
| 8415 |     29 | constdefs
 | 
| 13159 |     30 |   total  :: "('a \<Rightarrow> 'a \<Rightarrow> bool) => bool"
 | 
| 8415 |     31 |    "total r == (ALL x y. r x y | r y x)"
 | 
|  |     32 |   
 | 
| 13159 |     33 |   transf :: "('a \<Rightarrow> 'a \<Rightarrow> bool) => bool"
 | 
| 8415 |     34 |    "transf f == (ALL x y z. f x y & f y z --> f x z)"
 | 
|  |     35 | 
 | 
| 13159 |     36 | (* Some general lemmas *)
 | 
|  |     37 | 
 | 
|  |     38 | lemma multiset_append[simp]:
 | 
|  |     39 |  "multiset (xs@ys) x = multiset xs x + multiset ys x"
 | 
|  |     40 | by (induct xs, auto)
 | 
|  |     41 | 
 | 
|  |     42 | lemma multiset_compl_add[simp]:
 | 
|  |     43 |  "multiset [x:xs. ~p(x)] z + multiset [x:xs. p(x)] z = multiset xs z";
 | 
|  |     44 | by (induct xs, auto)
 | 
|  |     45 | 
 | 
|  |     46 | lemma set_via_multiset: "set xs = {x. multiset xs x ~= 0}";
 | 
|  |     47 | by (induct xs, auto)
 | 
|  |     48 | 
 | 
|  |     49 | (* Equivalence of two definitions of `sorted' *)
 | 
|  |     50 | 
 | 
|  |     51 | lemma sorted1_is_sorted:
 | 
|  |     52 |  "transf(le) ==> sorted1 le xs = sorted le xs";
 | 
|  |     53 | apply(induct xs)
 | 
|  |     54 |  apply simp
 | 
|  |     55 | apply(simp split: list.split)
 | 
|  |     56 | apply(unfold transf_def);
 | 
|  |     57 | apply(blast)
 | 
|  |     58 | done
 | 
|  |     59 | 
 | 
|  |     60 | lemma sorted_append[simp]:
 | 
|  |     61 |  "sorted le (xs@ys) = (sorted le xs \<and> sorted le ys \<and>
 | 
|  |     62 |                        (\<forall>x \<in> set xs. \<forall>y \<in> set ys. le x y))"
 | 
|  |     63 | by (induct xs, auto)
 | 
|  |     64 | 
 | 
| 969 |     65 | end
 |