author | clasohm |
Wed, 02 Mar 1994 12:26:55 +0100 | |
changeset 48 | 21291189b51e |
parent 42 | 87f6e8b93221 |
child 49 | 9f35f2744fa8 |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: HOL/list |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1993 University of Cambridge |
|
5 |
||
6 |
Definition of type 'a list by a least fixed point |
|
7 |
||
8 |
We use List(A) == lfp(%Z. {NUMB(0)} <+> A <*> Z) |
|
9 |
and not List == lfp(%Z. {NUMB(0)} <+> range(Leaf) <*> Z) |
|
10 |
so that List can serve as a "functor" for defining other recursive types |
|
11 |
*) |
|
12 |
||
13 |
List = Sexp + |
|
14 |
||
15 |
types |
|
16 |
list 1 |
|
17 |
||
18 |
arities |
|
19 |
list :: (term) term |
|
20 |
||
21 |
||
22 |
consts |
|
23 |
||
24 |
List_Fun :: "['a node set set, 'a node set set] => 'a node set set" |
|
25 |
List :: "'a node set set => 'a node set set" |
|
26 |
Rep_List :: "'a list => 'a node set" |
|
27 |
Abs_List :: "'a node set => 'a list" |
|
28 |
NIL :: "'a node set" |
|
29 |
CONS :: "['a node set, 'a node set] => 'a node set" |
|
30 |
Nil :: "'a list" |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
31 |
"#" :: "['a, 'a list] => 'a list" (infixr 65) |
0 | 32 |
List_case :: "['a node set, 'b, ['a node set, 'a node set]=>'b] => 'b" |
33 |
List_rec :: "['a node set, 'b, ['a node set, 'a node set, 'b]=>'b] => 'b" |
|
34 |
list_rec :: "['a list, 'b, ['a, 'a list, 'b]=>'b] => 'b" |
|
35 |
Rep_map :: "('b => 'a node set) => ('b list => 'a node set)" |
|
36 |
Abs_map :: "('a node set => 'b) => 'a node set => 'b list" |
|
37 |
null :: "'a list => bool" |
|
38 |
hd :: "'a list => 'a" |
|
40 | 39 |
tl,ttl :: "'a list => 'a list" |
34 | 40 |
mem :: "['a, 'a list] => bool" (infixl 55) |
0 | 41 |
list_all :: "('a => bool) => ('a list => bool)" |
42 |
map :: "('a=>'b) => ('a list => 'b list)" |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
43 |
"@" :: "['a list, 'a list] => 'a list" (infixl 65) |
0 | 44 |
list_case :: "['a list, 'b, ['a, 'a list]=>'b] => 'b" |
34 | 45 |
filter :: "['a => bool, 'a list] => 'a list" |
0 | 46 |
|
47 |
(* List Enumeration *) |
|
48 |
||
49 |
"[]" :: "'a list" ("[]") |
|
50 |
"@List" :: "args => 'a list" ("[(_)]") |
|
51 |
||
34 | 52 |
(* Special syntax for list_all and filter *) |
39 | 53 |
"@Alls" :: "[idt, 'a list, bool] => bool" ("(2Alls _:_./ _)" 10) |
54 |
"@filter" :: "[idt, 'a list, bool] => 'a list" ("(1[_:_ ./ _])") |
|
0 | 55 |
|
56 |
translations |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
57 |
"[x, xs]" == "x#[xs]" |
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
58 |
"[x]" == "x#[]" |
0 | 59 |
"[]" == "Nil" |
60 |
||
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
61 |
"case xs of Nil => a | y#ys => b" == "list_case(xs,a,%y ys.b)" |
42 | 62 |
|
34 | 63 |
"[x:xs . P]" == "filter(%x.P,xs)" |
64 |
"Alls x:xs.P" == "list_all(%x.P,xs)" |
|
0 | 65 |
|
66 |
rules |
|
67 |
||
68 |
List_Fun_def "List_Fun(A) == (%Z. {Numb(0)} <+> A <*> Z)" |
|
69 |
List_def "List(A) == lfp(List_Fun(A))" |
|
70 |
||
71 |
(* Faking a Type Definition ... *) |
|
72 |
||
73 |
Rep_List "Rep_List(xs): List(range(Leaf))" |
|
74 |
Rep_List_inverse "Abs_List(Rep_List(xs)) = xs" |
|
75 |
Abs_List_inverse "M: List(range(Leaf)) ==> Rep_List(Abs_List(M)) = M" |
|
76 |
||
77 |
(* Defining the Concrete Constructors *) |
|
78 |
||
79 |
NIL_def "NIL == In0(Numb(0))" |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
80 |
CONS_def "CONS(M, N) == In1(M $ N)" |
0 | 81 |
|
82 |
(* Defining the Abstract Constructors *) |
|
83 |
||
84 |
Nil_def "Nil == Abs_List(NIL)" |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
85 |
Cons_def "x#xs == Abs_List(CONS(Leaf(x), Rep_List(xs)))" |
0 | 86 |
|
87 |
List_case_def "List_case(M, c, d) == Case(M, %x.c, %u. Split(u, %x y.d(x, y)))" |
|
88 |
||
89 |
(* List Recursion -- the trancl is Essential; see list.ML *) |
|
90 |
||
91 |
List_rec_def |
|
92 |
"List_rec(M, c, d) == wfrec(trancl(pred_Sexp), M, \ |
|
93 |
\ %z g. List_case(z, c, %x y. d(x, y, g(y))))" |
|
94 |
||
95 |
list_rec_def |
|
96 |
"list_rec(l, c, d) == \ |
|
97 |
\ List_rec(Rep_List(l), c, %x y r. d(Inv(Leaf, x), Abs_List(y), r))" |
|
98 |
||
99 |
(* Generalized Map Functionals *) |
|
100 |
||
101 |
Rep_map_def |
|
102 |
"Rep_map(f, xs) == list_rec(xs, NIL, %x l r. CONS(f(x), r))" |
|
103 |
Abs_map_def |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
104 |
"Abs_map(g, M) == List_rec(M, Nil, %N L r. g(N)#r)" |
0 | 105 |
|
106 |
null_def "null(xs) == list_rec(xs, True, %x xs r.False)" |
|
107 |
hd_def "hd(xs) == list_rec(xs, @x.True, %x xs r.x)" |
|
108 |
tl_def "tl(xs) == list_rec(xs, @xs.True, %x xs r.xs)" |
|
40 | 109 |
(* a total version of tl: *) |
110 |
ttl_def "ttl(xs) == list_rec(xs, [], %x xs r.xs)" |
|
34 | 111 |
mem_def "x mem xs == \ |
112 |
\ list_rec(xs, False, %y ys r. if(y=x, True, r))" |
|
0 | 113 |
list_all_def "list_all(P, xs) == list_rec(xs, True, %x l r. P(x) & r)" |
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
114 |
map_def "map(f, xs) == list_rec(xs, [], %x l r. f(x)#r)" |
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
115 |
append_def "xs@ys == list_rec(xs, ys, %x l r. x#r)" |
34 | 116 |
filter_def "filter(P,xs) == \ |
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
117 |
\ list_rec(xs, [], %x xs r. if(P(x), x#r, r))" |
0 | 118 |
list_case_def "list_case(xs, a, f) == list_rec(xs, a, %x xs r.f(x, xs))" |
119 |
||
120 |
end |