author | clasohm |
Wed, 02 Mar 1994 12:26:55 +0100 | |
changeset 48 | 21291189b51e |
parent 46 | a73f8a7784bd |
child 195 | df6b3bd14dcb |
permissions | -rw-r--r-- |
(* Title: HOL/ex/insort.thy ID: $Id$ Author: Tobias Nipkow Copyright 1994 TU Muenchen Insertion sort *) InSort = Sorting + consts ins :: "[['a,'a]=>bool, 'a, 'a list] => 'a list" insort :: "[['a,'a]=>bool, 'a list] => 'a list" rules ins_Nil "ins(f,x,[]) = [x]" ins_Cons "ins(f,x,y#ys) = \ \ if(f(x,y), x#y#ys, y# ins(f,x,ys))" insort_Nil "insort(f,[]) = []" insort_Cons "insort(f,x#xs) = ins(f,x,insort(f,xs))" end