Skip to content

Commit 9a5ec9f

Browse files
authored
Add two lemmas to Data.List.Membership.Setoid.Properties (#2465)
* fixes first part of issue #2463; second part creates dependency cycle * refactor to break dependency cycle * renamed in line with review suggestions * fix bug
1 parent 41001ea commit 9a5ec9f

File tree

4 files changed

+128
-60
lines changed

4 files changed

+128
-60
lines changed

CHANGELOG.md

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -93,6 +93,11 @@ New modules
9393

9494
NB Imports of the existing proof procedures `solve` and `prove` etc. should still be via the top-level interfaces in `Algebra.Solver.*Monoid`.
9595

96+
* Refactored out from `Data.List.Relation.Unary.All.Properties` in order to break a dependency cycle with `Data.List.Membership.Setoid.Properties`:
97+
```agda
98+
Data.List.Relation.Unary.All.Properties.Core
99+
```
100+
96101
Additions to existing modules
97102
-----------------------------
98103

@@ -113,6 +118,12 @@ Additions to existing modules
113118
Env = Vec Carrier
114119
```
115120

121+
* In `Data.List.Membership.Setoid.Properties`:
122+
```agda
123+
Any-∈-swap : Any (_∈ ys) xs → Any (_∈ xs) ys
124+
All-∉-swap : All (_∉ ys) xs → All (_∉ xs) ys
125+
```
126+
116127
* In `Data.List.Properties`:
117128
```agda
118129
product≢0 : All NonZero ns → NonZero (product ns)

src/Data/List/Membership/Setoid/Properties.agda

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -17,6 +17,7 @@ import Data.List.Membership.Setoid as Membership
1717
import Data.List.Relation.Binary.Equality.Setoid as Equality
1818
open import Data.List.Relation.Unary.All as All using (All)
1919
open import Data.List.Relation.Unary.Any as Any using (Any; here; there)
20+
import Data.List.Relation.Unary.All.Properties.Core as All
2021
import Data.List.Relation.Unary.Any.Properties as Any
2122
import Data.List.Relation.Unary.Unique.Setoid as Unique
2223
open import Data.Nat.Base using (suc; z<s; _<_)
@@ -446,3 +447,18 @@ module _ (S : Setoid c ℓ) where
446447
∈-∷=⁻ (here x≈z) y≉v (there y∈) = there y∈
447448
∈-∷=⁻ (there x∈xs) y≉v (here y≈z) = here y≈z
448449
∈-∷=⁻ (there x∈xs) y≉v (there y∈) = there (∈-∷=⁻ x∈xs y≉v y∈)
450+
451+
------------------------------------------------------------------------
452+
-- Any/All symmetry wrt _∈_/_∉_
453+
454+
module _ (S : Setoid c ℓ) where
455+
456+
open Setoid S using (sym)
457+
open Membership S
458+
459+
Any-∈-swap : {xs ys} Any (_∈ ys) xs Any (_∈ xs) ys
460+
Any-∈-swap = Any.swap ∘ Any.map (Any.map sym)
461+
462+
All-∉-swap : {xs ys} All (_∉ ys) xs All (_∉ xs) ys
463+
All-∉-swap p = All.¬Any⇒All¬ _ ((All.All¬⇒¬Any p) ∘ Any-∈-swap)
464+

src/Data/List/Relation/Unary/All/Properties.agda

Lines changed: 5 additions & 60 deletions
Original file line numberDiff line numberDiff line change
@@ -61,6 +61,11 @@ private
6161
x y : A
6262
xs ys : List A
6363

64+
------------------------------------------------------------------------
65+
-- Re-export Core Properties
66+
67+
open import Data.List.Relation.Unary.All.Properties.Core public
68+
6469
------------------------------------------------------------------------
6570
-- Properties regarding Null
6671

@@ -85,66 +90,6 @@ null⇒Null {xs = _ ∷ _} ()
8590

8691
-- See also Data.List.Relation.Unary.All.Properties.WithK.[]=-irrelevant.
8792

88-
------------------------------------------------------------------------
89-
-- Lemmas relating Any, All and negation.
90-
91-
¬Any⇒All¬ : xs ¬ Any P xs All (¬_ ∘ P) xs
92-
¬Any⇒All¬ [] ¬p = []
93-
¬Any⇒All¬ (x ∷ xs) ¬p = ¬p ∘ here ∷ ¬Any⇒All¬ xs (¬p ∘ there)
94-
95-
All¬⇒¬Any : {xs} All (¬_ ∘ P) xs ¬ Any P xs
96-
All¬⇒¬Any (¬p ∷ _) (here p) = ¬p p
97-
All¬⇒¬Any (_ ∷ ¬p) (there p) = All¬⇒¬Any ¬p p
98-
99-
¬All⇒Any¬ : Decidable P xs ¬ All P xs Any (¬_ ∘ P) xs
100-
¬All⇒Any¬ dec [] ¬∀ = contradiction [] ¬∀
101-
¬All⇒Any¬ dec (x ∷ xs) ¬∀ with dec x
102-
... | true because [p] = there (¬All⇒Any¬ dec xs (¬∀ ∘ _∷_ (invert [p])))
103-
... | false because [¬p] = here (invert [¬p])
104-
105-
Any¬⇒¬All : {xs} Any (¬_ ∘ P) xs ¬ All P xs
106-
Any¬⇒¬All (here ¬p) = ¬p ∘ All.head
107-
Any¬⇒¬All (there ¬p) = Any¬⇒¬All ¬p ∘ All.tail
108-
109-
¬Any↠All¬ : {xs} (¬ Any P xs) ↠ All (¬_ ∘ P) xs
110-
¬Any↠All¬ = mk↠ₛ {to = ¬Any⇒All¬ _} (λ y All¬⇒¬Any y , to∘from y)
111-
where
112-
to∘from : {xs} (¬p : All (¬_ ∘ P) xs) ¬Any⇒All¬ xs (All¬⇒¬Any ¬p) ≡ ¬p
113-
to∘from [] = refl
114-
to∘from (¬p ∷ ¬ps) = cong₂ _∷_ refl (to∘from ¬ps)
115-
116-
-- If equality of functions were extensional, then the surjection
117-
-- could be strengthened to a bijection.
118-
119-
from∘to : Extensionality _ _
120-
xs (¬p : ¬ Any P xs) All¬⇒¬Any (¬Any⇒All¬ xs ¬p) ≡ ¬p
121-
from∘to ext [] ¬p = ext λ ()
122-
from∘to ext (x ∷ xs) ¬p = ext λ
123-
{ (here p) refl
124-
; (there p) cong (λ f f p) $ from∘to ext xs (¬p ∘ there)
125-
}
126-
127-
Any¬⇔¬All : {xs} Decidable P Any (¬_ ∘ P) xs ⇔ (¬ All P xs)
128-
Any¬⇔¬All dec = mk⇔ Any¬⇒¬All (¬All⇒Any¬ dec _)
129-
130-
private
131-
-- If equality of functions were extensional, then the logical
132-
-- equivalence could be strengthened to a surjection.
133-
to∘from : Extensionality _ _ (dec : Decidable P)
134-
(¬∀ : ¬ All P xs) Any¬⇒¬All (¬All⇒Any¬ dec xs ¬∀) ≡ ¬∀
135-
to∘from ext P ¬∀ = ext λ ∀P contradiction ∀P ¬∀
136-
137-
module _ {_~_ : REL A B ℓ} where
138-
139-
All-swap : {xs ys}
140-
All (λ x All (x ~_) ys) xs
141-
All (λ y All (_~ y) xs) ys
142-
All-swap {ys = []} _ = []
143-
All-swap {ys = y ∷ ys} [] = All.universal (λ _ []) (y ∷ ys)
144-
All-swap {ys = y ∷ ys} ((x~y ∷ x~ys) ∷ pxs) =
145-
(x~y ∷ (All.map All.head pxs)) ∷
146-
All-swap (x~ys ∷ (All.map All.tail pxs))
147-
14893
------------------------------------------------------------------------
14994
-- Defining properties of lookup and _[_]=_
15095
--
Lines changed: 96 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,96 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Properties related to All
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
module Data.List.Relation.Unary.All.Properties.Core where
10+
11+
open import Axiom.Extensionality.Propositional using (Extensionality)
12+
open import Data.Bool.Base using (true; false)
13+
open import Data.List.Base using (List; []; _∷_)
14+
open import Data.List.Relation.Unary.All as All using (All; []; _∷_)
15+
open import Data.List.Relation.Unary.Any as Any using (Any; here; there)
16+
open import Data.Product.Base as Product using (_,_)
17+
open import Function.Base using (_∘_; _$_)
18+
open import Function.Bundles using (_↠_; mk↠ₛ; _⇔_; mk⇔)
19+
open import Level using (Level)
20+
open import Relation.Binary.Core using (REL)
21+
open import Relation.Binary.PropositionalEquality.Core
22+
using (_≡_; refl; cong; cong₂)
23+
open import Relation.Nullary.Reflects using (invert)
24+
open import Relation.Nullary.Negation.Core using (¬_; contradiction)
25+
open import Relation.Nullary.Decidable.Core using (_because_)
26+
open import Relation.Unary using (Decidable; Pred; Universal)
27+
28+
private
29+
variable
30+
a b p ℓ : Level
31+
A : Set a
32+
B : Set b
33+
P : Pred A p
34+
x y : A
35+
xs ys : List A
36+
37+
------------------------------------------------------------------------
38+
-- Lemmas relating Any, All and negation.
39+
40+
¬Any⇒All¬ : xs ¬ Any P xs All (¬_ ∘ P) xs
41+
¬Any⇒All¬ [] ¬p = []
42+
¬Any⇒All¬ (x ∷ xs) ¬p = ¬p ∘ here ∷ ¬Any⇒All¬ xs (¬p ∘ there)
43+
44+
All¬⇒¬Any : {xs} All (¬_ ∘ P) xs ¬ Any P xs
45+
All¬⇒¬Any (¬p ∷ _) (here p) = ¬p p
46+
All¬⇒¬Any (_ ∷ ¬p) (there p) = All¬⇒¬Any ¬p p
47+
48+
¬All⇒Any¬ : Decidable P xs ¬ All P xs Any (¬_ ∘ P) xs
49+
¬All⇒Any¬ dec [] ¬∀ = contradiction [] ¬∀
50+
¬All⇒Any¬ dec (x ∷ xs) ¬∀ with dec x
51+
... | true because [p] = there (¬All⇒Any¬ dec xs (¬∀ ∘ _∷_ (invert [p])))
52+
... | false because [¬p] = here (invert [¬p])
53+
54+
Any¬⇒¬All : {xs} Any (¬_ ∘ P) xs ¬ All P xs
55+
Any¬⇒¬All (here ¬p) = ¬p ∘ All.head
56+
Any¬⇒¬All (there ¬p) = Any¬⇒¬All ¬p ∘ All.tail
57+
58+
¬Any↠All¬ : {xs} (¬ Any P xs) ↠ All (¬_ ∘ P) xs
59+
¬Any↠All¬ = mk↠ₛ {to = ¬Any⇒All¬ _} (λ y All¬⇒¬Any y , to∘from y)
60+
where
61+
to∘from : {xs} (¬p : All (¬_ ∘ P) xs) ¬Any⇒All¬ xs (All¬⇒¬Any ¬p) ≡ ¬p
62+
to∘from [] = refl
63+
to∘from (¬p ∷ ¬ps) = cong₂ _∷_ refl (to∘from ¬ps)
64+
65+
-- If equality of functions were extensional, then the surjection
66+
-- could be strengthened to a bijection.
67+
68+
from∘to : Extensionality _ _
69+
xs (¬p : ¬ Any P xs) All¬⇒¬Any (¬Any⇒All¬ xs ¬p) ≡ ¬p
70+
from∘to ext [] ¬p = ext λ ()
71+
from∘to ext (x ∷ xs) ¬p = ext λ
72+
{ (here p) refl
73+
; (there p) cong (λ f f p) $ from∘to ext xs (¬p ∘ there)
74+
}
75+
76+
Any¬⇔¬All : {xs} Decidable P Any (¬_ ∘ P) xs ⇔ (¬ All P xs)
77+
Any¬⇔¬All dec = mk⇔ Any¬⇒¬All (¬All⇒Any¬ dec _)
78+
79+
private
80+
-- If equality of functions were extensional, then the logical
81+
-- equivalence could be strengthened to a surjection.
82+
to∘from : Extensionality _ _ (dec : Decidable P)
83+
(¬∀ : ¬ All P xs) Any¬⇒¬All (¬All⇒Any¬ dec xs ¬∀) ≡ ¬∀
84+
to∘from ext P ¬∀ = ext λ ∀P contradiction ∀P ¬∀
85+
86+
module _ {_~_ : REL A B ℓ} where
87+
88+
All-swap : {xs ys}
89+
All (λ x All (x ~_) ys) xs
90+
All (λ y All (_~ y) xs) ys
91+
All-swap {ys = []} _ = []
92+
All-swap {ys = y ∷ ys} [] = All.universal (λ _ []) (y ∷ ys)
93+
All-swap {ys = y ∷ ys} ((x~y ∷ x~ys) ∷ pxs) =
94+
(x~y ∷ (All.map All.head pxs)) ∷
95+
All-swap (x~ys ∷ (All.map All.tail pxs))
96+

0 commit comments

Comments
 (0)