0
|
1 |
(* Title: CCL/ex/list
|
|
2 |
ID: $Id$
|
|
3 |
Author: Martin Coen, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1993 University of Cambridge
|
|
5 |
|
|
6 |
For list.thy.
|
|
7 |
*)
|
|
8 |
|
|
9 |
open List;
|
|
10 |
|
|
11 |
val list_defs = [map_def,comp_def,append_def,filter_def,flat_def,
|
|
12 |
insert_def,isort_def,partition_def,qsort_def];
|
|
13 |
|
|
14 |
(****)
|
|
15 |
|
|
16 |
val listBs = map (fn s=>prove_goalw List.thy list_defs s (fn _ => [SIMP_TAC term_ss 1]))
|
|
17 |
["(f o g) = (%a.f(g(a)))",
|
|
18 |
"(f o g)(a) = f(g(a))",
|
|
19 |
"map(f,[]) = []",
|
|
20 |
"map(f,x.xs) = f(x).map(f,xs)",
|
|
21 |
"[] @ m = m",
|
|
22 |
"x.xs @ m = x.(xs @ m)",
|
|
23 |
"filter(f,[]) = []",
|
|
24 |
"filter(f,x.xs) = if f`x then x.filter(f,xs) else filter(f,xs)",
|
|
25 |
"flat([]) = []",
|
|
26 |
"flat(x.xs) = x @ flat(xs)",
|
|
27 |
"insert(f,a,[]) = a.[]",
|
|
28 |
"insert(f,a,x.xs) = if f`a`x then a.x.xs else x.insert(f,a,xs)"];
|
|
29 |
|
|
30 |
val list_congs = ccl_mk_congs List.thy ["map","op @","filter","flat","insert","napply"];
|
|
31 |
|
|
32 |
val list_ss = nat_ss addrews listBs addcongs list_congs;
|
|
33 |
|
|
34 |
(****)
|
|
35 |
|
|
36 |
val [prem] = goal List.thy "n:Nat ==> map(f) ^ n ` [] = []";
|
|
37 |
br (prem RS Nat_ind) 1;
|
|
38 |
by (ALLGOALS (ASM_SIMP_TAC list_ss));
|
|
39 |
val nmapBnil = result();
|
|
40 |
|
|
41 |
val [prem] = goal List.thy "n:Nat ==> map(f)^n`(x.xs) = f^n`x.map(f)^n`xs";
|
|
42 |
br (prem RS Nat_ind) 1;
|
|
43 |
by (ALLGOALS (ASM_SIMP_TAC list_ss));
|
|
44 |
val nmapBcons = result();
|
|
45 |
|
|
46 |
(***)
|
|
47 |
|
|
48 |
val prems = goalw List.thy [map_def]
|
|
49 |
"[| !!x.x:A==>f(x):B; l : List(A) |] ==> map(f,l) : List(B)";
|
|
50 |
by (typechk_tac prems 1);
|
|
51 |
val mapT = result();
|
|
52 |
|
|
53 |
val prems = goalw List.thy [append_def]
|
|
54 |
"[| l : List(A); m : List(A) |] ==> l @ m : List(A)";
|
|
55 |
by (typechk_tac prems 1);
|
|
56 |
val appendT = result();
|
|
57 |
|
|
58 |
val prems = goal List.thy
|
|
59 |
"[| l : {l:List(A). m : {m:List(A).P(l @ m)}} |] ==> l @ m : {x:List(A). P(x)}";
|
|
60 |
by (cut_facts_tac prems 1);
|
|
61 |
by (fast_tac (set_cs addSIs [SubtypeI,appendT] addSEs [SubtypeE]) 1);
|
|
62 |
val appendTS = result();
|
|
63 |
|
|
64 |
val prems = goalw List.thy [filter_def]
|
|
65 |
"[| f:A->Bool; l : List(A) |] ==> filter(f,l) : List(A)";
|
|
66 |
by (typechk_tac prems 1);
|
|
67 |
val filterT = result();
|
|
68 |
|
|
69 |
val prems = goalw List.thy [flat_def]
|
|
70 |
"l : List(List(A)) ==> flat(l) : List(A)";
|
|
71 |
by (typechk_tac (appendT::prems) 1);
|
|
72 |
val flatT = result();
|
|
73 |
|
|
74 |
val prems = goalw List.thy [insert_def]
|
|
75 |
"[| f : A->A->Bool; a:A; l : List(A) |] ==> insert(f,a,l) : List(A)";
|
|
76 |
by (typechk_tac prems 1);
|
|
77 |
val insertT = result();
|
|
78 |
|
|
79 |
val prems = goal List.thy
|
|
80 |
"[| f : {f:A->A->Bool. a : {a:A. l : {l:List(A).P(insert(f,a,l))}}} |] ==> \
|
|
81 |
\ insert(f,a,l) : {x:List(A). P(x)}";
|
|
82 |
by (cut_facts_tac prems 1);
|
|
83 |
by (fast_tac (set_cs addSIs [SubtypeI,insertT] addSEs [SubtypeE]) 1);
|
|
84 |
val insertTS = result();
|
|
85 |
|
|
86 |
val prems = goalw List.thy [partition_def]
|
|
87 |
"[| f:A->Bool; l : List(A) |] ==> partition(f,l) : List(A)*List(A)";
|
|
88 |
by (typechk_tac prems 1);
|
|
89 |
by clean_ccs_tac;
|
|
90 |
br (ListPRI RS wfstI RS (ListPR_wf RS wmap_wf RS wfI)) 2;
|
|
91 |
br (ListPRI RS wfstI RS (ListPR_wf RS wmap_wf RS wfI)) 1;
|
|
92 |
by (REPEAT (atac 1));
|
|
93 |
val partitionT = result();
|
|
94 |
|
|
95 |
(*** Correctness Conditions for Insertion Sort ***)
|
|
96 |
|
|
97 |
|
|
98 |
val prems = goalw List.thy [isort_def]
|
|
99 |
"f:A->A->Bool ==> isort(f) : PROD l:List(A).{x: List(A). Ord(f,x) & Perm(x,l)}";
|
|
100 |
by (gen_ccs_tac ([insertTS,insertT]@prems) 1);
|
|
101 |
|
|
102 |
|
|
103 |
(*** Correctness Conditions for Quick Sort ***)
|
|
104 |
|
|
105 |
val prems = goalw List.thy [qsort_def]
|
|
106 |
"f:A->A->Bool ==> qsort(f) : PROD l:List(A).{x: List(A). Ord(f,x) & Perm(x,l)}";
|
|
107 |
by (gen_ccs_tac ([partitionT,appendTS,appendT]@prems) 1);
|
|
108 |
|