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"
|
|
31 |
Cons :: "['a, 'a list] => 'a list"
|
|
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"
|
|
39 |
tl :: "'a list => 'a list"
|
|
40 |
list_all :: "('a => bool) => ('a list => bool)"
|
|
41 |
map :: "('a=>'b) => ('a list => 'b list)"
|
13
|
42 |
"@" :: "['a list, 'a list] => 'a list" (infixr 65)
|
0
|
43 |
list_case :: "['a list, 'b, ['a, 'a list]=>'b] => 'b"
|
|
44 |
|
|
45 |
(* List Enumeration *)
|
|
46 |
|
|
47 |
"[]" :: "'a list" ("[]")
|
|
48 |
"@List" :: "args => 'a list" ("[(_)]")
|
|
49 |
|
|
50 |
|
|
51 |
translations
|
|
52 |
"[x, xs]" == "Cons(x, [xs])"
|
|
53 |
"[x]" == "Cons(x, [])"
|
|
54 |
"[]" == "Nil"
|
|
55 |
|
|
56 |
|
|
57 |
rules
|
|
58 |
|
|
59 |
List_Fun_def "List_Fun(A) == (%Z. {Numb(0)} <+> A <*> Z)"
|
|
60 |
List_def "List(A) == lfp(List_Fun(A))"
|
|
61 |
|
|
62 |
(* Faking a Type Definition ... *)
|
|
63 |
|
|
64 |
Rep_List "Rep_List(xs): List(range(Leaf))"
|
|
65 |
Rep_List_inverse "Abs_List(Rep_List(xs)) = xs"
|
|
66 |
Abs_List_inverse "M: List(range(Leaf)) ==> Rep_List(Abs_List(M)) = M"
|
|
67 |
|
|
68 |
(* Defining the Concrete Constructors *)
|
|
69 |
|
|
70 |
NIL_def "NIL == In0(Numb(0))"
|
|
71 |
CONS_def "CONS(M, N) == In1(M . N)"
|
|
72 |
|
|
73 |
(* Defining the Abstract Constructors *)
|
|
74 |
|
|
75 |
Nil_def "Nil == Abs_List(NIL)"
|
|
76 |
Cons_def "Cons(x, xs) == Abs_List(CONS(Leaf(x), Rep_List(xs)))"
|
|
77 |
|
|
78 |
List_case_def "List_case(M, c, d) == Case(M, %x.c, %u. Split(u, %x y.d(x, y)))"
|
|
79 |
|
|
80 |
(* List Recursion -- the trancl is Essential; see list.ML *)
|
|
81 |
|
|
82 |
List_rec_def
|
|
83 |
"List_rec(M, c, d) == wfrec(trancl(pred_Sexp), M, \
|
|
84 |
\ %z g. List_case(z, c, %x y. d(x, y, g(y))))"
|
|
85 |
|
|
86 |
list_rec_def
|
|
87 |
"list_rec(l, c, d) == \
|
|
88 |
\ List_rec(Rep_List(l), c, %x y r. d(Inv(Leaf, x), Abs_List(y), r))"
|
|
89 |
|
|
90 |
(* Generalized Map Functionals *)
|
|
91 |
|
|
92 |
Rep_map_def
|
|
93 |
"Rep_map(f, xs) == list_rec(xs, NIL, %x l r. CONS(f(x), r))"
|
|
94 |
Abs_map_def
|
|
95 |
"Abs_map(g, M) == List_rec(M, Nil, %N L r. Cons(g(N), r))"
|
|
96 |
|
|
97 |
null_def "null(xs) == list_rec(xs, True, %x xs r.False)"
|
|
98 |
hd_def "hd(xs) == list_rec(xs, @x.True, %x xs r.x)"
|
|
99 |
tl_def "tl(xs) == list_rec(xs, @xs.True, %x xs r.xs)"
|
|
100 |
list_all_def "list_all(P, xs) == list_rec(xs, True, %x l r. P(x) & r)"
|
|
101 |
map_def "map(f, xs) == list_rec(xs, Nil, %x l r. Cons(f(x), r))"
|
13
|
102 |
append_def "xs@ys == list_rec(xs, ys, %x l r. Cons(x,r))"
|
0
|
103 |
list_case_def "list_case(xs, a, f) == list_rec(xs, a, %x xs r.f(x, xs))"
|
|
104 |
|
|
105 |
end
|