| author | krauss | 
| Tue, 28 Aug 2007 23:53:07 +0200 | |
| changeset 24466 | 619f78b717cb | 
| parent 21404 | eb85850d3eb7 | 
| child 41413 | 64cd30d6b0b8 | 
| permissions | -rw-r--r-- | 
| 1476 | 1 | (* Title: HOL/ex/sorting.thy | 
| 969 | 2 | ID: $Id$ | 
| 1476 | 3 | Author: Tobias Nipkow | 
| 969 | 4 | Copyright 1994 TU Muenchen | 
| 5 | *) | |
| 6 | ||
| 15815 | 7 | header{*Sorting: Basic Theory*}
 | 
| 8 | ||
| 9 | theory Sorting | |
| 10 | imports Main Multiset | |
| 11 | begin | |
| 13159 | 12 | |
| 969 | 13 | consts | 
| 13159 | 14 |   sorted1:: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool"
 | 
| 15 |   sorted :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool"
 | |
| 969 | 16 | |
| 5184 | 17 | primrec | 
| 2517 | 18 | "sorted1 le [] = True" | 
| 19 | "sorted1 le (x#xs) = ((case xs of [] => True | y#ys => le x y) & | |
| 20 | sorted1 le xs)" | |
| 969 | 21 | |
| 5184 | 22 | primrec | 
| 2517 | 23 | "sorted le [] = True" | 
| 15815 | 24 | "sorted le (x#xs) = ((\<forall>y \<in> set xs. le x y) & sorted le xs)" | 
| 969 | 25 | |
| 26 | ||
| 19736 | 27 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
changeset | 28 |   total  :: "('a \<Rightarrow> 'a \<Rightarrow> bool) => bool" where
 | 
| 19736 | 29 | "total r = (\<forall>x y. r x y | r y x)" | 
| 8415 | 30 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
changeset | 31 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
changeset | 32 |   transf :: "('a \<Rightarrow> 'a \<Rightarrow> bool) => bool" where
 | 
| 19736 | 33 | "transf f = (\<forall>x y z. f x y & f y z --> f x z)" | 
| 8415 | 34 | |
| 13159 | 35 | |
| 36 | ||
| 37 | (* Equivalence of two definitions of `sorted' *) | |
| 38 | ||
| 15815 | 39 | lemma sorted1_is_sorted: "transf(le) ==> sorted1 le xs = sorted le xs"; | 
| 13159 | 40 | apply(induct xs) | 
| 41 | apply simp | |
| 42 | apply(simp split: list.split) | |
| 43 | apply(unfold transf_def); | |
| 44 | apply(blast) | |
| 45 | done | |
| 46 | ||
| 15815 | 47 | lemma sorted_append [simp]: | 
| 19736 | 48 | "sorted le (xs@ys) = | 
| 49 | (sorted le xs & sorted le ys & (\<forall>x \<in> set xs. \<forall>y \<in> set ys. le x y))" | |
| 50 | by (induct xs) auto | |
| 13159 | 51 | |
| 969 | 52 | end |