1
- {-# OPTIONS --erasure #-}
2
1
3
2
-- | Proofs on 'Map'.
4
3
module Data.Map.Prop where
@@ -14,13 +13,14 @@ open import Haskell.Data.Maybe using
14
13
15
14
import Data.Map.Maybe as Maybe
16
15
import Haskell.Prelude as List using (map)
16
+ open import Data.Set using (Set )
17
17
import Data.Set as Set
18
18
19
19
{-----------------------------------------------------------------------------
20
20
Proofs
21
21
involving 1 value type
22
22
------------------------------------------------------------------------------}
23
- module _ {k a : Set } {{_ : Ord k}} where
23
+ module _ {k a : Type } {{_ : Ord k}} where
24
24
25
25
--
26
26
prop-member-null
@@ -210,7 +210,7 @@ module _ {k a : Set} {{_ : Ord k}} where
210
210
Proofs
211
211
involving keysSet
212
212
------------------------------------------------------------------------------}
213
- module _ {k a : Set } {{_ : Ord k}} where
213
+ module _ {k a : Type } {{_ : Ord k}} where
214
214
215
215
--
216
216
prop-member-keysSet
@@ -325,12 +325,11 @@ module _ {k a : Set} {{_ : Ord k}} where
325
325
≡ Set.member key (Set.union (keysSet ma) (keysSet mb))
326
326
lem1 key
327
327
rewrite prop-member-keysSet {key} {union ma mb}
328
- rewrite prop-lookup-union key ma mb
329
- rewrite Set.prop-member-union {k} key (keysSet ma) (keysSet mb)
330
- rewrite prop-member-keysSet {key} {ma}
331
- rewrite prop-member-keysSet {key} {mb}
332
- with lookup key ma
333
- with lookup key mb
328
+ | prop-lookup-union key ma mb
329
+ | Set.prop-member-union {k} key (keysSet ma) (keysSet mb)
330
+ | prop-member-keysSet {key} {ma}
331
+ | prop-member-keysSet {key} {mb}
332
+ with lookup key ma | lookup key mb
334
333
... | Nothing | Nothing = refl
335
334
... | Just a | Nothing = refl
336
335
... | Nothing | Just b = refl
@@ -355,11 +354,11 @@ module _ {k a : Set} {{_ : Ord k}} where
355
354
Proofs
356
355
involving withoutKeys and restrictKeys
357
356
------------------------------------------------------------------------------}
358
- module _ {k a : Set } {{_ : Ord k}} where
357
+ module _ {k a : Type } {{_ : Ord k}} where
359
358
360
359
--
361
360
prop-lookup-withoutKeys
362
- : ∀ (key : k) (m : Map k a) (ks : Set.ℙ k)
361
+ : ∀ (key : k) (m : Map k a) (ks : Set k)
363
362
→ lookup key (withoutKeys m ks)
364
363
≡ Maybe.filt (not (Set.member key ks)) (lookup key m)
365
364
--
@@ -377,7 +376,7 @@ module _ {k a : Set} {{_ : Ord k}} where
377
376
378
377
--
379
378
prop-lookup-restrictKeys
380
- : ∀ (key : k) (m : Map k a) (ks : Set.ℙ k)
379
+ : ∀ (key : k) (m : Map k a) (ks : Set k)
381
380
→ lookup key (restrictKeys m ks)
382
381
≡ Maybe.filt (Set.member key ks) (lookup key m)
383
382
--
@@ -441,7 +440,7 @@ module _ {k a : Set} {{_ : Ord k}} where
441
440
442
441
--
443
442
prop-restrictKeys-union
444
- : ∀ (ma mb : Map k a) (ks : Set.ℙ k)
443
+ : ∀ (ma mb : Map k a) (ks : Set k)
445
444
→ restrictKeys (union ma mb) ks
446
445
≡ union (restrictKeys ma ks) (restrictKeys mb ks)
447
446
--
@@ -499,7 +498,7 @@ module _ {k a : Set} {{_ : Ord k}} where
499
498
500
499
--
501
500
prop-withoutKeys-union
502
- : ∀ (ma mb : Map k a) (ks : Set.ℙ k)
501
+ : ∀ (ma mb : Map k a) (ks : Set k)
503
502
→ withoutKeys (union ma mb) ks
504
503
≡ union (withoutKeys ma ks) (withoutKeys mb ks)
505
504
--
@@ -531,7 +530,7 @@ module _ {k a : Set} {{_ : Ord k}} where
531
530
532
531
--
533
532
prop-withoutKeys-difference
534
- : ∀ (m : Map k a) (ka kb : Set.ℙ k)
533
+ : ∀ (m : Map k a) (ka kb : Set k)
535
534
→ withoutKeys m (Set.difference ka kb)
536
535
≡ union (withoutKeys m ka) (restrictKeys m kb)
537
536
--
@@ -582,7 +581,7 @@ module _ {k a : Set} {{_ : Ord k}} where
582
581
583
582
--
584
583
prop-withoutKeys-withoutKeys
585
- : ∀ (m : Map k a) (ka kb : Set.ℙ k)
584
+ : ∀ (m : Map k a) (ka kb : Set k)
586
585
→ withoutKeys (withoutKeys m ka) kb
587
586
≡ withoutKeys m (Set.union ka kb)
588
587
--
@@ -625,7 +624,7 @@ module _ {k a : Set} {{_ : Ord k}} where
625
624
626
625
--
627
626
@0 prop-withoutKeys-intersection
628
- : ∀ (m : Map k a) (ka kb : Set.ℙ k)
627
+ : ∀ (m : Map k a) (ka kb : Set k)
629
628
→ withoutKeys m (Set.intersection ka kb)
630
629
≡ union (withoutKeys m ka) (withoutKeys m kb)
631
630
--
@@ -686,8 +685,8 @@ module _ {k a : Set} {{_ : Ord k}} where
686
685
≡ lookup key ma
687
686
eq-key key
688
687
rewrite prop-lookup-union key ma (restrictKeys mb (keysSet ma))
689
- rewrite prop-lookup-restrictKeys key mb (keysSet ma)
690
- rewrite prop-member-keysSet {k} {a} {key} {ma}
688
+ | prop-lookup-restrictKeys key mb (keysSet ma)
689
+ | prop-member-keysSet {k} {a} {key} {ma}
691
690
with lookup key ma
692
691
... | Nothing = refl
693
692
... | Just a
0 commit comments