-
Notifications
You must be signed in to change notification settings - Fork 381
/
Copy pathQuantifiers.idr
315 lines (267 loc) · 10.6 KB
/
Quantifiers.idr
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
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
module Data.List.Quantifiers
import Data.DPair
import Data.List.Elem
%default total
------------------------------------------------------------------------
-- Types and basic properties
namespace Any
||| A proof that some element of a list satisfies some property
|||
||| @ p the property to be satisfied
public export
data Any : (0 p : a -> Type) -> List a -> Type where
||| A proof that the satisfying element is the first one in the `List`
Here : {0 xs : List a} -> p x -> Any p (x :: xs)
||| A proof that the satisfying element is in the tail of the `List`
There : {0 xs : List a} -> Any p xs -> Any p (x :: xs)
export
Uninhabited (Any p Nil) where
uninhabited (Here _) impossible
uninhabited (There _) impossible
export
{0 p : a -> Type} -> Uninhabited (p x) => Uninhabited (Any p xs) => Uninhabited (Any p $ x::xs) where
uninhabited (Here y) = uninhabited y
uninhabited (There y) = uninhabited y
||| Modify the property given a pointwise function
export
mapProperty : (f : {0 x : a} -> p x -> q x) -> Any p l -> Any q l
mapProperty f (Here p) = Here (f p)
mapProperty f (There p) = There (mapProperty f p)
||| Given a decision procedure for a property, determine if an element of a
||| list satisfies it.
|||
||| @ p the property to be satisfied
||| @ dec the decision procedure
||| @ xs the list to examine
public export
any : (dec : (x : a) -> Dec (p x)) -> (xs : List a) -> Dec (Any p xs)
any _ Nil = No uninhabited
any p (x::xs) with (p x)
any p (x::xs) | Yes prf = Yes (Here prf)
any p (x::xs) | No ctra =
case any p xs of
Yes prf' => Yes (There prf')
No ctra' => No $ \case
Here px => ctra px
There pxs => ctra' pxs
||| Forget the membership proof
export
toExists : Any p xs -> Exists p
toExists (Here prf) = Evidence _ prf
toExists (There prf) = toExists prf
namespace All
||| A proof that all elements of a list satisfy a property. It is a list of
||| proofs, corresponding element-wise to the `List`.
public export
data All : (0 p : a -> Type) -> List a -> Type where
Nil : All p Nil
(::) : {0 xs : List a} -> p x -> All p xs -> All p (x :: xs)
Either (Uninhabited $ p x) (Uninhabited $ All p xs) => Uninhabited (All p $ x::xs) where
uninhabited @{Left _} (px::pxs) = uninhabited px
uninhabited @{Right _} (px::pxs) = uninhabited pxs
||| Modify the property given a pointwise function
export
mapProperty : (f : {0 x : a} -> p x -> q x) -> All p l -> All q l
mapProperty f [] = []
mapProperty f (p::pl) = f p :: mapProperty f pl
||| Modify the property given a pointwise interface function
public export
imapProperty : {0 a : Type}
-> {0 p,q : a -> Type}
-> (0 i : a -> Type)
-> (f : {0 x : a} -> i x => p x -> q x)
-> {0 as : List a}
-> All i as => All p as -> All q as
imapProperty i f @{[]} [] = []
imapProperty i f @{ix :: ixs} (x :: xs) = f @{ix} x :: imapProperty i f @{ixs} xs
||| Forget property source for a homogeneous collection of properties
public export
forget : All (const type) types -> List type
forget [] = []
forget (x :: xs) = x :: forget xs
||| Given a decision procedure for a property, decide whether all elements of
||| a list satisfy it.
|||
||| @ p the property
||| @ dec the decision procedure
||| @ xs the list to examine
public export
all : (dec : (x : a) -> Dec (p x)) -> (xs : List a) -> Dec (All p xs)
all _ Nil = Yes Nil
all d (x::xs) with (d x)
all d (x::xs) | No ctra = No $ \(p::_) => ctra p
all d (x::xs) | Yes prf =
case all d xs of
Yes prf' => Yes (prf :: prf')
No ctra' => No $ \(_::ps) => ctra' ps
export
zipPropertyWith : (f : {0 x : a} -> p x -> q x -> r x) ->
All p xs -> All q xs -> All r xs
zipPropertyWith f [] [] = []
zipPropertyWith f (px :: pxs) (qx :: qxs)
= f px qx :: zipPropertyWith f pxs qxs
export
All (Show . p) xs => Show (All p xs) where
show pxs = "[" ++ show' "" pxs ++ "]"
where
show' : String -> All (Show . p) xs' => All p xs' -> String
show' acc @{[]} [] = acc
show' acc @{[_]} [px] = acc ++ show px
show' acc @{_ :: _} (px :: pxs) = show' (acc ++ show px ++ ", ") pxs
export
All (Eq . p) xs => Eq (All p xs) where
(==) [] [] = True
(==) @{_ :: _} (h1::t1) (h2::t2) = h1 == h2 && t1 == t2
%hint
allEq : All (Ord . p) xs => All (Eq . p) xs
allEq @{[]} = []
allEq @{_ :: _} = %search :: allEq
export
All (Ord . p) xs => Ord (All p xs) where
compare [] [] = EQ
compare @{_ :: _} (h1::t1) (h2::t2) = case compare h1 h2 of
EQ => compare t1 t2
o => o
export
All (Semigroup . p) xs => Semigroup (All p xs) where
(<+>) [] [] = []
(<+>) @{_ :: _} (h1::t1) (h2::t2) = (h1 <+> h2) :: (t1 <+> t2)
%hint
allSemigroup : All (Monoid . p) xs => All (Semigroup . p) xs
allSemigroup @{[]} = []
allSemigroup @{_ :: _} = %search :: allSemigroup
export
All (Monoid . p) xs => Monoid (All p xs) where
neutral @{[]} = []
neutral @{_::_} = neutral :: neutral
||| A heterogeneous list of arbitrary types
public export
HList : List Type -> Type
HList = All id
||| Concatenate lists of proofs.
public export
(++) : All p xs -> All p ys -> All p (xs ++ ys)
[] ++ pys = pys
(px :: pxs) ++ pys = px :: (pxs ++ pys)
export
splitAt : (xs : List a) -> All p (xs ++ ys) -> (All p xs, All p ys)
splitAt [] pxs = ([], pxs)
splitAt (_ :: xs) (px :: pxs) = mapFst (px ::) (splitAt xs pxs)
export
take : (xs : List a) -> All p (xs ++ ys) -> All p xs
take xs pxs = fst (splitAt xs pxs)
export
drop : (xs : List a) -> All p (xs ++ ys) -> All p ys
drop xs pxs = snd (splitAt xs pxs)
------------------------------------------------------------------------
-- Relationship between all and any
||| If there does not exist an element that satifies the property, then it is
||| the case that all elements do not satisfy it.
export
negAnyAll : {xs : List a} -> Not (Any p xs) -> All (Not . p) xs
negAnyAll {xs=Nil} _ = Nil
negAnyAll {xs=x::xs} f = (f . Here) :: negAnyAll (f . There)
||| If there exists an element that doesn't satify the property, then it is
||| not the case that all elements satisfy it.
export
anyNegAll : Any (Not . p) xs -> Not (All p xs)
anyNegAll (Here ctra) (p::_) = ctra p
anyNegAll (There np) (_::ps) = anyNegAll np ps
||| If none of the elements satisfy the property, then not any single one can.
export
allNegAny : All (Not . p) xs -> Not (Any p xs)
allNegAny [] p = absurd p
allNegAny (np :: npxs) (Here p) = absurd (np p)
allNegAny (np :: npxs) (There p) = allNegAny npxs p
||| Given a proof of membership for some element, extract the property proof for it
public export
indexAll : Elem x xs -> All p xs -> p x
indexAll Here (p::_ ) = p
indexAll (There e) ( _::ps) = indexAll e ps
--- Relations between listwise `All` and elementwise `Subset` ---
||| Push in the property from the list level with element level
public export
pushIn : (xs : List a) -> (0 _ : All p xs) -> List $ Subset a p
pushIn [] [] = []
pushIn (x::xs) (p::ps) = Element x p :: pushIn xs ps
||| Pull the elementwise property out to the list level
public export
pullOut : (xs : List $ Subset a p) -> Subset (List a) (All p)
pullOut [] = Element [] []
pullOut (Element x p :: xs) = let Element ss ps = pullOut xs in Element (x::ss) (p::ps)
export
pushInOutInverse : (xs : List a) -> (0 prf : All p xs) -> pullOut (pushIn xs prf) = Element xs prf
pushInOutInverse [] [] = Refl
pushInOutInverse (x::xs) (p::ps) = rewrite pushInOutInverse xs ps in Refl
export
pushOutInInverse : (xs : List $ Subset a p) -> uncurry Quantifiers.pushIn (pullOut xs) = xs
pushOutInInverse [] = Refl
pushOutInInverse (Element x p :: xs) with (pushOutInInverse xs)
pushOutInInverse (Element x p :: xs) | subprf with (pullOut xs)
pushOutInInverse (Element x p :: xs) | subprf | Element ss ps = rewrite subprf in Refl
------------------------------------------------------------------------
-- Partitioning lists according to All
||| Two lists, `xs` and `ys`, whose elements are interleaved in the list `xys`.
public export
data Interleaving : (xs, ys, xys : List a) -> Type where
Nil : Interleaving [] [] []
Left : Interleaving xs ys xys -> Interleaving (x :: xs) ys (x :: xys)
Right : Interleaving xs ys xys -> Interleaving xs (y :: ys) (y :: xys)
||| A record for storing the result of splitting a list `xys` according to some
||| property `p`.
||| The `prfs` and `contras` are related to the original list (`xys`) via
||| `Interleaving`.
|||
||| @ xys the list which has been split
||| @ p the property used for the split
public export
record Split {a : Type} (p : a -> Type) (xys : List a) where
constructor MkSplit
{ayes, naws : List a}
{auto interleaving : Interleaving ayes naws xys}
||| A proof that all elements in `ayes` satisfies the property used for the
||| split.
prfs : All p ayes
||| A proof that all elements in `naws` do not satisfy the property used for
||| the split.
contras : All (Not . p) naws
||| Split the list according to the given decidable property, putting the
||| resulting proofs and contras in an accumulator.
|||
||| @ dec a function which returns a decidable property
||| @ xs a list of elements to split
||| @ a the accumulator
public export
splitOnto : (dec : (x : a) -> Dec (p x))
-> (xs : List a)
-> (a : Split p acc)
-> Split p (reverseOnto acc xs)
splitOnto dec [] a = a
splitOnto dec (x :: xs) (MkSplit prfs contras) =
case dec x of
(Yes prf) => splitOnto dec xs (MkSplit (prf :: prfs) contras)
(No contra) => splitOnto dec xs (MkSplit prfs (contra :: contras))
||| Split the list according to the given decidable property, starting with an
||| empty accumulator.
||| Use `splitOnto` if you want to specify the accumulator.
|||
||| @ dec a function which returns a decidable property
||| @ xs a list of elements to split
||| @ a the accumulator
public export
split : (dec : (x : a) -> Dec (p x))
-> (xs : List a)
-> Split p (reverse xs)
split dec xs = splitOnto dec xs (MkSplit [] [])
||| If any `a` either satisfies p or q then given a List of as,
||| either all values satisfy p
||| or at least one of them sastifies q
public export
decide : ((x : a) -> Either (p x) (q x)) ->
(xs : List a) -> Either (All p xs) (Any q xs)
decide dec [] = Left []
decide dec (x :: xs) = case dec x of
Left px => case decide dec xs of
Left pxs => Left (px :: pxs)
Right pxs => Right (There pxs)
Right px => Right (Here px)