forked from cedille/ial
-
Notifications
You must be signed in to change notification settings - Fork 0
/
list.agda
212 lines (165 loc) Β· 7.01 KB
/
list.agda
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
module list where
open import level
open import bool
open import eq
open import maybe
open import nat
open import unit
open import product
open import empty
open import sum
----------------------------------------------------------------------
-- datatypes
----------------------------------------------------------------------
data π {β} (A : Set β) : Set β where
[] : π A
_::_ : (x : A) (xs : π A) β π A
{-# BUILTIN LIST π #-}
list = π
----------------------------------------------------------------------
-- syntax
----------------------------------------------------------------------
infixr 6 _::_ _++_
infixr 5 _shorter_ _longer_
----------------------------------------------------------------------
-- operations
----------------------------------------------------------------------
[_] : β {β} {A : Set β} β A β π A
[ x ] = x :: []
is-empty : β{β}{A : Set β} β π A β πΉ
is-empty [] = tt
is-empty (_ :: _) = ff
tail : β {β} {A : Set β} β π A β π A
tail [] = []
tail (x :: xs) = xs
head : β{β}{A : Set β} β (l : π A) β .(is-empty l β‘ ff) β A
head [] ()
head (x :: xs) _ = x
head2 : β{β}{A : Set β} β (l : π A) β maybe A
head2 [] = nothing
head2 (a :: _) = just a
last : β{β}{A : Set β} β (l : π A) β is-empty l β‘ ff β A
last [] ()
last (x :: []) _ = x
last (x :: (y :: xs)) _ = last (y :: xs) refl
_++_ : β {β} {A : Set β} β π A β π A β π A
[] ++ ys = ys
(x :: xs) ++ ys = x :: (xs ++ ys)
concat : β{β}{A : Set β} β π (π A) β π A
concat [] = []
concat (l :: ls) = l ++ concat ls
repeat : β{β}{A : Set β} β β β A β π A
repeat 0 a = []
repeat (suc n) a = a :: (repeat n a)
map : β {β β'} {A : Set β} {B : Set β'} β (A β B) β π A β π B
map f [] = []
map f (x :: xs) = f x :: map f xs
-- The hom part of the list functor.
list-funct : {A B : Set} β (A β B) β (π A β π B)
list-funct f l = map f l
{- (maybe-map f xs) returns (just ys) if f returns (just y_i) for each
x_i in the list xs. Otherwise, (maybe-map f xs) returns nothing. -}
πmaybe-map : β {β β'} {A : Set β} {B : Set β'} β (A β maybe B) β π A β maybe (π B)
πmaybe-map f [] = just []
πmaybe-map f (x :: xs) with f x
πmaybe-map f (x :: xs) | nothing = nothing
πmaybe-map f (x :: xs) | just y with πmaybe-map f xs
πmaybe-map f (x :: xs) | just y | nothing = nothing
πmaybe-map f (x :: xs) | just y | just ys = just (y :: ys)
foldr : β{β β'}{A : Set β}{B : Set β'} β (A β B β B) β B β π A β B
foldr f b [] = b
foldr f b (a :: as) = f a (foldr f b as)
length : β{β}{A : Set β} β π A β β
length [] = 0
length (x :: xs) = suc (length xs)
reverse-helper : β {β}{A : Set β} β π A β π A β π A
reverse-helper h [] = h
reverse-helper h (x :: xs) = reverse-helper (x :: h) xs
reverse : β {β}{A : Set β} β π A β π A
reverse l = reverse-helper [] l
list-member : β{β}{A : Set β}(eq : A β A β πΉ)(a : A)(l : π A) β πΉ
list-member eq a [] = ff
list-member eq a (x :: xs) = eq a x || list-member eq a xs
list-minus : β{β}{A : Set β}(eq : A β A β πΉ)(l1 l2 : π A) β π A
list-minus eq [] l2 = []
list-minus eq (x :: xs) l2 =
let r = list-minus eq xs l2 in
if list-member eq x l2 then r else x :: r
_longer_ : β{β}{A : Set β}(l1 l2 : π A) β πΉ
[] longer y = ff
(x :: xs) longer [] = tt
(x :: xs) longer (y :: ys) = xs longer ys
_shorter_ : β{β}{A : Set β}(l1 l2 : π A) β πΉ
x shorter y = y longer x
-- return tt iff all elements in the list satisfy the given predicate pred.
list-all : β{β}{A : Set β}(pred : A β πΉ)(l : π A) β πΉ
list-all pred [] = tt
list-all pred (x :: xs) = pred x && list-all pred xs
all-pred : {X : Set} β (X β Set) β π X β Set
all-pred f [] = β€
all-pred f (xβ :: xs) = (f xβ) β§ (all-pred f xs)
-- return tt iff at least one element in the list satisfies the given predicate pred.
list-any : β{β}{A : Set β}(pred : A β πΉ)(l : π A) β πΉ
list-any pred [] = ff
list-any pred (x :: xs) = pred x || list-any pred xs
list-and : (l : π πΉ) β πΉ
list-and [] = tt
list-and (x :: xs) = x && (list-and xs)
list-or : (l : π πΉ) β πΉ
list-or [] = ff
list-or (x :: l) = x || list-or l
list-max : β{β}{A : Set β} (lt : A β A β πΉ) β π A β A β A
list-max lt [] x = x
list-max lt (y :: ys) x = list-max lt ys (if lt y x then x else y)
isSublist : β{β}{A : Set β} β π A β π A β (A β A β πΉ) β πΉ
isSublist l1 l2 eq = list-all (Ξ» a β list-member eq a l2) l1
=π : β{β}{A : Set β} β (A β A β πΉ) β (l1 : π A) β (l2 : π A) β πΉ
=π eq (a :: as) (b :: bs) = eq a b && =π eq as bs
=π eq [] [] = tt
=π eq _ _ = ff
filter : β{β}{A : Set β} β (A β πΉ) β π A β π A
filter p [] = []
filter p (x :: xs) = let r = filter p xs in
if p x then x :: r else r
-- remove all elements equal to the given one
remove : β{β}{A : Set β}(eq : A β A β πΉ)(a : A)(l : π A) β π A
remove eq a l = filter (Ξ» x β ~ (eq a x)) l
{- nthTail n l returns the part of the list after the first n elements,
or [] if the list has fewer than n elements -}
nthTail : β{β}{A : Set β} β β β π A β π A
nthTail 0 l = l
nthTail n [] = []
nthTail (suc n) (x :: l) = nthTail n l
nth : β{β}{A : Set β} β β β π A β maybe A
nth _ [] = nothing
nth 0 (x :: xs) = just x
nth (suc n) (x :: xs) = nth n xs
-- nats-down N returns N :: (N-1) :: ... :: 0 :: []
nats-down : β β π β
nats-down 0 = [ 0 ]
nats-down (suc x) = suc x :: nats-down x
zip : β{ββ ββ}{A : Set ββ}{B : Set ββ} β π A β π B β π (A Γ B)
zip [] [] = []
zip [] (x :: lβ) = []
zip (x :: lβ) [] = []
zip (x :: lβ) (y :: lβ) = (x , y) :: zip lβ lβ
unzip : β{ββ ββ}{A : Set ββ}{B : Set ββ} β π (A Γ B) β (π A Γ π B)
unzip [] = ([] , [])
unzip ((x , y) :: ps) with unzip ps
... | (xs , ys) = x :: xs , y :: ys
map-β : {ββ ββ ββ : Level} β {A : Set ββ}{B : Set ββ}{C : Set ββ} β (A β C) β (B β C) β π (A β B) β π C
map-β f g [] = []
map-β f g (injβ x :: l) = f x :: map-β f g l
map-β f g (injβ y :: l) = g y :: map-β f g l
proj-ββ : {β β' : Level}{A : Set β}{B : Set β'} β π (A β B) β (π A)
proj-ββ [] = []
proj-ββ (injβ x :: l) = x :: proj-ββ l
proj-ββ (injβ y :: l) = proj-ββ l
proj-ββ : {β β' : Level}{A : Set β}{B : Set β'} β π (A β B) β (π B)
proj-ββ [] = []
proj-ββ (injβ x :: l) = proj-ββ l
proj-ββ (injβ y :: l) = y :: proj-ββ l
drop-nothing : β{β}{A : Set β} β π (maybe A) β π A
drop-nothing [] = []
drop-nothing (nothing :: aa) = drop-nothing aa
drop-nothing (just a :: aa) = a :: drop-nothing aa