Skip to content

Commit 94906c0

Browse files
authored
refactor: use tighter lemma from Function.Consequences (agda#2574)
1 parent 963d701 commit 94906c0

File tree

1 file changed

+17
-17
lines changed

1 file changed

+17
-17
lines changed

src/Data/List/Relation/Unary/Enumerates/Setoid/Properties.agda

+17-17
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,8 @@
66

77
{-# OPTIONS --cubical-compatible --safe #-}
88

9+
module Data.List.Relation.Unary.Enumerates.Setoid.Properties where
10+
911
open import Data.List.Base
1012
open import Data.List.Membership.Setoid.Properties as Membership
1113
open import Data.List.Relation.Unary.Any using (index)
@@ -14,32 +16,32 @@ open import Data.List.Relation.Unary.Enumerates.Setoid
1416
open import Data.Sum.Base using (inj₁; inj₂)
1517
open import Data.Sum.Relation.Binary.Pointwise
1618
using (_⊎ₛ_; inj₁; inj₂)
17-
open import Data.Product.Base using (_,_; proj₁; proj₂)
19+
open import Data.Product.Base using (_,_)
1820
open import Data.Product.Relation.Binary.Pointwise.NonDependent
1921
using (_×ₛ_)
2022
open import Function.Base using (_∘_)
2123
open import Function.Bundles using (Surjection)
2224
open import Function.Definitions using (Surjective)
23-
open import Function.Consequences using (strictlySurjective⇒surjective)
25+
open import Function.Consequences using (inverseˡ⇒surjective)
2426
open import Level
2527
open import Relation.Binary.Bundles using (Setoid; DecSetoid)
2628
open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_)
2729
open import Relation.Binary.Properties.Setoid using (respʳ-flip)
2830

29-
module Data.List.Relation.Unary.Enumerates.Setoid.Properties where
30-
31-
3231
private
3332
variable
3433
a b ℓ₁ ℓ₂ : Level
34+
A : Set a
35+
xs ys : List A
36+
3537

3638
------------------------------------------------------------------------
3739
-- map
3840

3941
module _ (S : Setoid a ℓ₁) (T : Setoid b ℓ₂) (surj : Surjection S T) where
4042
open Surjection surj
4143

42-
map⁺ : {xs} IsEnumeration S xs IsEnumeration T (map to xs)
44+
map⁺ : IsEnumeration S xs IsEnumeration T (map to xs)
4345
map⁺ _∈xs y with (x , fx≈y) strictlySurjective y =
4446
∈-resp-≈ T fx≈y (∈-map⁺ S T cong (x ∈xs))
4547

@@ -48,15 +50,15 @@ module _ (S : Setoid a ℓ₁) (T : Setoid b ℓ₂) (surj : Surjection S T) whe
4850

4951
module _ (S : Setoid a ℓ₁) where
5052

51-
++⁺ˡ : {xs} ys IsEnumeration S xs IsEnumeration S (xs ++ ys)
53+
++⁺ˡ : ys IsEnumeration S xs IsEnumeration S (xs ++ ys)
5254
++⁺ˡ _ _∈xs v = Membership.∈-++⁺ˡ S (v ∈xs)
5355

54-
++⁺ʳ : xs {ys} IsEnumeration S ys IsEnumeration S (xs ++ ys)
56+
++⁺ʳ : xs IsEnumeration S ys IsEnumeration S (xs ++ ys)
5557
++⁺ʳ _ _∈ys v = Membership.∈-++⁺ʳ S _ (v ∈ys)
5658

5759
module _ (S : Setoid a ℓ₁) (T : Setoid b ℓ₂) where
5860

59-
++⁺ : {xs ys} IsEnumeration S xs IsEnumeration T ys
61+
++⁺ : IsEnumeration S xs IsEnumeration T ys
6062
IsEnumeration (S ⊎ₛ T) (map inj₁ xs ++ map inj₂ ys)
6163
++⁺ _∈xs _ (inj₁ x) = ∈-++⁺ˡ (S ⊎ₛ T) (∈-map⁺ S (S ⊎ₛ T) inj₁ (x ∈xs))
6264
++⁺ _ _∈ys (inj₂ y) = ∈-++⁺ʳ (S ⊎ₛ T) _ (∈-map⁺ T (S ⊎ₛ T) inj₂ (y ∈ys))
@@ -66,7 +68,7 @@ module _ (S : Setoid a ℓ₁) (T : Setoid b ℓ₂) where
6668

6769
module _ (S : Setoid a ℓ₁) (T : Setoid b ℓ₂) where
6870

69-
cartesianProduct⁺ : {xs ys} IsEnumeration S xs IsEnumeration T ys
71+
cartesianProduct⁺ : IsEnumeration S xs IsEnumeration T ys
7072
IsEnumeration (S ×ₛ T) (cartesianProduct xs ys)
7173
cartesianProduct⁺ _∈xs _∈ys (x , y) = ∈-cartesianProduct⁺ S T (x ∈xs) (y ∈ys)
7274

@@ -76,17 +78,15 @@ module _ (S : Setoid a ℓ₁) (T : Setoid b ℓ₂) where
7678
module _ (S? : DecSetoid a ℓ₁) where
7779
open DecSetoid S? renaming (setoid to S)
7880

79-
deduplicate⁺ : {xs} IsEnumeration S xs
80-
IsEnumeration S (deduplicate _≟_ xs)
81+
deduplicate⁺ : IsEnumeration S xs IsEnumeration S (deduplicate _≟_ xs)
8182
deduplicate⁺ = ∈-deduplicate⁺ S _≟_ (respʳ-flip S) ∘_
8283

8384
------------------------------------------------------------------------
8485
-- lookup
8586

8687
module _ (S : Setoid a ℓ₁) where
87-
open Setoid S
88+
open Setoid S using (_≈_; sym)
8889

89-
lookup-surjective : {xs} IsEnumeration S xs
90-
Surjective _≡_ _≈_ (lookup xs)
91-
lookup-surjective _∈xs = strictlySurjective⇒surjective
92-
trans (λ { ≡.refl refl}) (λ y index (y ∈xs) , sym (lookup-index (y ∈xs)))
90+
lookup-surjective : IsEnumeration S xs Surjective _≡_ _≈_ (lookup xs)
91+
lookup-surjective _∈xs = inverseˡ⇒surjective _≈_
92+
λ where ≡.refl sym (lookup-index (_ ∈xs))

0 commit comments

Comments
 (0)