| author | oheimb | 
| Thu, 24 Sep 1998 15:36:16 +0200 | |
| changeset 5548 | 5cd3396802f5 | 
| parent 5278 | a903b66822e2 | 
| child 6434 | f2a2a40e56c8 | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: HOL/ex/qsort.ML | 
| 969 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Tobias Nipkow | 
| 969 | 4 | Copyright 1994 TU Muenchen | 
| 5 | ||
| 6 | Two verifications of Quicksort | |
| 7 | *) | |
| 8 | ||
| 2526 | 9 | Addsimps [ball_Un]; | 
| 2511 | 10 | |
| 11 | (* Towards a proof of qsort_ind | |
| 12 | ||
| 13 | A short proof of Konrad's wf_minimal in WF1.ML: | |
| 14 | ||
| 15 | val [p1] = goalw WF.thy [wf_def] | |
| 16 | "wf(R) ==> (!Q. (? x. Q x) --> (? min. Q min & (!b. (b,min):R --> (~ Q b))))"; | |
| 17 | by (rtac allI 1); | |
| 4423 | 18 | by (cut_inst_tac[("x","%u.~Q u")](p1 RS spec) 1);
 | 
| 2511 | 19 | by (fast_tac HOL_cs 1); | 
| 20 | val wf_minimal = result(); | |
| 21 | ||
| 22 | *) | |
| 23 | ||
| 24 | ||
| 1266 | 25 | Addsimps ([Qsort.qsort_Nil,Qsort.qsort_Cons]@conj_comms); | 
| 969 | 26 | |
| 5069 | 27 | Goal "!x. mset (qsort le xs) x = mset xs x"; | 
| 2031 | 28 | by (res_inst_tac[("xs","xs"),("p","le")]Qsort.qsort_ind 1);
 | 
| 4686 | 29 | by (ALLGOALS Asm_simp_tac); | 
| 2511 | 30 | qed "qsort_permutes"; | 
| 969 | 31 | |
| 5069 | 32 | Goal "set(qsort le xs) = set xs"; | 
| 4423 | 33 | by (simp_tac (simpset() addsimps [set_via_mset,qsort_permutes]) 1); | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3465diff
changeset | 34 | qed "set_qsort"; | 
| 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3465diff
changeset | 35 | Addsimps [set_qsort]; | 
| 969 | 36 | |
| 2511 | 37 | goal List.thy | 
| 3842 | 38 | "(!x:set[x:xs. P(x)].Q(x)) = (!x:set xs. P(x)-->Q(x))"; | 
| 5184 | 39 | by (induct_tac "xs" 1); | 
| 4686 | 40 | by (ALLGOALS Asm_simp_tac); | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3465diff
changeset | 41 | qed"Ball_set_filter"; | 
| 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3465diff
changeset | 42 | Addsimps [Ball_set_filter]; | 
| 969 | 43 | |
| 5278 | 44 | Goal "sorted le (xs@ys) = (sorted le xs & sorted le ys & \ | 
| 3465 | 45 | \ (!x:set xs. !y:set ys. le x y))"; | 
| 5184 | 46 | by (induct_tac "xs" 1); | 
| 2511 | 47 | by (ALLGOALS Asm_simp_tac); | 
| 48 | qed "sorted_append"; | |
| 49 | Addsimps [sorted_append]; | |
| 1673 
d22110ddd0af
repaired critical proofs depending on the order inside non-confluent SimpSets
 oheimb parents: 
1465diff
changeset | 50 | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5069diff
changeset | 51 | Goal "[| total(le); transf(le) |] ==> sorted le (qsort le xs)"; | 
| 2031 | 52 | by (res_inst_tac[("xs","xs"),("p","le")]Qsort.qsort_ind 1);
 | 
| 2511 | 53 | by (ALLGOALS Asm_simp_tac); | 
| 2031 | 54 | by (rewrite_goals_tac [Sorting.total_def,Sorting.transf_def]); | 
| 55 | by (Fast_tac 1); | |
| 2511 | 56 | qed "sorted_qsort"; |