46
|
1 |
(* Title: HOL/ex/insort.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Tobias Nipkow
|
|
4 |
Copyright 1994 TU Muenchen
|
|
5 |
|
|
6 |
Correctness proof of insertion sort.
|
|
7 |
*)
|
|
8 |
|
|
9 |
val insort_ss = sorting_ss addsimps
|
|
10 |
[InSort.ins_Nil,InSort.ins_Cons,InSort.insort_Nil,InSort.insort_Cons];
|
|
11 |
|
|
12 |
goalw InSort.thy [Sorting.total_def]
|
|
13 |
"!!f. [| total(f); ~f(x,y) |] ==> f(y,x)";
|
|
14 |
by(fast_tac HOL_cs 1);
|
|
15 |
val totalD = result();
|
|
16 |
|
|
17 |
goalw InSort.thy [Sorting.transf_def]
|
|
18 |
"!!f. [| transf(f); f(b,a) |] ==> !x. f(a,x) --> f(b,x)";
|
|
19 |
by(fast_tac HOL_cs 1);
|
|
20 |
val transfD = result();
|
|
21 |
|
|
22 |
goal InSort.thy "list_all(p,ins(f,x,xs)) = (list_all(p,xs) & p(x))";
|
|
23 |
by(list_ind_tac "xs" 1);
|
|
24 |
by(asm_simp_tac insort_ss 1);
|
|
25 |
by(asm_simp_tac (insort_ss setloop (split_tac [expand_if])) 1);
|
|
26 |
by(fast_tac HOL_cs 1);
|
|
27 |
val insort_ss = insort_ss addsimps [result()];
|
|
28 |
|
|
29 |
goal InSort.thy "(!x. p(x) --> q(x)) --> list_all(p,xs) --> list_all(q,xs)";
|
|
30 |
by(list_ind_tac "xs" 1);
|
|
31 |
by(ALLGOALS(asm_simp_tac (insort_ss setloop (split_tac [expand_if]))));
|
|
32 |
val list_all_imp = result();
|
|
33 |
|
|
34 |
val prems = goal InSort.thy
|
|
35 |
"[| total(f); transf(f) |] ==> sorted(f,ins(f,x,xs)) = sorted(f,xs)";
|
|
36 |
by(list_ind_tac "xs" 1);
|
|
37 |
by(ALLGOALS(asm_simp_tac (insort_ss setloop (split_tac [expand_if]))));
|
|
38 |
by(cut_facts_tac prems 1);
|
|
39 |
by(cut_inst_tac [("p","f(xa)"),("q","f(x)")] list_all_imp 1);
|
|
40 |
by(fast_tac (HOL_cs addDs [totalD,transfD]) 1);
|
|
41 |
val insort_ss = insort_ss addsimps [result()];
|
|
42 |
|
|
43 |
goal InSort.thy "!!f. [| total(f); transf(f) |] ==> sorted(f,insort(f,xs))";
|
|
44 |
by(list_ind_tac "xs" 1);
|
|
45 |
by(ALLGOALS(asm_simp_tac insort_ss));
|
|
46 |
result();
|