author | wenzelm |
Tue, 24 Jan 2006 00:43:31 +0100 | |
changeset 18768 | 6e97b57cdcba |
parent 17456 | bcf7544875b2 |
permissions | -rw-r--r-- |
17456 | 1 |
(* Title: CCL/ex/List.ML |
0 | 2 |
ID: $Id$ |
1459 | 3 |
Author: Martin Coen, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1993 University of Cambridge |
5 |
*) |
|
6 |
||
7 |
val list_defs = [map_def,comp_def,append_def,filter_def,flat_def, |
|
8 |
insert_def,isort_def,partition_def,qsort_def]; |
|
9 |
||
10 |
(****) |
|
11 |
||
17456 | 12 |
val listBs = map (fn s=>prove_goalw (the_context ()) list_defs s (fn _ => [simp_tac term_ss 1])) |
3837 | 13 |
["(f o g) = (%a. f(g(a)))", |
0 | 14 |
"(f o g)(a) = f(g(a))", |
15 |
"map(f,[]) = []", |
|
290
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
clasohm
parents:
8
diff
changeset
|
16 |
"map(f,x$xs) = f(x)$map(f,xs)", |
0 | 17 |
"[] @ m = m", |
290
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
clasohm
parents:
8
diff
changeset
|
18 |
"x$xs @ m = x$(xs @ m)", |
0 | 19 |
"filter(f,[]) = []", |
290
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
clasohm
parents:
8
diff
changeset
|
20 |
"filter(f,x$xs) = if f`x then x$filter(f,xs) else filter(f,xs)", |
0 | 21 |
"flat([]) = []", |
290
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
clasohm
parents:
8
diff
changeset
|
22 |
"flat(x$xs) = x @ flat(xs)", |
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
clasohm
parents:
8
diff
changeset
|
23 |
"insert(f,a,[]) = a$[]", |
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
clasohm
parents:
8
diff
changeset
|
24 |
"insert(f,a,x$xs) = if f`a`x then a$x$xs else x$insert(f,a,xs)"]; |
0 | 25 |
|
8
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
26 |
val list_ss = nat_ss addsimps listBs; |
0 | 27 |
|
28 |
(****) |
|
29 |
||
17456 | 30 |
val [prem] = goal (the_context ()) "n:Nat ==> map(f) ^ n ` [] = []"; |
1459 | 31 |
by (rtac (prem RS Nat_ind) 1); |
8
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
32 |
by (ALLGOALS (asm_simp_tac list_ss)); |
757 | 33 |
qed "nmapBnil"; |
0 | 34 |
|
17456 | 35 |
val [prem] = goal (the_context ()) "n:Nat ==> map(f)^n`(x$xs) = (f^n`x)$(map(f)^n`xs)"; |
1459 | 36 |
by (rtac (prem RS Nat_ind) 1); |
8
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
37 |
by (ALLGOALS (asm_simp_tac list_ss)); |
757 | 38 |
qed "nmapBcons"; |
0 | 39 |
|
40 |
(***) |
|
41 |
||
17456 | 42 |
val prems = goalw (the_context ()) [map_def] |
3837 | 43 |
"[| !!x. x:A==>f(x):B; l : List(A) |] ==> map(f,l) : List(B)"; |
0 | 44 |
by (typechk_tac prems 1); |
757 | 45 |
qed "mapT"; |
0 | 46 |
|
17456 | 47 |
val prems = goalw (the_context ()) [append_def] |
0 | 48 |
"[| l : List(A); m : List(A) |] ==> l @ m : List(A)"; |
49 |
by (typechk_tac prems 1); |
|
757 | 50 |
qed "appendT"; |
0 | 51 |
|
17456 | 52 |
val prems = goal (the_context ()) |
0 | 53 |
"[| l : {l:List(A). m : {m:List(A).P(l @ m)}} |] ==> l @ m : {x:List(A). P(x)}"; |
54 |
by (cut_facts_tac prems 1); |
|
55 |
by (fast_tac (set_cs addSIs [SubtypeI,appendT] addSEs [SubtypeE]) 1); |
|
757 | 56 |
qed "appendTS"; |
0 | 57 |
|
17456 | 58 |
val prems = goalw (the_context ()) [filter_def] |
0 | 59 |
"[| f:A->Bool; l : List(A) |] ==> filter(f,l) : List(A)"; |
60 |
by (typechk_tac prems 1); |
|
757 | 61 |
qed "filterT"; |
0 | 62 |
|
17456 | 63 |
val prems = goalw (the_context ()) [flat_def] |
0 | 64 |
"l : List(List(A)) ==> flat(l) : List(A)"; |
65 |
by (typechk_tac (appendT::prems) 1); |
|
757 | 66 |
qed "flatT"; |
0 | 67 |
|
17456 | 68 |
val prems = goalw (the_context ()) [insert_def] |
0 | 69 |
"[| f : A->A->Bool; a:A; l : List(A) |] ==> insert(f,a,l) : List(A)"; |
70 |
by (typechk_tac prems 1); |
|
757 | 71 |
qed "insertT"; |
0 | 72 |
|
17456 | 73 |
val prems = goal (the_context ()) |
0 | 74 |
"[| f : {f:A->A->Bool. a : {a:A. l : {l:List(A).P(insert(f,a,l))}}} |] ==> \ |
75 |
\ insert(f,a,l) : {x:List(A). P(x)}"; |
|
76 |
by (cut_facts_tac prems 1); |
|
77 |
by (fast_tac (set_cs addSIs [SubtypeI,insertT] addSEs [SubtypeE]) 1); |
|
757 | 78 |
qed "insertTS"; |
0 | 79 |
|
17456 | 80 |
val prems = goalw (the_context ()) [partition_def] |
0 | 81 |
"[| f:A->Bool; l : List(A) |] ==> partition(f,l) : List(A)*List(A)"; |
82 |
by (typechk_tac prems 1); |
|
83 |
by clean_ccs_tac; |
|
1459 | 84 |
by (rtac (ListPRI RS wfstI RS (ListPR_wf RS wmap_wf RS wfI)) 2); |
85 |
by (rtac (ListPRI RS wfstI RS (ListPR_wf RS wmap_wf RS wfI)) 1); |
|
0 | 86 |
by (REPEAT (atac 1)); |
757 | 87 |
qed "partitionT"; |
0 | 88 |
|
89 |
(*** Correctness Conditions for Insertion Sort ***) |
|
90 |
||
91 |
||
17456 | 92 |
val prems = goalw (the_context ()) [isort_def] |
0 | 93 |
"f:A->A->Bool ==> isort(f) : PROD l:List(A).{x: List(A). Ord(f,x) & Perm(x,l)}"; |
94 |
by (gen_ccs_tac ([insertTS,insertT]@prems) 1); |
|
95 |
||
96 |
||
97 |
(*** Correctness Conditions for Quick Sort ***) |
|
98 |
||
17456 | 99 |
val prems = goalw (the_context ()) [qsort_def] |
0 | 100 |
"f:A->A->Bool ==> qsort(f) : PROD l:List(A).{x: List(A). Ord(f,x) & Perm(x,l)}"; |
101 |
by (gen_ccs_tac ([partitionT,appendTS,appendT]@prems) 1); |