src/HOL/ex/Qsort.ML
author wenzelm
Fri, 10 Oct 1997 19:02:28 +0200
changeset 3842 b55686a7b22c
parent 3647 a64c8fbcd98f
child 3919 c036caebfc75
permissions -rw-r--r--
fixed dots;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
     1
(*  Title:      HOL/ex/qsort.ML
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
     3
    Author:     Tobias Nipkow
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1994 TU Muenchen
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     5
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     6
Two verifications of Quicksort
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     7
*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     8
2526
43650141d67d Ball_Un -> ball_Un
nipkow
parents: 2517
diff changeset
     9
Addsimps [ball_Un];
2511
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    10
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    11
(* Towards a proof of qsort_ind
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    12
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    13
A short proof of Konrad's wf_minimal in WF1.ML:
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    14
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    15
val [p1] = goalw WF.thy [wf_def]
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    16
"wf(R) ==> (!Q. (? x. Q x) --> (? min. Q min & (!b. (b,min):R --> (~ Q b))))";
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    17
by (rtac allI 1);
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    18
by(cut_inst_tac[("x","%u.~Q u")](p1 RS spec) 1);
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    19
by (fast_tac HOL_cs 1);
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    20
val wf_minimal = result();
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    21
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    22
*)
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    23
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    24
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 969
diff changeset
    25
Addsimps ([Qsort.qsort_Nil,Qsort.qsort_Cons]@conj_comms);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    26
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    27
goal Qsort.thy "!x. mset (qsort le xs) x = mset xs x";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1820
diff changeset
    28
by (res_inst_tac[("xs","xs"),("p","le")]Qsort.qsort_ind 1);
03a843f0f447 Ran expandshort
paulson
parents: 1820
diff changeset
    29
by (ALLGOALS(asm_simp_tac (!simpset setloop (split_tac [expand_if]))));
2511
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    30
qed "qsort_permutes";
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    31
3465
e85c24717cad set_of_list -> set
nipkow
parents: 2526
diff changeset
    32
goal Qsort.thy "set(qsort le xs) = set xs";
3647
a64c8fbcd98f Renamed theorems of the form set_of_list_XXX to set_XXX
paulson
parents: 3465
diff changeset
    33
by(simp_tac (!simpset addsimps [set_via_mset,qsort_permutes]) 1);
a64c8fbcd98f Renamed theorems of the form set_of_list_XXX to set_XXX
paulson
parents: 3465
diff changeset
    34
qed "set_qsort";
a64c8fbcd98f Renamed theorems of the form set_of_list_XXX to set_XXX
paulson
parents: 3465
diff changeset
    35
Addsimps [set_qsort];
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    36
2511
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    37
goal List.thy
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3647
diff changeset
    38
  "(!x:set[x:xs. P(x)].Q(x)) = (!x:set xs. P(x)-->Q(x))";
2511
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    39
by (list.induct_tac "xs" 1);
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    40
by (ALLGOALS(asm_simp_tac (!simpset setloop (split_tac [expand_if]))));
3647
a64c8fbcd98f Renamed theorems of the form set_of_list_XXX to set_XXX
paulson
parents: 3465
diff changeset
    41
qed"Ball_set_filter";
a64c8fbcd98f Renamed theorems of the form set_of_list_XXX to set_XXX
paulson
parents: 3465
diff changeset
    42
Addsimps [Ball_set_filter];
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    43
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    44
goal Qsort.thy
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    45
 "sorted le (xs@ys) = (sorted le xs & sorted le ys & \
3465
e85c24717cad set_of_list -> set
nipkow
parents: 2526
diff changeset
    46
\                     (!x:set xs. !y:set ys. le x y))";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1820
diff changeset
    47
by (list.induct_tac "xs" 1);
2511
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    48
by (ALLGOALS Asm_simp_tac);
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    49
qed "sorted_append";
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    50
Addsimps [sorted_append];
1673
d22110ddd0af repaired critical proofs depending on the order inside non-confluent SimpSets
oheimb
parents: 1465
diff changeset
    51
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    52
goal Qsort.thy 
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    53
 "!!le. [| total(le); transf(le) |] ==>  sorted le (qsort le xs)";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1820
diff changeset
    54
by (res_inst_tac[("xs","xs"),("p","le")]Qsort.qsort_ind 1);
2511
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    55
by (ALLGOALS Asm_simp_tac);
2031
03a843f0f447 Ran expandshort
paulson
parents: 1820
diff changeset
    56
by (rewrite_goals_tac [Sorting.total_def,Sorting.transf_def]);
03a843f0f447 Ran expandshort
paulson
parents: 1820
diff changeset
    57
by (Fast_tac 1);
2511
282e9a9eae9d Got rid of Alls in favour of !x:set_of_list
nipkow
parents: 2031
diff changeset
    58
qed "sorted_qsort";