module Data.List.Any where
open import Data.Empty
open import Data.Fin
open import Function
open import Function.Equality using (_⟨$⟩_)
open import Function.Equivalence as Equiv using (module Equivalent)
open import Function.Inverse as Inv
using (module Inverse)
open import Data.List as List using (List; []; _∷_)
open import Data.Product as Prod using (∃; _×_; _,_)
open import Level
open import Relation.Nullary
import Relation.Nullary.Decidable as Dec
open import Relation.Unary using () renaming (_⊆_ to _⋐_)
open import Relation.Binary
import Relation.Binary.InducedPreorders as Ind
open import Relation.Binary.List.Pointwise as ListEq using ([]; _∷_)
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_)
data Any {A} (P : A → Set) : List A → Set where
here : ∀ {x xs} (px : P x) → Any P (x ∷ xs)
there : ∀ {x xs} (pxs : Any P xs) → Any P (x ∷ xs)
map : ∀ {A} {P Q : A → Set} → P ⋐ Q → Any P ⋐ Any Q
map g (here px) = here (g px)
map g (there pxs) = there (map g pxs)
tail : ∀ {A x xs} {P : A → Set} → ¬ P x → Any P (x ∷ xs) → Any P xs
tail ¬px (here px) = ⊥-elim (¬px px)
tail ¬px (there pxs) = pxs
any : ∀ {A} {P : A → Set} →
(∀ x → Dec (P x)) → (xs : List A) → Dec (Any P xs)
any p [] = no λ()
any p (x ∷ xs) with p x
any p (x ∷ xs) | yes px = yes (here px)
any p (x ∷ xs) | no ¬px = Dec.map′ there (tail ¬px) (any p xs)
index : ∀ {A} {P : A → Set} {xs} → Any P xs → Fin (List.length xs)
index (here px) = zero
index (there pxs) = suc (index pxs)
module Membership (S : Setoid zero zero) where
private
open module S = Setoid S using (_≈_) renaming (Carrier to A)
open module LS = Setoid (ListEq.setoid S)
using () renaming (_≈_ to _≋_)
lift-resp : ∀ {P} → P Respects _≈_ → Any P Respects _≋_
lift-resp resp [] ()
lift-resp resp (x≈y ∷ xs≈ys) (here px) = here (resp x≈y px)
lift-resp resp (x≈y ∷ xs≈ys) (there pxs) =
there (lift-resp resp xs≈ys pxs)
infix 4 _∈_ _∉_
_∈_ : A → List A → Set
x ∈ xs = Any (_≈_ x) xs
_∉_ : A → List A → Set
x ∉ xs = ¬ x ∈ xs
infix 4 _⊆_ _⊈_
_⊆_ : List A → List A → Set
xs ⊆ ys = ∀ {x} → x ∈ xs → x ∈ ys
_⊈_ : List A → List A → Set
xs ⊈ ys = ¬ xs ⊆ ys
∈-resp-≈ : ∀ {x} → (_≈_ x) Respects _≈_
∈-resp-≈ = flip S.trans
∈-resp-list-≈ : ∀ {x} → _∈_ x Respects _≋_
∈-resp-list-≈ = lift-resp ∈-resp-≈
⊆-preorder : Preorder _ _ _
⊆-preorder = Ind.InducedPreorder₂ (ListEq.setoid S) _∈_ ∈-resp-list-≈
module ⊆-Reasoning where
import Relation.Binary.PreorderReasoning as PreR
open PreR ⊆-preorder public
renaming (_∼⟨_⟩_ to _⊆⟨_⟩_)
infix 1 _∈⟨_⟩_
_∈⟨_⟩_ : ∀ x {xs ys} → x ∈ xs → xs IsRelatedTo ys → x ∈ ys
x ∈⟨ x∈xs ⟩ xs⊆ys = (begin xs⊆ys) x∈xs
map-with-∈ : ∀ {B : Set} (xs : List A) → (∀ {x} → x ∈ xs → B) → List B
map-with-∈ [] f = []
map-with-∈ (x ∷ xs) f = f (here S.refl) ∷ map-with-∈ xs (f ∘ there)
find : ∀ {P : A → Set} {xs} → Any P xs → ∃ λ x → x ∈ xs × P x
find (here px) = (_ , here S.refl , px)
find (there pxs) = Prod.map id (Prod.map there id) (find pxs)
lose : ∀ {P x xs} → P Respects _≈_ → x ∈ xs → P x → Any P xs
lose resp x∈xs px = map (flip resp px) x∈xs
module Membership-≡ where
private
open module M {A : Set} = Membership (PropEq.setoid A) public
hiding (lift-resp; lose; ⊆-preorder; module ⊆-Reasoning)
lose : ∀ {A} {P : A → Set} {x xs} → x ∈ xs → P x → Any P xs
lose {P = P} = M.lose (PropEq.subst P)
⊆-preorder : Set → Preorder _ _ _
⊆-preorder A = Ind.InducedPreorder₂ (PropEq.setoid (List A)) _∈_
(PropEq.subst (_∈_ _))
open Inv public
using (Kind) renaming (equivalent to set; inverse to bag)
[_]-Equality : Kind → Set → Setoid _ _
[ k ]-Equality A = Inv.InducedEquivalence₂ k (_∈_ {A = A})
infix 4 _≈[_]_
_≈[_]_ : {A : Set} → List A → Kind → List A → Set
xs ≈[ k ] ys = Setoid._≈_ ([ k ]-Equality _) xs ys
bag-=⇒set-= : {A : Set} {xs ys : List A} →
xs ≈[ bag ] ys → xs ≈[ set ] ys
bag-=⇒set-= xs≈ys = Inv.⇿⇒ xs≈ys
module ⊆-Reasoning where
import Relation.Binary.PreorderReasoning as PreR
private
open module PR {A : Set} = PreR (⊆-preorder A) public
renaming (_∼⟨_⟩_ to _⊆⟨_⟩_; _≈⟨_⟩_ to _≡⟨_⟩_)
infixr 2 _≈⟨_⟩_
infix 1 _∈⟨_⟩_
_∈⟨_⟩_ : ∀ {A : Set} x {xs ys : List A} →
x ∈ xs → xs IsRelatedTo ys → x ∈ ys
x ∈⟨ x∈xs ⟩ xs⊆ys = (begin xs⊆ys) x∈xs
_≈⟨_⟩_ : ∀ {k} {A : Set} xs {ys zs : List A} →
xs ≈[ k ] ys → ys IsRelatedTo zs → xs IsRelatedTo zs
xs ≈⟨ xs≈ys ⟩ ys≈zs =
xs ⊆⟨ _⟨$⟩_ (Equivalent.to (Inv.⇒⇔ xs≈ys)) ⟩ ys≈zs
satisfied : ∀ {A} {P : A → Set} {xs} → Any P xs → ∃ P
satisfied = Prod.map id Prod.proj₂ ∘ Membership-≡.find