diff --git a/CHANGELOG.md b/CHANGELOG.md index e02ff2dc30..156756ff00 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -82,16 +82,9 @@ Deprecated names scanl-defn ↦ Data.List.Scans.Properties.scanl-defn ``` -* In `Data.List.NonEmpty.Base`: - ```agda - inits : List A → List⁺ (List A) - tails : List A → List⁺ (List A) - ``` - -* In `Data.List.NonEmpty.Properties`: +* In `Data.List.Relation.Unary.All.Properties`: ```agda - toList-inits : toList ∘ List⁺.inits ≗ List.inits - toList-tails : toList ∘ List⁺.tails ≗ List.tails + map-compose ↦ map-∘ ``` * In `Data.Nat.Divisibility.Core`: @@ -412,6 +405,12 @@ Additions to existing modules tail∘tails : List A → List (List A) ``` +* In `Data.List.Membership.Propositional.Properties.Core`: + ```agda + find∘∃∈-Any : (p : ∃ λ x → x ∈ xs × P x) → find (∃∈-Any p) ≡ p + ∃∈-Any∘find : (p : Any P xs) → ∃∈-Any (find p) ≡ p + ``` + * In `Data.List.Membership.Setoid.Properties`: ```agda reverse⁺ : x ∈ xs → x ∈ reverse xs @@ -424,6 +423,12 @@ Additions to existing modules tails : List A → List⁺ (List A) ``` +* In `Data.List.NonEmpty.Properties`: + ```agda + toList-inits : toList ∘ List⁺.inits ≗ List.inits + toList-tails : toList ∘ List⁺.tails ≗ List.tails + ``` + * In `Data.List.Properties`: ```agda length-catMaybes : length (catMaybes xs) ≤ length xs @@ -477,6 +482,21 @@ Additions to existing modules mapMaybeIsInj₂∘mapInj₁ : mapMaybe isInj₂ (map inj₁ xs) ≡ [] ``` +* In `Data.List.Relation.Binary.Pointwise.Base`: + ```agda + unzip : Pointwise (R ; S) ⇒ (Pointwise R ; Pointwise S) + ``` + +* In `Data.Maybe.Relation.Binary.Pointwise`: + ```agda + pointwise⊆any : Pointwise R (just x) ⊆ Any (R x) + ``` + +* In `Data.List.Relation.Binary.Sublist.Setoid`: + ```agda + ⊆-upper-bound : ∀ {xs ys zs} (τ : xs ⊆ zs) (σ : ys ⊆ zs) → UpperBound τ σ + ``` + * In `Data.List.Relation.Binary.Sublist.Setoid.Properties`: ```agda ⊆-trans-idˡ : (trans-reflˡ : ∀ {x y} (p : x ≈ y) → trans ≈-refl p ≡ p) → @@ -498,6 +518,11 @@ Additions to existing modules reverse⁻ : reverse as ⊆ reverse bs → as ⊆ bs ``` +* In `Data.List.Relation.Unary.All`: + ```agda + universal-U : Universal (All U) + ``` + * In `Data.List.Relation.Unary.All.Properties`: ```agda All-catMaybes⁺ : All (Maybe.All P) xs → All P (catMaybes xs) @@ -510,6 +535,12 @@ Additions to existing modules tabulate⁺-< : (i < j → R (f i) (f j)) → AllPairs R (tabulate f) ``` +* In `Data.List.Relation.Unary.Any.Properties`: + ```agda + map-cong : (f g : P ⋐ Q) → (∀ {x} (p : P x) → f p ≡ g p) → + (p : Any P xs) → Any.map f p ≡ Any.map g p + ``` + * In `Data.List.Relation.Ternary.Appending.Setoid.Properties`: ```agda through→ : ∃[ xs ] Pointwise _≈_ as xs × Appending xs bs cs → @@ -536,21 +567,6 @@ Additions to existing modules ∃[ xs ] Appending Y U as bs xs × Appending V R xs cs ds ``` -* In `Data.List.Relation.Binary.Pointwise.Base`: - ```agda - unzip : Pointwise (R ; S) ⇒ (Pointwise R ; Pointwise S) - ``` - -* In `Data.Maybe.Relation.Binary.Pointwise`: - ```agda - pointwise⊆any : Pointwise R (just x) ⊆ Any (R x) - ``` - -* In `Data.List.Relation.Binary.Sublist.Setoid`: - ```agda - ⊆-upper-bound : ∀ {xs ys zs} (τ : xs ⊆ zs) (σ : ys ⊆ zs) → UpperBound τ σ - ``` - * In `Data.Nat.Divisibility`: ```agda quotient≢0 : m ∣ n → .{{NonZero n}} → NonZero quotient diff --git a/src/Data/List/Membership/Propositional.agda b/src/Data/List/Membership/Propositional.agda index c1ae08761e..2f5edc2d7e 100644 --- a/src/Data/List/Membership/Propositional.agda +++ b/src/Data/List/Membership/Propositional.agda @@ -10,7 +10,7 @@ module Data.List.Membership.Propositional {a} {A : Set a} where open import Data.List.Relation.Unary.Any using (Any) -open import Relation.Binary.PropositionalEquality.Core using (_≡_; _≢_; subst) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; _≢_; resp; subst) open import Relation.Binary.PropositionalEquality.Properties using (setoid) import Data.List.Membership.Setoid as SetoidMembership @@ -32,4 +32,4 @@ _≢∈_ x∈xs y∈xs = ∀ x≡y → subst (_∈ _) x≡y x∈xs ≢ y∈xs -- Other operations lose : ∀ {p} {P : A → Set p} {x xs} → x ∈ xs → P x → Any P xs -lose = SetoidMembership.lose (setoid A) (subst _) +lose = SetoidMembership.lose (setoid A) (resp _) diff --git a/src/Data/List/Membership/Propositional/Properties.agda b/src/Data/List/Membership/Propositional/Properties.agda index f93bf81f4a..03e31d8aae 100644 --- a/src/Data/List/Membership/Propositional/Properties.agda +++ b/src/Data/List/Membership/Propositional/Properties.agda @@ -10,26 +10,26 @@ module Data.List.Membership.Propositional.Properties where open import Algebra.Core using (Op₂) open import Algebra.Definitions using (Selective) -open import Effect.Monad using (RawMonad) -open import Data.Bool.Base using (Bool; false; true; T) open import Data.Fin.Base using (Fin) open import Data.List.Base as List -open import Data.List.Relation.Unary.Any as Any using (Any; here; there) -open import Data.List.Relation.Unary.Any.Properties - using (map↔; concat↔; >>=↔; ⊛↔; Any-cong; ⊗↔′; ¬Any[]) +open import Data.List.Effectful using (monad) open import Data.List.Membership.Propositional using (_∈_; _∉_; mapWith∈; _≢∈_) import Data.List.Membership.Setoid.Properties as Membership open import Data.List.Relation.Binary.Equality.Propositional using (_≋_; ≡⇒≋; ≋⇒≡) -open import Data.List.Effectful using (monad) -open import Data.Nat.Base using (ℕ; zero; suc; pred; s≤s; _≤_; _<_; _≤ᵇ_) -open import Data.Nat.Properties using (_≤?_; m≤n⇒m≤1+n; ≤ᵇ-reflects-≤; <⇒≢; ≰⇒>) +open import Data.List.Relation.Unary.Any as Any using (Any; here; there) +open import Data.List.Relation.Unary.Any.Properties + using (map↔; concat↔; >>=↔; ⊛↔; Any-cong; ⊗↔′; ¬Any[]) +open import Data.Nat.Base using (ℕ; suc; s≤s; _≤_; _<_; _≰_) +open import Data.Nat.Properties + using (suc-injective; m≤n⇒m≤1+n; _≤?_; <⇒≢; ≰⇒>) open import Data.Product.Base using (∃; ∃₂; _×_; _,_) open import Data.Product.Properties using (×-≡,≡↔≡) open import Data.Product.Function.NonDependent.Propositional using (_×-cong_) import Data.Product.Function.Dependent.Propositional as Σ open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂) +open import Effect.Monad using (RawMonad) open import Function.Base using (_∘_; _∘′_; _$_; id; flip; _⟨_⟩_) open import Function.Definitions using (Injective) import Function.Related.Propositional as Related @@ -40,15 +40,14 @@ open import Level using (Level) open import Relation.Binary.Core using (Rel) open import Relation.Binary.Definitions as Binary hiding (Decidable) open import Relation.Binary.PropositionalEquality.Core as ≡ - using (_≡_; _≢_; refl; sym; trans; cong; cong₂; subst; _≗_) + using (_≡_; _≢_; refl; sym; trans; cong; cong₂; resp; _≗_) open import Relation.Binary.PropositionalEquality.Properties as ≡ using (setoid) import Relation.Binary.Properties.DecTotalOrder as DTOProperties -open import Relation.Unary using (_⟨×⟩_; Decidable) -import Relation.Nullary.Reflects as Reflects +open import Relation.Nullary.Decidable.Core + using (Dec; yes; no; ¬¬-excluded-middle) +open import Relation.Nullary.Negation.Core using (¬_; contradiction) open import Relation.Nullary.Reflects using (invert) -open import Relation.Nullary using (¬_; Dec; does; yes; no; _because_) -open import Relation.Nullary.Negation using (contradiction) -open import Relation.Nullary.Decidable using (¬¬-excluded-middle) +open import Relation.Unary using (_⟨×⟩_; Decidable) private open module ListMonad {ℓ} = RawMonad (monad {ℓ = ℓ}) @@ -128,8 +127,9 @@ module _ {v : A} where ∈-insert xs = Membership.∈-insert (≡.setoid A) xs refl ∈-∃++ : ∀ {xs} → v ∈ xs → ∃₂ λ ys zs → xs ≡ ys ++ [ v ] ++ zs - ∈-∃++ v∈xs with Membership.∈-∃++ (≡.setoid A) v∈xs - ... | ys , zs , _ , refl , eq = ys , zs , ≋⇒≡ eq + ∈-∃++ v∈xs + with ys , zs , _ , refl , eq ← Membership.∈-∃++ (≡.setoid A) v∈xs + = ys , zs , ≋⇒≡ eq ------------------------------------------------------------------------ -- concat @@ -147,8 +147,9 @@ module _ {v : A} where Membership.∈-concat⁺′ (≡.setoid A) v∈vs (Any.map ≡⇒≋ vs∈xss) ∈-concat⁻′ : ∀ xss → v ∈ concat xss → ∃ λ xs → v ∈ xs × xs ∈ xss - ∈-concat⁻′ xss v∈c with Membership.∈-concat⁻′ (≡.setoid A) xss v∈c - ... | xs , v∈xs , xs∈xss = xs , v∈xs , Any.map ≋⇒≡ xs∈xss + ∈-concat⁻′ xss v∈c = + let xs , v∈xs , xs∈xss = Membership.∈-concat⁻′ (≡.setoid A) xss v∈c + in xs , v∈xs , Any.map ≋⇒≡ xs∈xss concat-∈↔ : ∀ {xss : List (List A)} → (∃ λ xs → v ∈ xs × xs ∈ xss) ↔ v ∈ concat xss @@ -183,8 +184,9 @@ module _ (f : A → B → C) where ∈-cartesianProduct⁻ : ∀ xs ys {xy@(x , y) : A × B} → xy ∈ cartesianProduct xs ys → x ∈ xs × y ∈ ys -∈-cartesianProduct⁻ xs ys xy∈p[xs,ys] with ∈-cartesianProductWith⁻ _,_ xs ys xy∈p[xs,ys] -... | (x , y , x∈xs , y∈ys , refl) = x∈xs , y∈ys +∈-cartesianProduct⁻ xs ys xy∈p[xs,ys] + with _ , _ , x∈xs , y∈ys , refl ← ∈-cartesianProductWith⁻ _,_ xs ys xy∈p[xs,ys] + = x∈xs , y∈ys ------------------------------------------------------------------------ -- applyUpTo @@ -205,8 +207,7 @@ module _ (f : ℕ → A) where ∈-upTo⁺ = ∈-applyUpTo⁺ id ∈-upTo⁻ : ∀ {n i} → i ∈ upTo n → i < n -∈-upTo⁻ p with ∈-applyUpTo⁻ id p -... | _ , i p) ∘ sym) + f′ⱼ∈xs j with i ≤? j + ... | yes i≤j = ∈-if-not-i (<⇒≢ (s≤s i≤j)) + ... | no i≰j = ∈-if-not-i (<⇒≢ (≰⇒> i≰j) ∘ sym) f′-injective′ : Injective _≡_ _≡_ f′ - f′-injective′ {j} {k} eq with i ≤ᵇ j | Reflects.invert (≤ᵇ-reflects-≤ i j) - | i ≤ᵇ k | Reflects.invert (≤ᵇ-reflects-≤ i k) - ... | true | p | true | q = ≡.cong pred (f-inj eq) - ... | true | p | false | q = contradiction (f-inj eq) (lemma p q) - ... | false | p | true | q = contradiction (f-inj eq) (lemma q p ∘ sym) - ... | false | p | false | q = f-inj eq + f′-injective′ {j} {k} eq with i ≤? j | i ≤? k + ... | yes i≤j | yes i≤k = suc-injective (f-inj eq) + ... | yes i≤j | no i≰k = contradiction (f-inj eq) (lemma i≤j i≰k) + ... | no i≰j | yes i≤k = contradiction (f-inj eq) (lemma i≤k i≰j ∘ sym) + ... | no i≰j | no i≰k = f-inj eq f′-inj : ℕ ↣ _ f′-inj = record diff --git a/src/Data/List/Membership/Propositional/Properties/Core.agda b/src/Data/List/Membership/Propositional/Properties/Core.agda index ca8ffbb1ea..5c77e40b96 100644 --- a/src/Data/List/Membership/Propositional/Properties/Core.agda +++ b/src/Data/List/Membership/Propositional/Properties/Core.agda @@ -12,75 +12,74 @@ module Data.List.Membership.Propositional.Properties.Core where -open import Function.Base using (flip; id; _∘_) -open import Function.Bundles open import Data.List.Base using (List) -open import Data.List.Relation.Unary.Any as Any using (Any; here; there) open import Data.List.Membership.Propositional -open import Data.Product.Base as Product - using (_,_; proj₁; proj₂; uncurry′; ∃; _×_) +open import Data.List.Relation.Unary.Any as Any using (Any; here; there) +open import Data.Product.Base as Product using (_,_; ∃; _×_) +open import Function.Base using (flip; id; _∘_) +open import Function.Bundles using (_↔_; mk↔ₛ′) open import Level using (Level) open import Relation.Binary.PropositionalEquality.Core - using (_≡_; refl; cong; subst) + using (_≡_; refl; cong; resp) open import Relation.Unary using (Pred; _⊆_) private variable a p q : Level A : Set a - ------------------------------------------------------------------------- --- Lemmas relating map and find. - -map∘find : ∀ {P : Pred A p} {xs} - (p : Any P xs) → let p′ = find p in - {f : _≡_ (proj₁ p′) ⊆ P} → - f refl ≡ proj₂ (proj₂ p′) → - Any.map f (proj₁ (proj₂ p′)) ≡ p -map∘find (here p) hyp = cong here hyp -map∘find (there p) hyp = cong there (map∘find p hyp) - -find∘map : ∀ {P : Pred A p} {Q : Pred A q} - {xs : List A} (p : Any P xs) (f : P ⊆ Q) → - find (Any.map f p) ≡ Product.map id (Product.map id f) (find p) -find∘map (here p) f = refl -find∘map (there p) f rewrite find∘map p f = refl + x : A + xs : List A ------------------------------------------------------------------------ -- find satisfies a simple equality when the predicate is a -- propositional equality. -find-∈ : ∀ {x : A} {xs : List A} (x∈xs : x ∈ xs) → - find x∈xs ≡ (x , x∈xs , refl) +find-∈ : (x∈xs : x ∈ xs) → find x∈xs ≡ (x , x∈xs , refl) find-∈ (here refl) = refl find-∈ (there x∈xs) rewrite find-∈ x∈xs = refl ------------------------------------------------------------------------ --- find and lose are inverses (more or less). +-- Lemmas relating map and find. -lose∘find : ∀ {P : Pred A p} {xs : List A} - (p : Any P xs) → - uncurry′ lose (proj₂ (find p)) ≡ p -lose∘find p = map∘find p refl +module _ {P : Pred A p} where -find∘lose : ∀ (P : Pred A p) {x xs} - (x∈xs : x ∈ xs) (pp : P x) → - find {P = P} (lose x∈xs pp) ≡ (x , x∈xs , pp) -find∘lose P x∈xs p - rewrite find∘map x∈xs (flip (subst P) p) - | find-∈ x∈xs - = refl + map∘find : (p : Any P xs) → let x , x∈xs , px = find p in + {f : (x ≡_) ⊆ P} → f refl ≡ px → + Any.map f x∈xs ≡ p + map∘find (here p) hyp = cong here hyp + map∘find (there p) hyp = cong there (map∘find p hyp) + + find∘map : ∀ {Q : Pred A q} {xs} (p : Any P xs) (f : P ⊆ Q) → + let x , x∈xs , px = find p in + find (Any.map f p) ≡ (x , x∈xs , f px) + find∘map (here p) f = refl + find∘map (there p) f rewrite find∘map p f = refl ------------------------------------------------------------------------ -- Any can be expressed using _∈_ module _ {P : Pred A p} where - ∃∈-Any : ∀ {xs} → (∃ λ x → x ∈ xs × P x) → Any P xs - ∃∈-Any = uncurry′ lose ∘ proj₂ + ∃∈-Any : (∃ λ x → x ∈ xs × P x) → Any P xs + ∃∈-Any (x , x∈xs , px) = lose {P = P} x∈xs px + + ∃∈-Any∘find : (p : Any P xs) → ∃∈-Any (find p) ≡ p + ∃∈-Any∘find p = map∘find p refl + + find∘∃∈-Any : (p : ∃ λ x → x ∈ xs × P x) → find (∃∈-Any p) ≡ p + find∘∃∈-Any p@(x , x∈xs , px) + rewrite find∘map x∈xs (flip (resp P) px) | find-∈ x∈xs = refl - Any↔ : ∀ {xs} → (∃ λ x → x ∈ xs × P x) ↔ Any P xs - Any↔ = mk↔ₛ′ ∃∈-Any find lose∘find from∘to - where - from∘to : ∀ v → find (∃∈-Any v) ≡ v - from∘to p = find∘lose _ (proj₁ (proj₂ p)) (proj₂ (proj₂ p)) + Any↔ : (∃ λ x → x ∈ xs × P x) ↔ Any P xs + Any↔ = mk↔ₛ′ ∃∈-Any find ∃∈-Any∘find find∘∃∈-Any + +------------------------------------------------------------------------ +-- Hence, find and lose are inverses (more or less). + +lose∘find : ∀ {P : Pred A p} {xs} (p : Any P xs) → ∃∈-Any (find p) ≡ p +lose∘find = ∃∈-Any∘find + +find∘lose : ∀ (P : Pred A p) {x xs} + (x∈xs : x ∈ xs) (px : P x) → + find (lose {P = P} x∈xs px) ≡ (x , x∈xs , px) +find∘lose P {x} x∈xs px = find∘∃∈-Any (x , x∈xs , px) diff --git a/src/Data/List/Membership/Setoid.agda b/src/Data/List/Membership/Setoid.agda index cadc600646..b1ff4f2773 100644 --- a/src/Data/List/Membership/Setoid.agda +++ b/src/Data/List/Membership/Setoid.agda @@ -7,17 +7,17 @@ {-# OPTIONS --cubical-compatible --safe #-} open import Relation.Binary.Bundles using (Setoid) -open import Relation.Binary.Definitions using (_Respects_) module Data.List.Membership.Setoid {c ℓ} (S : Setoid c ℓ) where -open import Function.Base using (_∘_; id; flip) -open import Data.List.Base as List using (List; []; _∷_; length; lookup) +open import Data.List.Base using (List; []; _∷_) open import Data.List.Relation.Unary.Any as Any - using (Any; index; map; here; there) + using (Any; map; here; there) open import Data.Product.Base as Product using (∃; _×_; _,_) -open import Relation.Unary using (Pred) +open import Function.Base using (_∘_; flip; const) +open import Relation.Binary.Definitions using (_Respects_) open import Relation.Nullary.Negation using (¬_) +open import Relation.Unary using (Pred) open Setoid S renaming (Carrier to A) @@ -49,8 +49,8 @@ mapWith∈ (x ∷ xs) f = f (here refl) ∷ mapWith∈ xs (f ∘ there) module _ {p} {P : Pred A p} where find : ∀ {xs} → Any P xs → ∃ λ x → x ∈ xs × P x - find (here px) = (_ , here refl , px) - find (there pxs) = Product.map id (Product.map there id) (find pxs) + find (here px) = _ , here refl , px + find (there pxs) = let x , x∈xs , px = find pxs in x , there x∈xs , px lose : P Respects _≈_ → ∀ {x xs} → x ∈ xs → P x → Any P xs lose resp x∈xs px = map (flip resp px) x∈xs diff --git a/src/Data/List/Membership/Setoid/Properties.agda b/src/Data/List/Membership/Setoid/Properties.agda index a0fcb4881b..48303dec37 100644 --- a/src/Data/List/Membership/Setoid/Properties.agda +++ b/src/Data/List/Membership/Setoid/Properties.agda @@ -13,15 +13,14 @@ open import Data.Bool.Base using (true; false) open import Data.Fin.Base using (Fin; zero; suc) open import Data.Fin.Properties using (suc-injective) open import Data.List.Base hiding (find) -open import Data.List.Relation.Unary.Any as Any using (Any; here; there) -open import Data.List.Relation.Unary.All as All using (All) -import Data.List.Relation.Unary.Any.Properties as Any import Data.List.Membership.Setoid as Membership import Data.List.Relation.Binary.Equality.Setoid as Equality +open import Data.List.Relation.Unary.All as All using (All) +open import Data.List.Relation.Unary.Any as Any using (Any; here; there) +import Data.List.Relation.Unary.Any.Properties as Any import Data.List.Relation.Unary.Unique.Setoid as Unique -open import Data.Nat.Base using (suc; z≤n; s≤s; _≤_; _<_) -open import Data.Nat.Properties using (≤-trans; n≤1+n) -open import Data.Product.Base as Product using (∃; _×_; _,_ ; ∃₂; proj₁; proj₂) +open import Data.Nat.Base using (suc; z