summaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorIain Lane <laney@ubuntu.com>2011-02-25 10:53:20 +0000
committerIain Lane <laney@ubuntu.com>2011-02-25 10:53:20 +0000
commita4ce69d4b3f30a0d508aaedadd8dd89366438712 (patch)
treeee35e782a8edacde5130d9eed82bd3f4c40a328e
parentcc1f5c802508a4098a9d4ea99c356277cbe5ff0d (diff)
Imported Upstream version 0.5
-rw-r--r--.boring1
-rw-r--r--AllNonAsciiChars.hs2
-rw-r--r--Everything.agda27
-rw-r--r--GenerateEverything.hs2
-rw-r--r--LICENCE5
-rw-r--r--README.agda16
-rw-r--r--lib.cabal13
-rw-r--r--release-notes12
-rw-r--r--src/Algebra/Props/BooleanAlgebra.agda30
-rw-r--r--src/Algebra/Props/BooleanAlgebra/Expression.agda181
-rw-r--r--src/Algebra/Props/DistributiveLattice.agda23
-rw-r--r--src/Algebra/Props/Lattice.agda57
-rw-r--r--src/Category/Applicative.agda7
-rw-r--r--src/Category/Applicative/Indexed.agda40
-rw-r--r--src/Category/Functor.agda9
-rw-r--r--src/Category/Monad.agda23
-rw-r--r--src/Category/Monad/Continuation.agda23
-rw-r--r--src/Category/Monad/Identity.agda6
-rw-r--r--src/Category/Monad/Indexed.agda14
-rw-r--r--src/Category/Monad/Partiality.agda436
-rw-r--r--src/Category/Monad/State.agda56
-rw-r--r--src/Data/BoundedVec/Inefficient.agda14
-rw-r--r--src/Data/Colist.agda74
-rw-r--r--src/Data/Container.agda92
-rw-r--r--src/Data/Container/AlternativeBagAndSetEquality.agda145
-rw-r--r--src/Data/Container/Any.agda12
-rw-r--r--src/Data/Context.agda188
-rw-r--r--src/Data/Fin.agda9
-rw-r--r--src/Data/Fin/Dec.agda24
-rw-r--r--src/Data/Fin/Props.agda18
-rw-r--r--src/Data/Fin/Subset.agda66
-rw-r--r--src/Data/Fin/Subset/Props.agda165
-rw-r--r--src/Data/Fin/Substitution/Lemmas.agda12
-rw-r--r--src/Data/List.agda20
-rw-r--r--src/Data/List/All.agda25
-rw-r--r--src/Data/List/All/Properties.agda20
-rw-r--r--src/Data/List/Any.agda62
-rw-r--r--src/Data/List/Any/BagAndSetEquality.agda239
-rw-r--r--src/Data/List/Any/Membership.agda93
-rw-r--r--src/Data/List/Any/Properties.agda233
-rw-r--r--src/Data/List/NonEmpty.agda54
-rw-r--r--src/Data/List/NonEmpty/Properties.agda32
-rw-r--r--src/Data/List/Properties.agda85
-rw-r--r--src/Data/Maybe.agda25
-rw-r--r--src/Data/Nat.agda4
-rw-r--r--src/Data/Nat/DivMod.agda4
-rw-r--r--src/Data/Plus.agda84
-rw-r--r--src/Data/Product/N-ary.agda66
-rw-r--r--src/Data/ReflexiveClosure.agda48
-rw-r--r--src/Data/Star.agda49
-rw-r--r--src/Data/Star/Nat.agda5
-rw-r--r--src/Data/Star/Properties.agda30
-rw-r--r--src/Data/Vec.agda53
-rw-r--r--src/Data/Vec/N-ary.agda149
-rw-r--r--src/Data/Vec/Properties.agda140
-rw-r--r--src/Data/W.agda31
-rw-r--r--src/Function/Inverse/TypeIsomorphisms.agda79
-rw-r--r--src/Function/Surjection.agda10
-rw-r--r--src/IO.agda4
-rw-r--r--src/Reflection.agda173
-rw-r--r--src/Relation/Binary/List/Pointwise.agda52
-rw-r--r--src/Relation/Binary/Product/StrictLex.agda9
-rw-r--r--src/Relation/Binary/Reflection.agda28
-rw-r--r--src/Relation/Binary/Sigma/Pointwise.agda168
-rw-r--r--src/Relation/Binary/Vec/Pointwise.agda204
-rw-r--r--src/Relation/Nullary/Negation.agda2
-rw-r--r--src/Universe.agda19
67 files changed, 3009 insertions, 1092 deletions
diff --git a/.boring b/.boring
index ed47212..9507b61 100644
--- a/.boring
+++ b/.boring
@@ -1,5 +1,6 @@
\.l?agda\.el$
\.agdai$
(^|/)MAlonzo($|/)
+^dist($|/)
^html($|/)
^Everything\.agda$
diff --git a/AllNonAsciiChars.hs b/AllNonAsciiChars.hs
index 50de1f9..fb90dbd 100644
--- a/AllNonAsciiChars.hs
+++ b/AllNonAsciiChars.hs
@@ -19,7 +19,7 @@ readUTF8File f = do
main :: IO ()
main = do
agdaFiles <- find always
- (extension ~~? ".agda" ||? extension ~~? ".lagda")
+ (extension ==? ".agda" ||? extension ==? ".lagda")
"src"
nonAsciiChars <-
filter (not . isAscii) . concat <$> mapM readUTF8File agdaFiles
diff --git a/Everything.agda b/Everything.agda
index f60361b..7e92f6c 100644
--- a/Everything.agda
+++ b/Everything.agda
@@ -26,6 +26,9 @@ import Algebra.Props.AbelianGroup
-- Some derivable properties
import Algebra.Props.BooleanAlgebra
+-- Boolean algebra expressions
+import Algebra.Props.BooleanAlgebra.Expression
+
-- Some derivable properties
import Algebra.Props.DistributiveLattice
@@ -127,15 +130,15 @@ import Data.Conat
-- Containers, based on the work of Abbott and others
import Data.Container
--- Alternative definition of bag and set equality
-import Data.Container.AlternativeBagAndSetEquality
-
-- Properties related to ◇
import Data.Container.Any
-- Container combinators
import Data.Container.Combinator
+-- Contexts, variables, substitutions, etc.
+import Data.Context
+
-- Coinductive vectors
import Data.Covec
@@ -266,12 +269,21 @@ import Data.Nat.Properties
-- Showing natural numbers
import Data.Nat.Show
+-- Transitive closures
+import Data.Plus
+
-- Products
import Data.Product
+-- N-ary products
+import Data.Product.N-ary
+
-- Rational numbers
import Data.Rational
+-- Reflexive closures
+import Data.ReflexiveClosure
+
-- Signs
import Data.Sign
@@ -332,6 +344,9 @@ import Data.Vec.N-ary
-- Some Vec-related properties
import Data.Vec.Properties
+-- W-types
+import Data.W
+
-- Types used (only) when calling out to Haskell via the FFI
import Foreign.Haskell
@@ -484,6 +499,9 @@ import Relation.Binary.StrictToNonStrict
-- Sums of binary relations
import Relation.Binary.Sum
+-- Pointwise lifting of relations to vectors
+import Relation.Binary.Vec.Pointwise
+
-- Operations on nullary relations (like negation and decidability)
import Relation.Nullary
@@ -507,3 +525,6 @@ import Relation.Unary
-- Sizes for Agda's sized types
import Size
+
+-- Universes
+import Universe
diff --git a/GenerateEverything.hs b/GenerateEverything.hs
index eaed773..4515f86 100644
--- a/GenerateEverything.hs
+++ b/GenerateEverything.hs
@@ -21,7 +21,7 @@ main = do
header <- readFile headerFile
modules <- filter isLibraryModule . List.sort <$>
find always
- (extension ~~? ".agda" ||? extension ~~? ".lagda")
+ (extension ==? ".agda" ||? extension ==? ".lagda")
srcDir
headers <- mapM extractHeader modules
diff --git a/LICENCE b/LICENCE
index a83b34d..507653d 100644
--- a/LICENCE
+++ b/LICENCE
@@ -1,7 +1,8 @@
-Copyright (c) 2007-2010 Nils Anders Danielsson, Ulf Norell, Shin-Cheng
+Copyright (c) 2007-2011 Nils Anders Danielsson, Ulf Norell, Shin-Cheng
Mu, Samuel Bronson, Dan Doel, Patrik Jansson, Liang-Ting Chen,
Jean-Philippe Bernardy, Andrés Sicard-Ramírez, Nicolas Pouillard,
-Darin Morrison, Peter Berry
+Darin Morrison, Peter Berry, Daniel Brown, Simon Foster, Dominique
+Devriese
Permission is hereby granted, free of charge, to any person obtaining a
copy of this software and associated documentation files (the
diff --git a/README.agda b/README.agda
index 3749b79..e995e8e 100644
--- a/README.agda
+++ b/README.agda
@@ -1,17 +1,19 @@
module README where
------------------------------------------------------------------------
--- The Agda standard library, version 0.4
+-- The Agda standard library
--
-- Author: Nils Anders Danielsson, with contributions from
--- Jean-Philippe Bernardy, Peter Berry, Samuel Bronson, Liang-Ting
--- Chen, Dan Doel, Patrik Jansson, Darin Morrison, Shin-Cheng Mu, Ulf
--- Norell, Nicolas Pouillard and Andrés Sicard-Ramírez
+-- Jean-Philippe Bernardy, Peter Berry, Samuel Bronson, Daniel Brown,
+-- Liang-Ting Chen, Dominique Devriese, Dan Doel, Simon Foster, Patrik
+-- Jansson, Darin Morrison, Shin-Cheng Mu, Ulf Norell, Nicolas
+-- Pouillard and Andrés Sicard-Ramírez
------------------------------------------------------------------------
--- This version of the library has been tested using Agda 2.2.8.
+-- Note that the development version of the library often requires the
+-- latest development version of Agda.
--- Note that no guarantees are currently made about forwards or
+-- Note also that no guarantees are currently made about forwards or
-- backwards compatibility, the library is still at an experimental
-- stage.
@@ -66,6 +68,8 @@ module README where
-- binary relations).
-- • Size
-- Sizes used by the sized types mechanism.
+-- • Universe
+-- A definition of universes.
------------------------------------------------------------------------
-- A selection of useful library modules
diff --git a/lib.cabal b/lib.cabal
index 215ac74..a70149c 100644
--- a/lib.cabal
+++ b/lib.cabal
@@ -1,19 +1,18 @@
name: lib
-version: 0.0
+version: 0.0.2
cabal-version: >= 1.8
build-type: Simple
description: Helper programs.
-tested-with: GHC == 6.12.1
executable GenerateEverything
hs-source-dirs: .
main-is: GenerateEverything.hs
- build-depends: base == 4.2.*,
- FileManip == 0.3.*,
- filepath == 1.1.*
+ build-depends: base >= 4.2 && < 4.4,
+ filemanip == 0.3.*,
+ filepath >= 1.1 && < 1.3
executable AllNonAsciiChars
hs-source-dirs: .
main-is: AllNonAsciiChars.hs
- build-depends: base == 4.2.*,
- FileManip == 0.3.*
+ build-depends: base >= 4.2 && < 4.4,
+ filemanip == 0.3.*
diff --git a/release-notes b/release-notes
index 467be99..fe54e58 100644
--- a/release-notes
+++ b/release-notes
@@ -1,4 +1,16 @@
------------------------------------------------------------------------
+Version 0.5
+------------------------------------------------------------------------
+
+Version 0.5 of the standard library has now been released, see
+http://wiki.portal.chalmers.se/agda/pmwiki.php?n=Libraries.StandardLibrary.
+
+The library has been tested using Agda version 2.2.10.
+
+Note that no guarantees are made about backwards or forwards
+compatibility, the library is still at an experimental stage.
+
+------------------------------------------------------------------------
Version 0.4
------------------------------------------------------------------------
diff --git a/src/Algebra/Props/BooleanAlgebra.agda b/src/Algebra/Props/BooleanAlgebra.agda
index 8efac7b..d0ec453 100644
--- a/src/Algebra/Props/BooleanAlgebra.agda
+++ b/src/Algebra/Props/BooleanAlgebra.agda
@@ -11,13 +11,18 @@ module Algebra.Props.BooleanAlgebra
where
open BooleanAlgebra B
-import Algebra.Props.DistributiveLattice as DL
-open DL distributiveLattice public
+import Algebra.Props.DistributiveLattice
+private
+ open module DL = Algebra.Props.DistributiveLattice
+ distributiveLattice public
+ hiding (replace-equality)
open import Algebra.Structures
import Algebra.FunctionProperties as P; open P _≈_
import Relation.Binary.EqReasoning as EqR; open EqR setoid
open import Relation.Binary
open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence using (_⇔_; module Equivalent)
open import Data.Product
------------------------------------------------------------------------
@@ -237,6 +242,27 @@ deMorgan₂ x y = begin
¬ ¬ (¬ x ∧ ¬ y) ≈⟨ ¬-involutive _ ⟩
¬ x ∧ ¬ y ∎
+-- One can replace the underlying equality with an equivalent one.
+
+replace-equality :
+ {_≈′_ : Rel Carrier b₂} →
+ (∀ {x y} → x ≈ y ⇔ x ≈′ y) → BooleanAlgebra _ _
+replace-equality {_≈′_} ≈⇔≈′ = record
+ { _≈_ = _≈′_
+ ; _∨_ = _∨_
+ ; _∧_ = _∧_
+ ; ¬_ = ¬_
+ ; ⊤ = ⊤
+ ; ⊥ = ⊥
+ ; isBooleanAlgebra = record
+ { isDistributiveLattice = DistributiveLattice.isDistributiveLattice
+ (DL.replace-equality ≈⇔≈′)
+ ; ∨-complementʳ = λ x → to ⟨$⟩ ∨-complementʳ x
+ ; ∧-complementʳ = λ x → to ⟨$⟩ ∧-complementʳ x
+ ; ¬-cong = λ i≈j → to ⟨$⟩ ¬-cong (from ⟨$⟩ i≈j)
+ }
+ } where open module E {x y} = Equivalent (≈⇔≈′ {x} {y})
+
------------------------------------------------------------------------
-- (⊕, ∧, id, ⊥, ⊤) is a commutative ring
diff --git a/src/Algebra/Props/BooleanAlgebra/Expression.agda b/src/Algebra/Props/BooleanAlgebra/Expression.agda
new file mode 100644
index 0000000..4c0319d
--- /dev/null
+++ b/src/Algebra/Props/BooleanAlgebra/Expression.agda
@@ -0,0 +1,181 @@
+------------------------------------------------------------------------
+-- Boolean algebra expressions
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+open import Algebra
+
+module Algebra.Props.BooleanAlgebra.Expression
+ {b} (B : BooleanAlgebra b b)
+ where
+
+open BooleanAlgebra B
+
+open import Category.Applicative
+import Category.Applicative.Indexed as Applicative
+open import Category.Monad
+open import Category.Monad.Identity
+open import Data.Fin using (Fin)
+open import Data.Nat
+open import Data.Vec as Vec using (Vec)
+open import Data.Product
+import Data.Vec.Properties as VecProp
+open import Function
+open import Relation.Binary.PropositionalEquality as P using (_≗_)
+import Relation.Binary.Reflection as Reflection
+open import Relation.Binary.Vec.Pointwise as PW
+ using (Pointwise; module Pointwise; ext)
+
+-- Expressions made up of variables and the operations of a boolean
+-- algebra.
+
+infixr 7 _and_
+infixr 6 _or_
+
+data Expr n : Set b where
+ var : (x : Fin n) → Expr n
+ _or_ _and_ : (e₁ e₂ : Expr n) → Expr n
+ not : (e : Expr n) → Expr n
+ top bot : Expr n
+
+-- The semantics of an expression, parametrised by an applicative
+-- functor.
+
+module Semantics
+ {F : Set b → Set b}
+ (A : RawApplicative F)
+ where
+
+ open RawApplicative A
+
+ ⟦_⟧ : ∀ {n} → Expr n → Vec (F Carrier) n → F Carrier
+ ⟦ var x ⟧ ρ = Vec.lookup x ρ
+ ⟦ e₁ or e₂ ⟧ ρ = pure _∨_ ⊛ ⟦ e₁ ⟧ ρ ⊛ ⟦ e₂ ⟧ ρ
+ ⟦ e₁ and e₂ ⟧ ρ = pure _∧_ ⊛ ⟦ e₁ ⟧ ρ ⊛ ⟦ e₂ ⟧ ρ
+ ⟦ not e ⟧ ρ = pure ¬_ ⊛ ⟦ e ⟧ ρ
+ ⟦ top ⟧ ρ = pure ⊤
+ ⟦ bot ⟧ ρ = pure ⊥
+
+-- flip Semantics.⟦_⟧ e is natural.
+
+module Naturality
+ {F₁ F₂ : Set b → Set b}
+ {A₁ : RawApplicative F₁}
+ {A₂ : RawApplicative F₂}
+ (f : Applicative.Morphism A₁ A₂)
+ where
+
+ open P.≡-Reasoning
+ open Applicative.Morphism f
+ open Semantics A₁ renaming (⟦_⟧ to ⟦_⟧₁)
+ open Semantics A₂ renaming (⟦_⟧ to ⟦_⟧₂)
+ open RawApplicative A₁ renaming (pure to pure₁; _⊛_ to _⊛₁_)
+ open RawApplicative A₂ renaming (pure to pure₂; _⊛_ to _⊛₂_)
+
+ natural : ∀ {n} (e : Expr n) → op ∘ ⟦ e ⟧₁ ≗ ⟦ e ⟧₂ ∘ Vec.map op
+ natural (var x) ρ = begin
+ op (Vec.lookup x ρ) ≡⟨ P.sym $
+ Applicative.Morphism.op-<$> (VecProp.lookup-morphism x) op ρ ⟩
+ Vec.lookup x (Vec.map op ρ) ∎
+ natural (e₁ or e₂) ρ = begin
+ op (pure₁ _∨_ ⊛₁ ⟦ e₁ ⟧₁ ρ ⊛₁ ⟦ e₂ ⟧₁ ρ) ≡⟨ op-⊛ _ _ ⟩
+ op (pure₁ _∨_ ⊛₁ ⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (op-⊛ _ _) P.refl ⟩
+ op (pure₁ _∨_) ⊛₂ op (⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (P.cong₂ _⊛₂_ (op-pure _) (natural e₁ ρ))
+ (natural e₂ ρ) ⟩
+ pure₂ _∨_ ⊛₂ ⟦ e₁ ⟧₂ (Vec.map op ρ) ⊛₂ ⟦ e₂ ⟧₂ (Vec.map op ρ) ∎
+ natural (e₁ and e₂) ρ = begin
+ op (pure₁ _∧_ ⊛₁ ⟦ e₁ ⟧₁ ρ ⊛₁ ⟦ e₂ ⟧₁ ρ) ≡⟨ op-⊛ _ _ ⟩
+ op (pure₁ _∧_ ⊛₁ ⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (op-⊛ _ _) P.refl ⟩
+ op (pure₁ _∧_) ⊛₂ op (⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (P.cong₂ _⊛₂_ (op-pure _) (natural e₁ ρ))
+ (natural e₂ ρ) ⟩
+ pure₂ _∧_ ⊛₂ ⟦ e₁ ⟧₂ (Vec.map op ρ) ⊛₂ ⟦ e₂ ⟧₂ (Vec.map op ρ) ∎
+ natural (not e) ρ = begin
+ op (pure₁ ¬_ ⊛₁ ⟦ e ⟧₁ ρ) ≡⟨ op-⊛ _ _ ⟩
+ op (pure₁ ¬_) ⊛₂ op (⟦ e ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (op-pure _) (natural e ρ) ⟩
+ pure₂ ¬_ ⊛₂ ⟦ e ⟧₂ (Vec.map op ρ) ∎
+ natural top ρ = begin
+ op (pure₁ ⊤) ≡⟨ op-pure _ ⟩
+ pure₂ ⊤ ∎
+ natural bot ρ = begin
+ op (pure₁ ⊥) ≡⟨ op-pure _ ⟩
+ pure₂ ⊥ ∎
+
+-- An example of how naturality can be used: Any boolean algebra can
+-- be lifted, in a pointwise manner, to vectors of carrier elements.
+
+lift : ℕ → BooleanAlgebra b b
+lift n = record
+ { Carrier = Vec Carrier n
+ ; _≈_ = Pointwise _≈_
+ ; _∨_ = Vec.zipWith _∨_
+ ; _∧_ = Vec.zipWith _∧_
+ ; ¬_ = Vec.map ¬_
+ ; ⊤ = Vec.replicate ⊤
+ ; ⊥ = Vec.replicate ⊥
+ ; isBooleanAlgebra = record
+ { isDistributiveLattice = record
+ { isLattice = record
+ { isEquivalence = PW.isEquivalence isEquivalence
+ ; ∨-comm = λ _ _ → ext λ i →
+ solve i 2 (λ x y → x or y , y or x)
+ (∨-comm _ _) _ _
+ ; ∨-assoc = λ _ _ _ → ext λ i →
+ solve i 3
+ (λ x y z → (x or y) or z , x or (y or z))
+ (∨-assoc _ _ _) _ _ _
+ ; ∨-cong = λ xs≈us ys≈vs → ext λ i →
+ solve₁ i 4 (λ x y u v → x or y , u or v)
+ _ _ _ _
+ (∨-cong (Pointwise.app xs≈us i)
+ (Pointwise.app ys≈vs i))
+ ; ∧-comm = λ _ _ → ext λ i →
+ solve i 2 (λ x y → x and y , y and x)
+ (∧-comm _ _) _ _
+ ; ∧-assoc = λ _ _ _ → ext λ i →
+ solve i 3
+ (λ x y z → (x and y) and z ,
+ x and (y and z))
+ (∧-assoc _ _ _) _ _ _
+ ; ∧-cong = λ xs≈ys us≈vs → ext λ i →
+ solve₁ i 4 (λ x y u v → x and y , u and v)
+ _ _ _ _
+ (∧-cong (Pointwise.app xs≈ys i)
+ (Pointwise.app us≈vs i))
+ ; absorptive = (λ _ _ → ext λ i →
+ solve i 2 (λ x y → x or (x and y) , x)
+ (proj₁ absorptive _ _) _ _) ,
+ (λ _ _ → ext λ i →
+ solve i 2 (λ x y → x and (x or y) , x)
+ (proj₂ absorptive _ _) _ _)
+ }
+ ; ∨-∧-distribʳ = λ _ _ _ → ext λ i →
+ solve i 3
+ (λ x y z → (y and z) or x ,
+ (y or x) and (z or x))
+ (∨-∧-distribʳ _ _ _) _ _ _
+ }
+ ; ∨-complementʳ = λ _ → ext λ i →
+ solve i 1 (λ x → x or (not x) , top)
+ (∨-complementʳ _) _
+ ; ∧-complementʳ = λ _ → ext λ i →
+ solve i 1 (λ x → x and (not x) , bot)
+ (∧-complementʳ _) _
+ ; ¬-cong = λ xs≈ys → ext λ i →
+ solve₁ i 2 (λ x y → not x , not y) _ _
+ (¬-cong (Pointwise.app xs≈ys i))
+ }
+ }
+ where
+ ⟦_⟧Id : ∀ {n} → Expr n → Vec Carrier n → Carrier
+ ⟦_⟧Id = Semantics.⟦_⟧ (RawMonad.rawIApplicative IdentityMonad)
+
+ ⟦_⟧Vec : ∀ {m n} → Expr n → Vec (Vec Carrier m) n → Vec Carrier m
+ ⟦_⟧Vec = Semantics.⟦_⟧ Vec.applicative
+
+ open module R {n} (i : Fin n) =
+ Reflection setoid var
+ (λ e ρ → Vec.lookup i (⟦ e ⟧Vec ρ))
+ (λ e ρ → ⟦ e ⟧Id (Vec.map (Vec.lookup i) ρ))
+ (λ e ρ → sym $ reflexive $
+ Naturality.natural (VecProp.lookup-morphism i) e ρ)
diff --git a/src/Algebra/Props/DistributiveLattice.agda b/src/Algebra/Props/DistributiveLattice.agda
index 4dde88d..b48fc76 100644
--- a/src/Algebra/Props/DistributiveLattice.agda
+++ b/src/Algebra/Props/DistributiveLattice.agda
@@ -11,11 +11,17 @@ module Algebra.Props.DistributiveLattice
where
open DistributiveLattice DL
-import Algebra.Props.Lattice as L; open L lattice public
+import Algebra.Props.Lattice
+private
+ open module L = Algebra.Props.Lattice lattice public
+ hiding (replace-equality)
open import Algebra.Structures
import Algebra.FunctionProperties as P; open P _≈_
+open import Relation.Binary
import Relation.Binary.EqReasoning as EqR; open EqR setoid
open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence using (_⇔_; module Equivalent)
open import Data.Product
∨-∧-distrib : _∨_ DistributesOver _∧_
@@ -62,3 +68,18 @@ open import Data.Product
; _∨_ = _∧_
; isDistributiveLattice = ∧-∨-isDistributiveLattice
}
+
+-- One can replace the underlying equality with an equivalent one.
+
+replace-equality :
+ {_≈′_ : Rel Carrier dl₂} →
+ (∀ {x y} → x ≈ y ⇔ x ≈′ y) → DistributiveLattice _ _
+replace-equality {_≈′_} ≈⇔≈′ = record
+ { _≈_ = _≈′_
+ ; _∧_ = _∧_
+ ; _∨_ = _∨_
+ ; isDistributiveLattice = record
+ { isLattice = Lattice.isLattice (L.replace-equality ≈⇔≈′)
+ ; ∨-∧-distribʳ = λ x y z → to ⟨$⟩ ∨-∧-distribʳ x y z
+ }
+ } where open module E {x y} = Equivalent (≈⇔≈′ {x} {y})
diff --git a/src/Algebra/Props/Lattice.agda b/src/Algebra/Props/Lattice.agda
index e8f46e0..c0a7954 100644
--- a/src/Algebra/Props/Lattice.agda
+++ b/src/Algebra/Props/Lattice.agda
@@ -11,8 +11,11 @@ module Algebra.Props.Lattice {l₁ l₂} (L : Lattice l₁ l₂) where
open Lattice L
open import Algebra.Structures
import Algebra.FunctionProperties as P; open P _≈_
+open import Relation.Binary
import Relation.Binary.EqReasoning as EqR; open EqR setoid
open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence using (_⇔_; module Equivalent)
open import Data.Product
∧-idempotent : Idempotent _∧_
@@ -47,3 +50,57 @@ open import Data.Product
; _∨_ = _∧_
; isLattice = ∧-∨-isLattice
}
+
+-- Every lattice can be turned into a poset.
+
+poset : Poset _ _ _
+poset = record
+ { Carrier = Carrier
+ ; _≈_ = _≈_
+ ; _≤_ = λ x y → x ≈ x ∧ y
+ ; isPartialOrder = record
+ { isPreorder = record
+ { isEquivalence = isEquivalence
+ ; reflexive = λ {i} {j} i≈j → begin
+ i ≈⟨ sym $ ∧-idempotent _ ⟩
+ i ∧ i ≈⟨ ∧-cong refl i≈j ⟩
+ i ∧ j ∎
+ ; trans = λ {i} {j} {k} i≈i∧j j≈j∧k → begin
+ i ≈⟨ i≈i∧j ⟩
+ i ∧ j ≈⟨ ∧-cong refl j≈j∧k ⟩
+ i ∧ (j ∧ k) ≈⟨ sym (∧-assoc _ _ _) ⟩
+ (i ∧ j) ∧ k ≈⟨ ∧-cong (sym i≈i∧j) refl ⟩
+ i ∧ k ∎
+ }
+ ; antisym = λ {x} {y} x≈x∧y y≈y∧x → begin
+ x ≈⟨ x≈x∧y ⟩
+ x ∧ y ≈⟨ ∧-comm _ _ ⟩
+ y ∧ x ≈⟨ sym y≈y∧x ⟩
+ y ∎
+ }
+ }
+
+-- One can replace the underlying equality with an equivalent one.
+
+replace-equality : {_≈′_ : Rel Carrier l₂} →
+ (∀ {x y} → x ≈ y ⇔ x ≈′ y) → Lattice _ _
+replace-equality {_≈′_} ≈⇔≈′ = record
+ { _≈_ = _≈′_
+ ; _∧_ = _∧_
+ ; _∨_ = _∨_
+ ; isLattice = record
+ { isEquivalence = record
+ { refl = to ⟨$⟩ refl
+ ; sym = λ x≈y → to ⟨$⟩ sym (from ⟨$⟩ x≈y)
+ ; trans = λ x≈y y≈z → to ⟨$⟩ trans (from ⟨$⟩ x≈y) (from ⟨$⟩ y≈z)
+ }
+ ; ∨-comm = λ x y → to ⟨$⟩ ∨-comm x y
+ ; ∨-assoc = λ x y z → to ⟨$⟩ ∨-assoc x y z
+ ; ∨-cong = λ x≈y u≈v → to ⟨$⟩ ∨-cong (from ⟨$⟩ x≈y) (from ⟨$⟩ u≈v)
+ ; ∧-comm = λ x y → to ⟨$⟩ ∧-comm x y
+ ; ∧-assoc = λ x y z → to ⟨$⟩ ∧-assoc x y z
+ ; ∧-cong = λ x≈y u≈v → to ⟨$⟩ ∧-cong (from ⟨$⟩ x≈y) (from ⟨$⟩ u≈v)
+ ; absorptive = (λ x y → to ⟨$⟩ proj₁ absorptive x y)
+ , (λ x y → to ⟨$⟩ proj₂ absorptive x y)
+ }
+ } where open module E {x y} = Equivalent (≈⇔≈′ {x} {y})
diff --git a/src/Category/Applicative.agda b/src/Category/Applicative.agda
index 2cb1971..6cafb2e 100644
--- a/src/Category/Applicative.agda
+++ b/src/Category/Applicative.agda
@@ -2,6 +2,8 @@
-- Applicative functors
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- Note that currently the applicative functor laws are not included
-- here.
@@ -10,8 +12,9 @@ module Category.Applicative where
open import Data.Unit
open import Category.Applicative.Indexed
-RawApplicative : (Set → Set) → Set₁
+RawApplicative : ∀ {f} → (Set f → Set f) → Set _
RawApplicative F = RawIApplicative {I = ⊤} (λ _ _ → F)
-module RawApplicative {F : Set → Set} (app : RawApplicative F) where
+module RawApplicative {f} {F : Set f → Set f}
+ (app : RawApplicative F) where
open RawIApplicative app public
diff --git a/src/Category/Applicative/Indexed.agda b/src/Category/Applicative/Indexed.agda
index 680c7e5..20be955 100644
--- a/src/Category/Applicative/Indexed.agda
+++ b/src/Category/Applicative/Indexed.agda
@@ -2,19 +2,24 @@
-- Indexed applicative functors
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- Note that currently the applicative functor laws are not included
-- here.
module Category.Applicative.Indexed where
-open import Function
-open import Data.Product
open import Category.Functor
+open import Data.Product
+open import Function
+open import Level
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
-IFun : Set → Set₁
-IFun I = I → I → Set → Set
+IFun : ∀ {i} → Set i → (ℓ : Level) → Set _
+IFun I ℓ = I → I → Set ℓ → Set ℓ
-record RawIApplicative {I : Set} (F : IFun I) : Set₁ where
+record RawIApplicative {i f} {I : Set i} (F : IFun I f) :
+ Set (i ⊔ suc f) where
infixl 4 _⊛_ _<⊛_ _⊛>_
infix 4 _⊗_
@@ -40,3 +45,28 @@ record RawIApplicative {I : Set} (F : IFun I) : Set₁ where
_⊗_ : ∀ {i j k A B} → F i j A → F j k B → F i k (A × B)
x ⊗ y = (_,_) <$> x ⊛ y
+
+ zipWith : ∀ {i j k A B C} → (A → B → C) → F i j A → F j k B → F i k C
+ zipWith f x y = f <$> x ⊛ y
+
+-- Applicative functor morphisms, specialised to propositional
+-- equality.
+
+record Morphism {i f} {I : Set i} {F₁ F₂ : IFun I f}
+ (A₁ : RawIApplicative F₁)
+ (A₂ : RawIApplicative F₂) : Set (i ⊔ suc f) where
+ module A₁ = RawIApplicative A₁
+ module A₂ = RawIApplicative A₂
+ field
+ op : ∀ {i j X} → F₁ i j X → F₂ i j X
+ op-pure : ∀ {i X} (x : X) → op (A₁.pure {i = i} x) ≡ A₂.pure x
+ op-⊛ : ∀ {i j k X Y} (f : F₁ i j (X → Y)) (x : F₁ j k X) →
+ op (A₁._⊛_ f x) ≡ A₂._⊛_ (op f) (op x)
+
+ op-<$> : ∀ {i j X Y} (f : X → Y) (x : F₁ i j X) →
+ op (A₁._<$>_ f x) ≡ A₂._<$>_ f (op x)
+ op-<$> f x = begin
+ op (A₁._⊛_ (A₁.pure f) x) ≡⟨ op-⊛ _ _ ⟩
+ A₂._⊛_ (op (A₁.pure f)) (op x) ≡⟨ P.cong₂ A₂._⊛_ (op-pure _) P.refl ⟩
+ A₂._⊛_ (A₂.pure f) (op x) ∎
+ where open P.≡-Reasoning
diff --git a/src/Category/Functor.agda b/src/Category/Functor.agda
index 39b9311..33a3814 100644
--- a/src/Category/Functor.agda
+++ b/src/Category/Functor.agda
@@ -2,17 +2,20 @@
-- Functors
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- Note that currently the functor laws are not included here.
module Category.Functor where
open import Function
+open import Level
-record RawFunctor (f : Set → Set) : Set₁ where
+record RawFunctor {ℓ} (F : Set ℓ → Set ℓ) : Set (suc ℓ) where
infixl 4 _<$>_ _<$_
field
- _<$>_ : ∀ {a b} → (a → b) → f a → f b
+ _<$>_ : ∀ {A B} → (A → B) → F A → F B
- _<$_ : ∀ {a b} → a → f b → f a
+ _<$_ : ∀ {A B} → A → F B → F A
x <$ y = const x <$> y
diff --git a/src/Category/Monad.agda b/src/Category/Monad.agda
index 54cc482..3ffb879 100644
--- a/src/Category/Monad.agda
+++ b/src/Category/Monad.agda
@@ -2,6 +2,8 @@
-- Monads
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- Note that currently the monad laws are not included here.
module Category.Monad where
@@ -10,20 +12,23 @@ open import Function
open import Category.Monad.Indexed
open import Data.Unit
-RawMonad : (Set → Set) → Set₁
-RawMonad M = RawIMonad {⊤} (λ _ _ → M)
+RawMonad : ∀ {f} → (Set f → Set f) → Set _
+RawMonad M = RawIMonad {I = ⊤} (λ _ _ → M)
-RawMonadZero : (Set → Set) → Set₁
-RawMonadZero M = RawIMonadZero {⊤} (λ _ _ → M)
+RawMonadZero : ∀ {f} → (Set f → Set f) → Set _
+RawMonadZero M = RawIMonadZero {I = ⊤} (λ _ _ → M)
-RawMonadPlus : (Set → Set) → Set₁
-RawMonadPlus M = RawIMonadPlus {⊤} (λ _ _ → M)
+RawMonadPlus : ∀ {f} → (Set f → Set f) → Set _
+RawMonadPlus M = RawIMonadPlus {I = ⊤} (λ _ _ → M)
-module RawMonad {M : Set → Set} (Mon : RawMonad M) where
+module RawMonad {f} {M : Set f → Set f}
+ (Mon : RawMonad M) where
open RawIMonad Mon public
-module RawMonadZero {M : Set → Set} (Mon : RawMonadZero M) where
+module RawMonadZero {f} {M : Set f → Set f}
+ (Mon : RawMonadZero M) where
open RawIMonadZero Mon public
-module RawMonadPlus {M : Set → Set} (Mon : RawMonadPlus M) where
+module RawMonadPlus {f} {M : Set f → Set f}
+ (Mon : RawMonadPlus M) where
open RawIMonadPlus Mon public
diff --git a/src/Category/Monad/Continuation.agda b/src/Category/Monad/Continuation.agda
index 2dceb53..ee1ab33 100644
--- a/src/Category/Monad/Continuation.agda
+++ b/src/Category/Monad/Continuation.agda
@@ -2,26 +2,28 @@
-- A delimited continuation monad
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Category.Monad.Continuation where
open import Category.Applicative
open import Category.Applicative.Indexed
open import Category.Monad
-open import Category.Monad.Indexed
open import Category.Monad.Identity
-
+open import Category.Monad.Indexed
open import Function
+open import Level
------------------------------------------------------------------------
-- Delimited continuation monads
-DContT : {I : Set} → (I → Set) → (Set → Set) → IFun I
+DContT : ∀ {i f} {I : Set i} → (I → Set f) → (Set f → Set f) → IFun I f
DContT K M r₂ r₁ a = (a → M (K r₁)) → M (K r₂)
-DCont : {I : Set} → (I → Set) → IFun I
+DCont : ∀ {i f} {I : Set i} → (I → Set f) → IFun I f
DCont K = DContT K Identity
-DContTIMonad : ∀ {I} (K : I → Set) {M} →
+DContTIMonad : ∀ {i f} {I : Set i} (K : I → Set f) {M} →
RawMonad M → RawIMonad (DContT K M)
DContTIMonad K Mon = record
{ return = λ a k → k a
@@ -29,14 +31,14 @@ DContTIMonad K Mon = record
}
where open RawMonad Mon
-DContIMonad : ∀ {I} (K : I → Set) → RawIMonad (DCont K)
+DContIMonad : ∀ {i f} {I : Set i} (K : I → Set f) → RawIMonad (DCont K)
DContIMonad K = DContTIMonad K IdentityMonad
------------------------------------------------------------------------
-- Delimited continuation operations
-record RawIMonadDCont {I : Set} (K : I → Set)
- (M : I → I → Set → Set) : Set₁ where
+record RawIMonadDCont {i f} {I : Set i} (K : I → Set f)
+ (M : IFun I f) : Set (i ⊔ suc f) where
field
monad : RawIMonad M
reset : ∀ {r₁ r₂ r₃} → M r₁ r₂ (K r₂) → M r₃ r₃ (K r₁)
@@ -45,7 +47,7 @@ record RawIMonadDCont {I : Set} (K : I → Set)
open RawIMonad monad public
-DContTIMonadDCont : ∀ {I} (K : I → Set) {M} →
+DContTIMonadDCont : ∀ {i f} {I : Set i} (K : I → Set f) {M} →
RawMonad M → RawIMonadDCont K (DContT K M)
DContTIMonadDCont K Mon = record
{ monad = DContTIMonad K Mon
@@ -55,5 +57,6 @@ DContTIMonadDCont K Mon = record
where
open RawIMonad Mon
-DContIMonadDCont : ∀ {I} (K : I → Set) → RawIMonadDCont K (DCont K)
+DContIMonadDCont : ∀ {i f} {I : Set i}
+ (K : I → Set f) → RawIMonadDCont K (DCont K)
DContIMonadDCont K = DContTIMonadDCont K IdentityMonad
diff --git a/src/Category/Monad/Identity.agda b/src/Category/Monad/Identity.agda
index 9ba73de..86cb6ef 100644
--- a/src/Category/Monad/Identity.agda
+++ b/src/Category/Monad/Identity.agda
@@ -2,14 +2,16 @@
-- The identity monad
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Category.Monad.Identity where
open import Category.Monad
-Identity : Set → Set
+Identity : ∀ {f} → Set f → Set f
Identity A = A
-IdentityMonad : RawMonad Identity
+IdentityMonad : ∀ {f} → RawMonad (Identity {f})
IdentityMonad = record
{ return = λ x → x
; _>>=_ = λ x f → f x
diff --git a/src/Category/Monad/Indexed.agda b/src/Category/Monad/Indexed.agda
index a7d4cd3..c5163d8 100644
--- a/src/Category/Monad/Indexed.agda
+++ b/src/Category/Monad/Indexed.agda
@@ -2,14 +2,18 @@
-- Indexed monads
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- Note that currently the monad laws are not included here.
module Category.Monad.Indexed where
-open import Function
open import Category.Applicative.Indexed
+open import Function
+open import Level
-record RawIMonad {I : Set} (M : IFun I) : Set₁ where
+record RawIMonad {i f} {I : Set i} (M : IFun I f) :
+ Set (i ⊔ suc f) where
infixl 1 _>>=_ _>>_
infixr 1 _=<<_
@@ -34,14 +38,16 @@ record RawIMonad {I : Set} (M : IFun I) : Set₁ where
open RawIApplicative rawIApplicative public
-record RawIMonadZero {I : Set} (M : IFun I) : Set₁ where
+record RawIMonadZero {i f} {I : Set i} (M : IFun I f) :
+ Set (i ⊔ suc f) where
field
monad : RawIMonad M
∅ : ∀ {i j A} → M i j A
open RawIMonad monad public
-record RawIMonadPlus {I : Set} (M : IFun I) : Set₁ where
+record RawIMonadPlus {i f} {I : Set i} (M : IFun I f) :
+ Set (i ⊔ suc f) where
infixr 3 _∣_
field
monadZero : RawIMonadZero M
diff --git a/src/Category/Monad/Partiality.agda b/src/Category/Monad/Partiality.agda
index 1e16d89..bbda0e4 100644
--- a/src/Category/Monad/Partiality.agda
+++ b/src/Category/Monad/Partiality.agda
@@ -2,31 +2,33 @@
-- The partiality monad
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Category.Monad.Partiality where
open import Coinduction
open import Category.Monad
open import Data.Bool
open import Data.Nat using (ℕ; zero; suc; _+_)
-open import Data.Product as Prod
-open import Data.Sum
+open import Data.Product as Prod hiding (map)
+open import Data.Sum hiding (map)
open import Function
open import Function.Equivalence using (_⇔_; equivalent)
open import Level
open import Relation.Binary as B hiding (Rel)
open import Relation.Binary.PropositionalEquality as P using (_≡_)
open import Relation.Nullary
-open import Relation.Nullary.Decidable
+open import Relation.Nullary.Decidable hiding (map)
open import Relation.Nullary.Negation
------------------------------------------------------------------------
-- The partiality monad
-data _⊥ (A : Set) : Set where
+data _⊥ {a} (A : Set a) : Set a where
now : (x : A) → A ⊥
later : (x : ∞ (A ⊥)) → A ⊥
-monad : RawMonad _⊥
+monad : ∀ {f} → RawMonad {f = f} _⊥
monad = record
{ return = now
; _>>=_ = _>>=_
@@ -36,23 +38,23 @@ monad = record
now x >>= f = f x
later x >>= f = later (♯ (♭ x >>= f))
-private module M = RawMonad monad
+private module M {f} = RawMonad (monad {f})
-- Non-termination.
-never : {A : Set} → A ⊥
+never : ∀ {a} {A : Set a} → A ⊥
never = later (♯ never)
-- run x for n steps peels off at most n "later" constructors from x.
-run_for_steps : ∀ {A} → A ⊥ → ℕ → A ⊥
+run_for_steps : ∀ {a} {A : Set a} → A ⊥ → ℕ → A ⊥
run now x for n steps = now x
run later x for zero steps = later x
run later x for suc n steps = run ♭ x for n steps
-- Is the computation done?
-isNow : ∀ {A} → A ⊥ → Bool
+isNow : ∀ {a} {A : Set a} → A ⊥ → Bool
isNow (now x) = true
isNow (later x) = false
@@ -93,44 +95,70 @@ Equality k = False (k ≟-Kind other geq)
------------------------------------------------------------------------
-- Equality/ordering
-module Equality {A : Set} -- The "return type".
- (_R_ : A → A → Set) where
+module Equality {a ℓ} {A : Set a} -- The "return type".
+ (_∼_ : A → A → Set ℓ) where
-- The three relations.
- data Rel : Kind → A ⊥ → A ⊥ → Set where
- now : ∀ {k x y} (xRy : x R y) → Rel k (now x) (now y)
+ data Rel : Kind → A ⊥ → A ⊥ → Set (a ⊔ ℓ) where
+ now : ∀ {k x y} (x∼y : x ∼ y) → Rel k (now x) (now y)
later : ∀ {k x y} (x∼y : ∞ (Rel k (♭ x) (♭ y))) → Rel k (later x) (later y)
laterʳ : ∀ {x y} (x≈y : Rel (other weak) x (♭ y) ) → Rel (other weak) x (later y)
laterˡ : ∀ {k x y} (x∼y : Rel (other k) (♭ x) y ) → Rel (other k) (later x) y
infix 4 _≅_ _≳_ _≲_ _≈_
- _≅_ : A ⊥ → A ⊥ → Set
+ _≅_ : A ⊥ → A ⊥ → Set _
_≅_ = Rel strong
- _≳_ : A ⊥ → A ⊥ → Set
+ _≳_ : A ⊥ → A ⊥ → Set _
_≳_ = Rel (other geq)
- _≲_ : A ⊥ → A ⊥ → Set
+ _≲_ : A ⊥ → A ⊥ → Set _
_≲_ = flip _≳_
- _≈_ : A ⊥ → A ⊥ → Set
+ _≈_ : A ⊥ → A ⊥ → Set _
_≈_ = Rel (other weak)
+ -- x ⇓ y means that x terminates with y.
+
+ infix 4 _⇓[_]_ _⇓_
+
+ _⇓[_]_ : A ⊥ → Kind → A → Set _
+ x ⇓[ k ] y = Rel k x (now y)
+
+ _⇓_ : A ⊥ → A → Set _
+ x ⇓ y = x ⇓[ other weak ] y
+
+ -- x ⇑ means that x does not terminate.
+
+ infix 4 _⇑[_] _⇑
+
+ _⇑[_] : A ⊥ → Kind → Set _
+ x ⇑[ k ] = Rel k x never
+
+ _⇑ : A ⊥ → Set _
+ x ⇑ = x ⇑[ other weak ]
+
------------------------------------------------------------------------
-- Lemmas relating the three relations
+private
+ module Dummy {a ℓ} {A : Set a} {_∼_ : A → A → Set ℓ} where
+
+ open Equality _∼_
+ open Equality.Rel
+
-- All relations include strong equality.
≅⇒ : ∀ {k} {x y : A ⊥} → x ≅ y → Rel k x y
- ≅⇒ (now xRy) = now xRy
+ ≅⇒ (now x∼y) = now x∼y
≅⇒ (later x≅y) = later (♯ ≅⇒ (♭ x≅y))
-- The weak equality includes the ordering.
≳⇒ : ∀ {k} {x y : A ⊥} → x ≳ y → Rel (other k) x y
- ≳⇒ (now xRy) = now xRy
+ ≳⇒ (now x∼y) = now x∼y
≳⇒ (later x≳y) = later (♯ ≳⇒ (♭ x≳y))
≳⇒ (laterˡ x≳y) = laterˡ (≳⇒ x≳y )
@@ -153,7 +181,7 @@ module Equality {A : Set} -- The "return type".
now⇒now : ∀ {k₁ k₂} {x} {y : A} →
Rel (other k₁) x (now y) → Rel (other k₂) x (now y)
- now⇒now (now xRy) = now xRy
+ now⇒now (now x∼y) = now x∼y
now⇒now (laterˡ x∼now) = laterˡ (now⇒now x∼now)
------------------------------------------------------------------------
@@ -177,122 +205,93 @@ module Equality {A : Set} -- The "return type".
later⁻¹ (laterˡ x∼ly) = laterʳ⁻¹ x∼ly
------------------------------------------------------------------------
--- Terminating computations
+-- The relations are equivalences or partial orders, given suitable
+-- assumptions about the underlying relation
- -- x ⇓ y means that x terminates with y.
+ module Equivalence where
- infix 4 _⇓[_]_ _⇓_
+ -- Reflexivity.
- _⇓[_]_ : A ⊥ → Kind → A → Set
- x ⇓[ k ] y = Rel k x (now y)
+ refl : Reflexive _∼_ → ∀ {k} → Reflexive (Rel k)
+ refl refl-∼ {x = now v} = now refl-∼
+ refl refl-∼ {x = later x} = later (♯ refl refl-∼)
- _⇓_ : A ⊥ → A → Set
- x ⇓ y = x ⇓[ other weak ] y
+ -- Symmetry.
- -- The number of later constructors (steps) in the terminating
- -- computation x.
+ sym : Symmetric _∼_ → ∀ {k} → Equality k → Symmetric (Rel k)
+ sym sym-∼ eq (now x∼y) = now (sym-∼ x∼y)
+ sym sym-∼ eq (later x∼y) = later (♯ sym sym-∼ eq (♭ x∼y))
+ sym sym-∼ eq (laterʳ x≈y) = laterˡ (sym sym-∼ eq x≈y )
+ sym sym-∼ eq (laterˡ {weak} x≈y) = laterʳ (sym sym-∼ eq x≈y )
+ sym sym-∼ () (laterˡ {geq} x≳y)
- steps : ∀ {k} {x : A ⊥} {y} → x ⇓[ k ] y → ℕ
- steps (now _) = zero
- steps .{x = later x} (laterˡ {x = x} x⇓) = suc (steps {x = ♭ x} x⇓)
-
-------------------------------------------------------------------------
--- Non-terminating computations
+ -- Transitivity.
- -- x ⇑ means that x does not terminate.
-
- infix 4 _⇑[_] _⇑
-
- _⇑[_] : A ⊥ → Kind → Set
- x ⇑[ k ] = Rel k x never
+ private
+ module Trans (trans-∼ : Transitive _∼_) where
- _⇑ : A ⊥ → Set
- x ⇑ = x ⇑[ other weak ]
+ now-trans : ∀ {k x y} {v : A} →
+ Rel k x y → Rel k y (now v) → Rel k x (now v)
+ now-trans (now x∼y) (now y∼z) = now (trans-∼ x∼y y∼z)
+ now-trans x∼ly (laterˡ y∼z) = now-trans (laterʳ⁻¹ x∼ly) y∼z
+ now-trans (laterˡ x∼y) y∼z = laterˡ (now-trans x∼y y∼z)
-------------------------------------------------------------------------
--- The relations are equivalences or partial orders
+ mutual
--- The precondition that the underlying relation is an equivalence can
--- be weakened for some of the properties.
+ later-trans : ∀ {k} {x y : A ⊥} {z} →
+ Rel k x y → Rel k y (later z) → Rel k x (later z)
+ later-trans (later x∼y) ly∼lz = later (♯ trans (♭ x∼y) (later⁻¹ ly∼lz))
+ later-trans (laterˡ x∼y) y∼lz = later (♯ trans x∼y (laterʳ⁻¹ y∼lz))
+ later-trans (laterʳ x≈y) ly≈lz = later-trans x≈y (laterˡ⁻¹ ly≈lz)
+ later-trans x≈y (laterʳ y≈z) = laterʳ ( trans x≈y y≈z )
-module Properties (S : Setoid zero zero) where
+ trans : ∀ {k} {x y z : A ⊥} → Rel k x y → Rel k y z → Rel k x z
+ trans {z = now v} x∼y y∼v = now-trans x∼y y∼v
+ trans {z = later z} x∼y y∼lz = later-trans x∼y y∼lz
- private
- open module R = Setoid S
- using () renaming (Carrier to A; _≈_ to _R_)
- open Equality _R_
+ open Trans public using (trans)
-- All the relations are preorders.
- preorder : Kind → Preorder _ _ _
- preorder k = record
+ preorder : IsPreorder _≡_ _∼_ → Kind → Preorder _ _ _
+ preorder pre k = record
{ Carrier = A ⊥
; _≈_ = _≡_
; _∼_ = Rel k
; isPreorder = record
{ isEquivalence = P.isEquivalence
- ; reflexive = refl _
- ; trans = trans
+ ; reflexive = refl′
+ ; trans = Equivalence.trans (IsPreorder.trans pre)
}
}
where
- refl : ∀ {k} (x : A ⊥) {y} → x ≡ y → Rel k x y
- refl (now v) P.refl = now R.refl
- refl (later x) P.refl = later (♯ refl (♭ x) P.refl)
-
- now-R-trans : ∀ {k x} {y z : A} →
- Rel k x (now y) → y R z → Rel k x (now z)
- now-R-trans (now xRy) yRz = now (R.trans xRy yRz)
- now-R-trans (laterˡ x∼y) yRz = laterˡ (now-R-trans x∼y yRz)
+ refl′ : ∀ {k} {x y : A ⊥} → x ≡ y → Rel k x y
+ refl′ P.refl = Equivalence.refl (IsPreorder.refl pre)
- now-trans : ∀ {k x y} {v : A} →
- Rel k x y → Rel k y (now v) → Rel k x (now v)
- now-trans x∼ly (laterˡ y∼z) = now-trans (laterʳ⁻¹ x∼ly) y∼z
- now-trans x∼y (now yRz) = now-R-trans x∼y yRz
-
- mutual
-
- later-trans : ∀ {k} {x y : A ⊥} {z} →
- Rel k x y → Rel k y (later z) → Rel k x (later z)
- later-trans (later x∼y) (later y∼z) = later (♯ trans (♭ x∼y) (♭ y∼z))
- later-trans (later x∼y) (laterˡ y∼lz) = later (♯ trans (♭ x∼y) (laterʳ⁻¹ y∼lz))
- later-trans (laterˡ x∼y) y∼lz = later (♯ trans x∼y (laterʳ⁻¹ y∼lz))
- later-trans (laterʳ x≈y) ly≈lz = later-trans x≈y (laterˡ⁻¹ ly≈lz)
- later-trans x≈y (laterʳ y≈z) = laterʳ ( trans x≈y y≈z )
-
- trans : ∀ {k} {x y z : A ⊥} → Rel k x y → Rel k y z → Rel k x z
- trans {z = now v} x∼y y∼v = now-trans x∼y y∼v
- trans {z = later z} x∼y y∼lz = later-trans x∼y y∼lz
-
- private module Pre {k} = Preorder (preorder k)
+ private
+ preorder′ : IsEquivalence _∼_ → Kind → Preorder _ _ _
+ preorder′ equiv =
+ preorder (Setoid.isPreorder (record { isEquivalence = equiv }))
-- The two equalities are equivalence relations.
- setoid : (k : Kind) {eq : Equality k} → Setoid _ _
- setoid k {eq} = record
+ setoid : IsEquivalence _∼_ →
+ (k : Kind) {eq : Equality k} → Setoid _ _
+ setoid equiv k {eq} = record
{ Carrier = A ⊥
; _≈_ = Rel k
; isEquivalence = record
{ refl = Pre.refl
- ; sym = sym eq
+ ; sym = Equivalence.sym (IsEquivalence.sym equiv) eq
; trans = Pre.trans
}
- }
- where
- sym : ∀ {k} {x y : A ⊥} → Equality k → Rel k x y → Rel k y x
- sym eq (now xRy) = now (R.sym xRy)
- sym eq (later x∼y) = later (♯ sym eq (♭ x∼y))
- sym eq (laterʳ x≈y) = laterˡ (sym eq x≈y )
- sym eq (laterˡ {weak} x≈y) = laterʳ (sym eq x≈y )
- sym () (laterˡ {geq} x≳y)
-
- private module S {k} {eq} = Setoid (setoid k {eq})
+ } where module Pre = Preorder (preorder′ equiv k)
-- The order is a partial order, with strong equality as the
-- underlying equality.
- ≳-poset : Poset _ _ _
- ≳-poset = record
+ ≳-poset : IsEquivalence _∼_ → Poset _ _ _
+ ≳-poset equiv = record
{ Carrier = A ⊥
; _≈_ = _≅_
; _≤_ = _≳_
@@ -306,8 +305,11 @@ module Properties (S : Setoid zero zero) where
}
}
where
+ module S = Setoid (setoid equiv strong)
+ module Pre = Preorder (preorder′ equiv (other geq))
+
antisym : {x y : A ⊥} → x ≳ y → x ≲ y → x ≅ y
- antisym (now xRy) (now _) = now xRy
+ antisym (now x∼y) (now _) = now x∼y
antisym (later x≳y) (later x≲y) = later (♯ antisym (♭ x≳y) (♭ x≲y))
antisym (later x≳y) (laterˡ x≲ly) = later (♯ antisym (♭ x≳y) (laterʳ⁻¹ x≲ly))
antisym (laterˡ x≳ly) (later x≲y) = later (♯ antisym (laterʳ⁻¹ x≳ly) (♭ x≲y))
@@ -315,7 +317,11 @@ module Properties (S : Setoid zero zero) where
-- Equational reasoning.
- module Reasoning where
+ module Reasoning (isEquivalence : IsEquivalence _∼_) where
+
+ private
+ module Pre {k} = Preorder (preorder′ isEquivalence k)
+ module S {k eq} = Setoid (setoid isEquivalence k {eq})
infix 2 _∎
infixr 2 _≅⟨_⟩_ _≳⟨_⟩_ _≈⟨_⟩_
@@ -345,16 +351,18 @@ module Properties (S : Setoid zero zero) where
now≉never : ∀ {k} {x : A} → ¬ Rel k (now x) never
now≉never (laterʳ hyp) = now≉never hyp
- -- A partial value is either now or never (classically).
+ -- A partial value is either now or never (classically, when the
+ -- underlying relation is reflexive).
- now-or-never : ∀ {k} (x : A ⊥) →
+ now-or-never : Reflexive _∼_ →
+ ∀ {k} (x : A ⊥) →
¬ ¬ ((∃ λ y → x ⇓[ other k ] y) ⊎ x ⇑[ other k ])
- now-or-never x = helper <$> excluded-middle
+ now-or-never refl x = helper <$> excluded-middle
where
open RawMonad ¬¬-Monad
not-now-is-never : (x : A ⊥) → (∄ λ y → x ≳ now y) → x ≳ never
- not-now-is-never (now x) hyp with hyp (, now R.refl)
+ not-now-is-never (now x) hyp with hyp (, now refl)
... | ()
not-now-is-never (later x) hyp =
later (♯ not-now-is-never (♭ x) (hyp ∘ Prod.map id laterˡ))
@@ -364,93 +372,112 @@ module Properties (S : Setoid zero zero) where
helper (no ≵now) = inj₂ $ ≳⇒ $ not-now-is-never x ≵now
------------------------------------------------------------------------
--- Lemmas related to steps
+-- Map-like results
- module Steps where
+ -- Map.
- open P.≡-Reasoning
- open Reasoning using (_≅⟨_⟩_)
+ map : ∀ {_∼′_ k} →
+ _∼′_ ⇒ _∼_ → Equality.Rel _∼′_ k ⇒ Equality.Rel _∼_ k
+ map ∼′⇒∼ (now x∼y) = now (∼′⇒∼ x∼y)
+ map ∼′⇒∼ (later x∼y) = later (♯ map ∼′⇒∼ (♭ x∼y))
+ map ∼′⇒∼ (laterʳ x≈y) = laterʳ (map ∼′⇒∼ x≈y)
+ map ∼′⇒∼ (laterˡ x∼y) = laterˡ (map ∼′⇒∼ x∼y)
- private
+ -- If a statement can be proved using propositional equality as the
+ -- underlying relation, then it can also be proved for any other
+ -- reflexive underlying relation.
- lemma : ∀ {k x y} {z : A}
- (x∼y : Rel (other k) (♭ x) y)
- (y⇓z : y ⇓[ other k ] z) →
- steps (Pre.trans (laterˡ {x = x} x∼y) y⇓z) ≡
- suc (steps (Pre.trans x∼y y⇓z))
- lemma x∼y (now yRz) = P.refl
- lemma x∼y (laterˡ y⇓z) = begin
- steps (Pre.trans (laterˡ (laterʳ⁻¹ x∼y)) y⇓z) ≡⟨ lemma (laterʳ⁻¹ x∼y) y⇓z ⟩
- suc (steps (Pre.trans (laterʳ⁻¹ x∼y) y⇓z)) ∎
-
- left-identity : ∀ {k x y} {z : A}
- (x≅y : x ≅ y) (y⇓z : y ⇓[ k ] z) →
- steps (x ≅⟨ x≅y ⟩ y⇓z) ≡ steps y⇓z
+ ≡⇒ : Reflexive _∼_ →
+ ∀ {k x y} → Equality.Rel _≡_ k x y → Rel k x y
+ ≡⇒ refl-∼ = map (flip (P.subst (_∼_ _)) refl-∼)
+
+------------------------------------------------------------------------
+-- Steps
+
+ -- The number of later constructors (steps) in the terminating
+ -- computation x.
+
+ steps : ∀ {k} {x : A ⊥} {y} → x ⇓[ k ] y → ℕ
+ steps (now _) = zero
+ steps .{x = later x} (laterˡ {x = x} x⇓) = suc (steps {x = ♭ x} x⇓)
+
+ module Steps {trans-∼ : Transitive _∼_} where
+
+ left-identity :
+ ∀ {k x y} {z : A}
+ (x≅y : x ≅ y) (y⇓z : y ⇓[ k ] z) →
+ steps (Equivalence.trans trans-∼ (≅⇒ x≅y) y⇓z) ≡ steps y⇓z
left-identity (now _) (now _) = P.refl
- left-identity (later x≅y) (laterˡ y⇓z) = begin
- steps (Pre.trans (laterˡ (≅⇒ (♭ x≅y))) y⇓z) ≡⟨ lemma (≅⇒ (♭ x≅y)) y⇓z ⟩
- suc (steps (_ ≅⟨ ♭ x≅y ⟩ y⇓z)) ≡⟨ P.cong suc $ left-identity (♭ x≅y) y⇓z ⟩
- suc (steps y⇓z) ∎
-
- right-identity : ∀ {k x} {y z : A}
- (x⇓y : x ⇓[ k ] y) (y≈z : now y ⇓[ k ] z) →
- steps (Pre.trans x⇓y y≈z) ≡ steps x⇓y
- right-identity (now xRy) (now yRz) = P.refl
- right-identity (laterˡ x∼y) (now yRz) =
- P.cong suc $ right-identity x∼y (now yRz)
+ left-identity (later x≅y) (laterˡ y⇓z) =
+ P.cong suc $ left-identity (♭ x≅y) y⇓z
+
+ right-identity :
+ ∀ {k x} {y z : A}
+ (x⇓y : x ⇓[ k ] y) (y≈z : now y ⇓[ k ] z) →
+ steps (Equivalence.trans trans-∼ x⇓y y≈z) ≡ steps x⇓y
+ right-identity (now x∼y) (now y∼z) = P.refl
+ right-identity (laterˡ x∼y) (now y∼z) =
+ P.cong suc $ right-identity x∼y (now y∼z)
------------------------------------------------------------------------
-- Laws related to bind
-- Never is a left and right "zero" of bind.
- left-zero : {B : Set} (f : B → A ⊥) → let open M in
+ left-zero : ∀ {B} (f : B → A ⊥) → let open M in
(never >>= f) ≅ never
left-zero f = later (♯ left-zero f)
- right-zero : {B : Set} (x : B ⊥) → let open M in
+ right-zero : ∀ {B} (x : B ⊥) → let open M in
(x >>= λ _ → never) ≅ never
- right-zero (now x) = S.refl
right-zero (later x) = later (♯ right-zero (♭ x))
+ right-zero (now x) = never≅never
+ where never≅never = later (♯ never≅never)
- -- Now is a left and right identity of bind.
+ -- Now is a left and right identity of bind (for a reflexive
+ -- underlying relation).
- left-identity : {B : Set} (x : B) (f : B → A ⊥) → let open M in
+ left-identity : Reflexive _∼_ →
+ ∀ {B} (x : B) (f : B → A ⊥) → let open M in
(now x >>= f) ≅ f x
- left-identity x f = S.refl
+ left-identity refl-∼ x f = Equivalence.refl refl-∼
- right-identity : (x : A ⊥) → let open M in
+ right-identity : Reflexive _∼_ →
+ (x : A ⊥) → let open M in
(x >>= now) ≅ x
- right-identity (now x) = now R.refl
- right-identity (later x) = later (♯ right-identity (♭ x))
+ right-identity refl (now x) = now refl
+ right-identity refl (later x) = later (♯ right-identity refl (♭ x))
- -- Bind is associative.
+ -- Bind is associative (for a reflexive underlying relation).
- associative : {B C : Set} (x : C ⊥) (f : C → B ⊥) (g : B → A ⊥) →
+ associative : Reflexive _∼_ →
+ ∀ {B C} (x : C ⊥) (f : C → B ⊥) (g : B → A ⊥) →
let open M in
(x >>= f >>= g) ≅ (x >>= λ y → f y >>= g)
- associative (now x) f g = S.refl
- associative (later x) f g = later (♯ associative (♭ x) f g)
+ associative refl-∼ (now x) f g = Equivalence.refl refl-∼
+ associative refl-∼ (later x) f g =
+ later (♯ associative refl-∼ (♭ x) f g)
-module Properties₂ (S₁ S₂ : Setoid zero zero) where
+open Dummy public
+
+private
+ module Dummy₂ {s ℓ} {A B : Set s}
+ {_∼A_ : A → A → Set ℓ}
+ {_∼B_ : B → B → Set ℓ} where
- open Setoid S₁ renaming (Carrier to A; _≈_ to _≈A_)
- open Setoid S₂ renaming (Carrier to B; _≈_ to _≈B_)
open Equality
private
- open module EqA = Equality _≈A_ using () renaming (_⇓[_]_ to _⇓[_]A_; _⇑[_] to _⇑[_]A)
- open module EqB = Equality _≈B_ using () renaming (_⇓[_]_ to _⇓[_]B_; _⇑[_] to _⇑[_]B)
- module PropA = Properties S₁
- open PropA.Reasoning
+ open module EqA = Equality _∼A_ using () renaming (_⇓[_]_ to _⇓[_]A_; _⇑[_] to _⇑[_]A)
+ open module EqB = Equality _∼B_ using () renaming (_⇓[_]_ to _⇓[_]B_; _⇑[_] to _⇑[_]B)
-- Bind preserves all the relations.
_>>=-cong_ :
∀ {k} {x₁ x₂ : A ⊥} {f₁ f₂ : A → B ⊥} → let open M in
- EqA.Rel k x₁ x₂ →
- (∀ {x₁ x₂} → x₁ ≈A x₂ → EqB.Rel k (f₁ x₁) (f₂ x₂)) →
- EqB.Rel k (x₁ >>= f₁) (x₂ >>= f₂)
- now x₁Rx₂ >>=-cong f₁∼f₂ = f₁∼f₂ x₁Rx₂
+ Rel _∼A_ k x₁ x₂ →
+ (∀ {x₁ x₂} → x₁ ∼A x₂ → Rel _∼B_ k (f₁ x₁) (f₂ x₂)) →
+ Rel _∼B_ k (x₁ >>= f₁) (x₂ >>= f₂)
+ now x₁∼x₂ >>=-cong f₁∼f₂ = f₁∼f₂ x₁∼x₂
later x₁∼x₂ >>=-cong f₁∼f₂ = later (♯ (♭ x₁∼x₂ >>=-cong f₁∼f₂))
laterʳ x₁≈x₂ >>=-cong f₁≈f₂ = laterʳ (x₁≈x₂ >>=-cong f₁≈f₂)
laterˡ x₁∼x₂ >>=-cong f₁∼f₂ = laterˡ (x₁∼x₂ >>=-cong f₁∼f₂)
@@ -458,63 +485,67 @@ module Properties₂ (S₁ S₂ : Setoid zero zero) where
-- Inversion lemmas for bind.
>>=-inversion-⇓ :
+ Reflexive _∼A_ →
∀ {k} x {f : A → B ⊥} {y} → let open M in
(x>>=f⇓ : (x >>= f) ⇓[ k ]B y) →
∃ λ z → ∃₂ λ (x⇓ : x ⇓[ k ]A z) (fz⇓ : f z ⇓[ k ]B y) →
- EqA.steps x⇓ + EqB.steps fz⇓ ≡ EqB.steps x>>=f⇓
- >>=-inversion-⇓ (now x) fx⇓ =
- (x , now (Setoid.refl S₁) , fx⇓ , P.refl)
- >>=-inversion-⇓ (later x) (laterˡ x>>=f⇓) =
+ steps x⇓ + steps fz⇓ ≡ steps x>>=f⇓
+ >>=-inversion-⇓ refl (now x) fx⇓ =
+ (x , now refl , fx⇓ , P.refl)
+ >>=-inversion-⇓ refl (later x) (laterˡ x>>=f⇓) =
Prod.map id (Prod.map laterˡ (Prod.map id (P.cong suc))) $
- >>=-inversion-⇓ (♭ x) x>>=f⇓
+ >>=-inversion-⇓ refl (♭ x) x>>=f⇓
- >>=-inversion-⇑ : ∀ {k} x {f : A → B ⊥} → let open M in
- EqB.Rel (other k) (x >>= f) never →
+ >>=-inversion-⇑ : IsEquivalence _∼A_ →
+ ∀ {k} x {f : A → B ⊥} → let open M in
+ Rel _∼B_ (other k) (x >>= f) never →
¬ ¬ (x ⇑[ other k ]A ⊎
∃ λ y → x ⇓[ other k ]A y × f y ⇑[ other k ]B)
- >>=-inversion-⇑ {k} x {f} ∼never = helper <$> PropA.now-or-never x
+ >>=-inversion-⇑ eqA {k} x {f} ∼never =
+ helper <$> now-or-never IsEqA.refl x
where
open RawMonad ¬¬-Monad using (_<$>_)
open M using (_>>=_)
+ open Reasoning eqA
+ module IsEqA = IsEquivalence eqA
k≳ = other geq
is-never : ∀ {x y} →
x ⇓[ k≳ ]A y → (x >>= f) ⇑[ k≳ ]B →
- ∃ λ z → y ≈A z × f z ⇑[ k≳ ]B
- is-never (now xRy) = λ fx⇑ → (_ , Setoid.sym S₁ xRy , fx⇑)
- is-never (laterˡ ≳now) = is-never ≳now ∘ EqB.later⁻¹
+ ∃ λ z → y ∼A z × f z ⇑[ k≳ ]B
+ is-never (now x∼y) = λ fx⇑ → (_ , IsEqA.sym x∼y , fx⇑)
+ is-never (laterˡ ≳now) = is-never ≳now ∘ later⁻¹
helper : (∃ λ y → x ⇓[ k≳ ]A y) ⊎ x ⇑[ k≳ ]A →
x ⇑[ other k ]A ⊎
∃ λ y → x ⇓[ other k ]A y × f y ⇑[ other k ]B
- helper (inj₂ ≳never) = inj₁ (EqA.≳⇒ ≳never)
- helper (inj₁ (y , ≳now)) with is-never ≳now (EqB.never⇒never ∼never)
- ... | (z , yRz , fz⇑) = inj₂ (z , EqA.≳⇒ (x ≳⟨ ≳now ⟩
- now y ≅⟨ now yRz ⟩
- now z ∎)
- , EqB.≳⇒ fz⇑)
+ helper (inj₂ ≳never) = inj₁ (≳⇒ ≳never)
+ helper (inj₁ (y , ≳now)) with is-never ≳now (never⇒never ∼never)
+ ... | (z , y∼z , fz⇑) = inj₂ (z , ≳⇒ (x ≳⟨ ≳now ⟩
+ now y ≅⟨ now y∼z ⟩
+ now z ∎)
+ , ≳⇒ fz⇑)
-------------------------------------------------------------------------
--- Instantiation of the modules above using propositional equality
+open Dummy₂ public
-module Propositional where
- private
- open module Eq {A : Set} = Equality (_≡_ {A = A}) public
- open module P₁ {A : Set} = Properties (P.setoid A) public
- open module P₂ {A B : Set} =
- Properties₂ (P.setoid A) (P.setoid B) public
+private
+ module Dummy₃ {ℓ} {A B : Set ℓ} {_∼_ : B → B → Set ℓ} where
- -- If a statement can be proved using propositional equality as the
- -- underlying relation, then it can also be proved for any other
- -- reflexive underlying relation.
+ open Equality
+
+ -- A variant of _>>=-cong_.
- ≡⇒ : ∀ {A : Set} {_≈_ : A → A → Set} → Reflexive _≈_ →
- ∀ {k x y} → Rel k x y → Equality.Rel _≈_ k x y
- ≡⇒ refl (now P.refl) = Equality.now refl
- ≡⇒ refl (later x∼y) = Equality.later (♯ ≡⇒ refl (♭ x∼y))
- ≡⇒ refl (laterʳ x≈y) = Equality.laterʳ (≡⇒ refl x≈y)
- ≡⇒ refl (laterˡ x∼y) = Equality.laterˡ (≡⇒ refl x∼y)
+ _≡->>=-cong_ :
+ ∀ {k} {x₁ x₂ : A ⊥} {f₁ f₂ : A → B ⊥} → let open M in
+ Rel _≡_ k x₁ x₂ →
+ (∀ x → Rel _∼_ k (f₁ x) (f₂ x)) →
+ Rel _∼_ k (x₁ >>= f₁) (x₂ >>= f₂)
+ _≡->>=-cong_ {k} {f₁ = f₁} {f₂} x₁≈x₂ f₁≈f₂ =
+ x₁≈x₂ >>=-cong λ {x} x≡x′ →
+ P.subst (λ y → Rel _∼_ k (f₁ x) (f₂ y)) x≡x′ (f₁≈f₂ x)
+
+open Dummy₃ public
------------------------------------------------------------------------
-- Productivity checker workaround
@@ -522,18 +553,18 @@ module Propositional where
-- The monad can be awkward to use, due to the limitations of guarded
-- coinduction. The following code provides a (limited) workaround.
-module Workaround where
+module Workaround {a} where
infixl 1 _>>=_
- data _⊥P : Set → Set₁ where
+ data _⊥P : Set a → Set (suc a) where
now : ∀ {A} (x : A) → A ⊥P
later : ∀ {A} (x : ∞ (A ⊥P)) → A ⊥P
_>>=_ : ∀ {A B} (x : A ⊥P) (f : A → B ⊥P) → B ⊥P
private
- data _⊥W : Set → Set₁ where
+ data _⊥W : Set a → Set (suc a) where
now : ∀ {A} (x : A) → A ⊥W
later : ∀ {A} (x : A ⊥P) → A ⊥W
@@ -564,8 +595,9 @@ module Workaround where
module Correct where
- open Propositional
- open Reasoning
+ private
+ open module Eq {A} = Equality {A = A} _≡_
+ open module R {A} = Reasoning {A = A} P.isEquivalence
now-hom : ∀ {A} (x : A) → ⟦ now x ⟧P ≅ now x
now-hom x = now x ∎
@@ -590,11 +622,11 @@ module Workaround where
------------------------------------------------------------------------
-- An alternative, but equivalent, formulation of equality/ordering
-module AlternativeEquality where
+module AlternativeEquality {a ℓ} where
private
- El : Setoid zero zero → Set
+ El : Setoid a ℓ → Set _
El = Setoid.Carrier
Eq : ∀ S → B.Rel (El S) _
@@ -621,14 +653,14 @@ module AlternativeEquality where
_∣_≈P_ : ∀ S → B.Rel (El S ⊥) _
_∣_≈P_ = flip RelP (other weak)
- data RelP S : Kind → B.Rel (El S ⊥) (suc zero) where
+ data RelP S : Kind → B.Rel (El S ⊥) (suc (a ⊔ ℓ)) where
-- Congruences.
now : ∀ {k x y} (xRy : x ⟨ Eq S ⟩ y) → RelP S k (now x) (now y)
later : ∀ {k x y} (x∼y : ∞ (RelP S k (♭ x) (♭ y))) →
RelP S k (later x) (later y)
- _>>=_ : ∀ {S′ : Setoid zero zero} {k} {x₁ x₂}
+ _>>=_ : ∀ {S′ : Setoid a ℓ} {k} {x₁ x₂}
{f₁ f₂ : El S′ → El S ⊥} →
let open M in
(x₁∼x₂ : RelP S′ k x₁ x₂)
@@ -689,7 +721,7 @@ module AlternativeEquality where
-- Proof WHNFs.
- data RelW S : Kind → B.Rel (El S ⊥) (suc zero) where
+ data RelW S : Kind → B.Rel (El S ⊥) (suc (a ⊔ ℓ)) where
now : ∀ {k x y} (xRy : x ⟨ Eq S ⟩ y) → RelW S k (now x) (now y)
later : ∀ {k x y} (x∼y : RelP S k (♭ x) (♭ y)) → RelW S k (later x) (later y)
laterʳ : ∀ {x y} (x≈y : RelW S (other weak) x (♭ y)) → RelW S (other weak) x (later y)
diff --git a/src/Category/Monad/State.agda b/src/Category/Monad/State.agda
index 7626540..670a910 100644
--- a/src/Category/Monad/State.agda
+++ b/src/Category/Monad/State.agda
@@ -2,23 +2,26 @@
-- The state monad
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Category.Monad.State where
+open import Category.Applicative.Indexed
open import Category.Monad
-open import Category.Monad.Indexed
open import Category.Monad.Identity
+open import Category.Monad.Indexed
open import Data.Product
-open import Function
open import Data.Unit
-open import Category.Applicative.Indexed
+open import Function
+open import Level
------------------------------------------------------------------------
-- Indexed state monads
-IStateT : {I : Set} → (I → Set) → (Set → Set) → IFun I
+IStateT : ∀ {i f} {I : Set i} → (I → Set f) → (Set f → Set f) → IFun I f
IStateT S M i j A = S i → M (A × S j)
-StateTIMonad : ∀ {I} (S : I → Set) {M} →
+StateTIMonad : ∀ {i f} {I : Set i} (S : I → Set f) {M} →
RawMonad M → RawIMonad (IStateT S M)
StateTIMonad S Mon = record
{ return = λ x s → return (x , s)
@@ -26,7 +29,7 @@ StateTIMonad S Mon = record
}
where open RawMonad Mon
-StateTIMonadZero : ∀ {I} (S : I → Set) {M} →
+StateTIMonadZero : ∀ {i f} {I : Set i} (S : I → Set f) {M} →
RawMonadZero M → RawIMonadZero (IStateT S M)
StateTIMonadZero S Mon = record
{ monad = StateTIMonad S (RawMonadZero.monad Mon)
@@ -34,7 +37,7 @@ StateTIMonadZero S Mon = record
}
where open RawMonadZero Mon
-StateTIMonadPlus : ∀ {I} (S : I → Set) {M} →
+StateTIMonadPlus : ∀ {i f} {I : Set i} (S : I → Set f) {M} →
RawMonadPlus M → RawIMonadPlus (IStateT S M)
StateTIMonadPlus S Mon = record
{ monadZero = StateTIMonadZero S (RawIMonadPlus.monadZero Mon)
@@ -45,19 +48,19 @@ StateTIMonadPlus S Mon = record
------------------------------------------------------------------------
-- State monad operations
-record RawIMonadState {I : Set} (S : I → Set)
- (M : I → I → Set → Set) : Set₁ where
+record RawIMonadState {i f} {I : Set i} (S : I → Set f)
+ (M : IFun I f) : Set (i ⊔ suc f) where
field
monad : RawIMonad M
get : ∀ {i} → M i i (S i)
- put : ∀ {i j} → S j → M i j ⊤
+ put : ∀ {i j} → S j → M i j (Lift ⊤)
open RawIMonad monad public
- modify : ∀ {i j} → (S i → S j) → M i j ⊤
+ modify : ∀ {i j} → (S i → S j) → M i j (Lift ⊤)
modify f = get >>= put ∘ f
-StateTIMonadState : ∀ {I} (S : I → Set) {M} →
+StateTIMonadState : ∀ {i f} {I : Set i} (S : I → Set f) {M} →
RawMonad M → RawIMonadState S (IStateT S M)
StateTIMonadState S Mon = record
{ monad = StateTIMonad S Mon
@@ -69,38 +72,41 @@ StateTIMonadState S Mon = record
------------------------------------------------------------------------
-- Ordinary state monads
-RawMonadState : Set → (Set → Set) → Set₁
-RawMonadState S M = RawIMonadState {⊤} (λ _ → S) (λ _ _ → M)
+RawMonadState : ∀ {f} → Set f → (Set f → Set f) → Set _
+RawMonadState S M = RawIMonadState {I = ⊤} (λ _ → S) (λ _ _ → M)
-module RawMonadState {S : Set} {M : Set → Set}
+module RawMonadState {f} {S : Set f} {M : Set f → Set f}
(Mon : RawMonadState S M) where
open RawIMonadState Mon public
-StateT : Set → (Set → Set) → Set → Set
-StateT S M = IStateT {⊤} (λ _ → S) M _ _
+StateT : ∀ {f} → Set f → (Set f → Set f) → Set f → Set f
+StateT S M = IStateT {I = ⊤} (λ _ → S) M _ _
-StateTMonad : ∀ S {M} → RawMonad M → RawMonad (StateT S M)
+StateTMonad : ∀ {f} (S : Set f) {M} → RawMonad M → RawMonad (StateT S M)
StateTMonad S = StateTIMonad (λ _ → S)
-StateTMonadZero : ∀ S {M} → RawMonadZero M → RawMonadZero (StateT S M)
+StateTMonadZero : ∀ {f} (S : Set f) {M} →
+ RawMonadZero M → RawMonadZero (StateT S M)
StateTMonadZero S = StateTIMonadZero (λ _ → S)
-StateTMonadPlus : ∀ S {M} → RawMonadPlus M → RawMonadPlus (StateT S M)
+StateTMonadPlus : ∀ {f} (S : Set f) {M} →
+ RawMonadPlus M → RawMonadPlus (StateT S M)
StateTMonadPlus S = StateTIMonadPlus (λ _ → S)
-StateTMonadState : ∀ S {M} → RawMonad M → RawMonadState S (StateT S M)
+StateTMonadState : ∀ {f} (S : Set f) {M} →
+ RawMonad M → RawMonadState S (StateT S M)
StateTMonadState S = StateTIMonadState (λ _ → S)
-State : Set → Set → Set
+State : ∀ {f} → Set f → Set f → Set f
State S = StateT S Identity
-StateMonad : (S : Set) → RawMonad (State S)
+StateMonad : ∀ {f} (S : Set f) → RawMonad (State S)
StateMonad S = StateTMonad S IdentityMonad
-StateMonadState : ∀ S → RawMonadState S (State S)
+StateMonadState : ∀ {f} (S : Set f) → RawMonadState S (State S)
StateMonadState S = StateTMonadState S IdentityMonad
-LiftMonadState : ∀ {S₁} S₂ {M} →
+LiftMonadState : ∀ {f S₁} (S₂ : Set f) {M} →
RawMonadState S₁ M →
RawMonadState S₁ (StateT S₂ M)
LiftMonadState S₂ Mon = record
diff --git a/src/Data/BoundedVec/Inefficient.agda b/src/Data/BoundedVec/Inefficient.agda
index a73a6a2..93efcc1 100644
--- a/src/Data/BoundedVec/Inefficient.agda
+++ b/src/Data/BoundedVec/Inefficient.agda
@@ -2,6 +2,8 @@
-- Bounded vectors (inefficient, concrete implementation)
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- Vectors of a specified maximum length.
module Data.BoundedVec.Inefficient where
@@ -14,26 +16,26 @@ open import Data.List
infixr 5 _∷_
-data BoundedVec (a : Set) : ℕ → Set where
- [] : ∀ {n} → BoundedVec a n
- _∷_ : ∀ {n} (x : a) (xs : BoundedVec a n) → BoundedVec a (suc n)
+data BoundedVec {a} (A : Set a) : ℕ → Set a where
+ [] : ∀ {n} → BoundedVec A n
+ _∷_ : ∀ {n} (x : A) (xs : BoundedVec A n) → BoundedVec A (suc n)
------------------------------------------------------------------------
-- Increasing the bound
-- Note that this operation is linear in the length of the list.
-↑ : ∀ {a n} → BoundedVec a n → BoundedVec a (suc n)
+↑ : ∀ {a n} {A : Set a} → BoundedVec A n → BoundedVec A (suc n)
↑ [] = []
↑ (x ∷ xs) = x ∷ ↑ xs
------------------------------------------------------------------------
-- Conversions
-fromList : ∀ {a} → (xs : List a) → BoundedVec a (length xs)
+fromList : ∀ {a} {A : Set a} → (xs : List A) → BoundedVec A (length xs)
fromList [] = []
fromList (x ∷ xs) = x ∷ fromList xs
-toList : ∀ {a n} → BoundedVec a n → List a
+toList : ∀ {a n} {A : Set a} → BoundedVec A n → List A
toList [] = []
toList (x ∷ xs) = x ∷ toList xs
diff --git a/src/Data/Colist.agda b/src/Data/Colist.agda
index 2e8be02..62a2dcf 100644
--- a/src/Data/Colist.agda
+++ b/src/Data/Colist.agda
@@ -2,6 +2,8 @@
-- Coinductive lists
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Data.Colist where
open import Category.Monad
@@ -19,6 +21,7 @@ open import Data.BoundedVec.Inefficient as BVec
open import Data.Product using (_,_)
open import Data.Sum using (_⊎_; inj₁; inj₂)
open import Function
+open import Level
open import Relation.Binary
import Relation.Binary.InducedPreorders as Ind
open import Relation.Binary.PropositionalEquality using (_≡_)
@@ -31,63 +34,65 @@ open RawMonad ¬¬-Monad
infixr 5 _∷_
-data Colist (A : Set) : Set where
+data Colist {a} (A : Set a) : Set a where
[] : Colist A
_∷_ : (x : A) (xs : ∞ (Colist A)) → Colist A
-data Any {A} (P : A → Set) : Colist A → Set where
+data Any {a p} {A : Set a} (P : A → Set p) :
+ Colist A → Set (a ⊔ p) where
here : ∀ {x xs} (px : P x) → Any P (x ∷ xs)
there : ∀ {x xs} (pxs : Any P (♭ xs)) → Any P (x ∷ xs)
-data All {A} (P : A → Set) : Colist A → Set where
+data All {a p} {A : Set a} (P : A → Set p) :
+ Colist A → Set (a ⊔ p) where
[] : All P []
_∷_ : ∀ {x xs} (px : P x) (pxs : ∞ (All P (♭ xs))) → All P (x ∷ xs)
------------------------------------------------------------------------
-- Some operations
-null : ∀ {A} → Colist A → Bool
+null : ∀ {a} {A : Set a} → Colist A → Bool
null [] = true
null (_ ∷ _) = false
-length : ∀ {A} → Colist A → Coℕ
+length : ∀ {a} {A : Set a} → Colist A → Coℕ
length [] = zero
length (x ∷ xs) = suc (♯ length (♭ xs))
-map : ∀ {A B} → (A → B) → Colist A → Colist B
+map : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → Colist A → Colist B
map f [] = []
map f (x ∷ xs) = f x ∷ ♯ map f (♭ xs)
-fromList : ∀ {A} → List A → Colist A
+fromList : ∀ {a} {A : Set a} → List A → Colist A
fromList [] = []
fromList (x ∷ xs) = x ∷ ♯ fromList xs
-take : ∀ {A} (n : ℕ) → Colist A → BoundedVec A n
+take : ∀ {a} {A : Set a} (n : ℕ) → Colist A → BoundedVec A n
take zero xs = []
take (suc n) [] = []
take (suc n) (x ∷ xs) = x ∷ take n (♭ xs)
-replicate : ∀ {A} → Coℕ → A → Colist A
+replicate : ∀ {a} {A : Set a} → Coℕ → A → Colist A
replicate zero x = []
replicate (suc n) x = x ∷ ♯ replicate (♭ n) x
-lookup : ∀ {A} → ℕ → Colist A → Maybe A
+lookup : ∀ {a} {A : Set a} → ℕ → Colist A → Maybe A
lookup n [] = nothing
lookup zero (x ∷ xs) = just x
lookup (suc n) (x ∷ xs) = lookup n (♭ xs)
infixr 5 _++_
-_++_ : ∀ {A} → Colist A → Colist A → Colist A
+_++_ : ∀ {a} {A : Set a} → Colist A → Colist A → Colist A
[] ++ ys = ys
(x ∷ xs) ++ ys = x ∷ ♯ (♭ xs ++ ys)
-concat : ∀ {A} → Colist (List⁺ A) → Colist A
+concat : ∀ {a} {A : Set a} → Colist (List⁺ A) → Colist A
concat [] = []
concat ([ x ]⁺ ∷ xss) = x ∷ ♯ concat (♭ xss)
concat ((x ∷ xs) ∷ xss) = x ∷ ♯ concat (xs ∷ xss)
-[_] : ∀ {A} → A → Colist A
+[_] : ∀ {a} {A : Set a} → A → Colist A
[ x ] = x ∷ ♯ []
------------------------------------------------------------------------
@@ -97,13 +102,13 @@ concat ((x ∷ xs) ∷ xss) = x ∷ ♯ concat (xs ∷ xss)
infix 4 _≈_
-data _≈_ {A} : (xs ys : Colist A) → Set where
+data _≈_ {a} {A : Set a} : (xs ys : Colist A) → Set a where
[] : [] ≈ []
_∷_ : ∀ x {xs ys} (xs≈ : ∞ (♭ xs ≈ ♭ ys)) → x ∷ xs ≈ x ∷ ys
-- The equality relation forms a setoid.
-setoid : Set → Setoid _ _
+setoid : ∀ {ℓ} → Set ℓ → Setoid _ _
setoid A = record
{ Carrier = Colist A
; _≈_ = _≈_
@@ -129,11 +134,12 @@ setoid A = record
module ≈-Reasoning where
import Relation.Binary.EqReasoning as EqR
private
- open module R {A : Set} = EqR (setoid A) public
+ open module R {a} {A : Set a} = EqR (setoid A) public
-- map preserves equality.
-map-cong : ∀ {A B} (f : A → B) → _≈_ =[ map f ]⇒ _≈_
+map-cong : ∀ {a b} {A : Set a} {B : Set b}
+ (f : A → B) → _≈_ =[ map f ]⇒ _≈_
map-cong f [] = []
map-cong f (x ∷ xs≈) = f x ∷ ♯ map-cong f (♭ xs≈)
@@ -144,34 +150,34 @@ map-cong f (x ∷ xs≈) = f x ∷ ♯ map-cong f (♭ xs≈)
infix 4 _∈_
-_∈_ : {A : Set} → A → Colist A → Set
+_∈_ : ∀ {a} → {A : Set a} → A → Colist A → Set a
x ∈ xs = Any (_≡_ x) xs
-- xs ⊆ ys means that xs is a subset of ys.
infix 4 _⊆_
-_⊆_ : {A : Set} → Colist A → Colist A → Set
+_⊆_ : ∀ {a} → {A : Set a} → Colist A → Colist A → Set a
xs ⊆ ys = ∀ {x} → x ∈ xs → x ∈ ys
-- xs ⊑ ys means that xs is a prefix of ys.
infix 4 _⊑_
-data _⊑_ {A : Set} : Colist A → Colist A → Set where
+data _⊑_ {a} {A : Set a} : Colist A → Colist A → Set a where
[] : ∀ {ys} → [] ⊑ ys
_∷_ : ∀ x {xs ys} (p : ∞ (♭ xs ⊑ ♭ ys)) → x ∷ xs ⊑ x ∷ ys
-- Prefixes are subsets.
-⊑⇒⊆ : {A : Set} → _⊑_ {A = A} ⇒ _⊆_
+⊑⇒⊆ : ∀ {a} → {A : Set a} → _⊑_ {A = A} ⇒ _⊆_
⊑⇒⊆ [] ()
⊑⇒⊆ (x ∷ xs⊑ys) (here ≡x) = here ≡x
⊑⇒⊆ (_ ∷ xs⊑ys) (there x∈xs) = there (⊑⇒⊆ (♭ xs⊑ys) x∈xs)
-- The prefix relation forms a poset.
-⊑-Poset : Set → Poset _ _ _
+⊑-Poset : ∀ {ℓ} → Set ℓ → Poset _ _ _
⊑-Poset A = record
{ Carrier = Colist A
; _≈_ = _≈_
@@ -201,33 +207,33 @@ data _⊑_ {A : Set} : Colist A → Colist A → Set where
module ⊑-Reasoning where
import Relation.Binary.PartialOrderReasoning as POR
private
- open module R {A : Set} = POR (⊑-Poset A)
+ open module R {a} {A : Set a} = POR (⊑-Poset A)
public renaming (_≤⟨_⟩_ to _⊑⟨_⟩_)
-- The subset relation forms a preorder.
-⊆-Preorder : Set → Preorder _ _ _
+⊆-Preorder : ∀ {ℓ} → Set ℓ → Preorder _ _ _
⊆-Preorder A =
Ind.InducedPreorder₂ (setoid A) _∈_
- (λ xs≈ys → ⊑⇒⊆ $ ⊑P.reflexive xs≈ys)
+ (λ xs≈ys → ⊑⇒⊆ (⊑P.reflexive xs≈ys))
where module ⊑P = Poset (⊑-Poset A)
module ⊆-Reasoning where
import Relation.Binary.PreorderReasoning as PreR
private
- open module R {A : Set} = PreR (⊆-Preorder A)
+ open module R {a} {A : Set a} = PreR (⊆-Preorder A)
public renaming (_∼⟨_⟩_ to _⊆⟨_⟩_)
infix 1 _∈⟨_⟩_
- _∈⟨_⟩_ : ∀ {A : Set} (x : A) {xs ys} →
+ _∈⟨_⟩_ : ∀ {a} {A : Set a} (x : A) {xs ys} →
x ∈ xs → xs IsRelatedTo ys → x ∈ ys
x ∈⟨ x∈xs ⟩ xs⊆ys = (begin xs⊆ys) x∈xs
-- take returns a prefix.
-take-⊑ : ∀ {A} n (xs : Colist A) →
- let toColist = fromList ∘ BVec.toList in
+take-⊑ : ∀ {a} {A : Set a} n (xs : Colist A) →
+ let toColist = fromList {a} ∘ BVec.toList in
toColist (take n xs) ⊑ xs
take-⊑ zero xs = []
take-⊑ (suc n) [] = []
@@ -238,25 +244,27 @@ take-⊑ (suc n) (x ∷ xs) = x ∷ ♯ take-⊑ n (♭ xs)
-- Finite xs means that xs has finite length.
-data Finite {A : Set} : Colist A → Set where
+data Finite {a} {A : Set a} : Colist A → Set a where
[] : Finite []
_∷_ : ∀ x {xs} (fin : Finite (♭ xs)) → Finite (x ∷ xs)
-- Infinite xs means that xs has infinite length.
-data Infinite {A : Set} : Colist A → Set where
+data Infinite {a} {A : Set a} : Colist A → Set a where
_∷_ : ∀ x {xs} (inf : ∞ (Infinite (♭ xs))) → Infinite (x ∷ xs)
-- Colists which are not finite are infinite.
not-finite-is-infinite :
- {A : Set} (xs : Colist A) → ¬ Finite xs → Infinite xs
+ ∀ {a} {A : Set a} (xs : Colist A) → ¬ Finite xs → Infinite xs
not-finite-is-infinite [] hyp with hyp []
... | ()
not-finite-is-infinite (x ∷ xs) hyp =
x ∷ ♯ not-finite-is-infinite (♭ xs) (hyp ∘ _∷_ x)
-- Colists are either finite or infinite (classically).
+--
+-- TODO: Make this definition universe polymorphic.
finite-or-infinite :
{A : Set} (xs : Colist A) → ¬ ¬ (Finite xs ⊎ Infinite xs)
@@ -269,7 +277,7 @@ finite-or-infinite xs = helper <$> excluded-middle
-- Colists are not both finite and infinite.
not-finite-and-infinite :
- {A : Set} {xs : Colist A} → Finite xs → Infinite xs → ⊥
+ ∀ {a} {A : Set a} {xs : Colist A} → Finite xs → Infinite xs → ⊥
not-finite-and-infinite [] ()
not-finite-and-infinite (x ∷ fin) (.x ∷ inf) =
not-finite-and-infinite fin (♭ inf)
diff --git a/src/Data/Container.agda b/src/Data/Container.agda
index 5f89f9c..bdc148c 100644
--- a/src/Data/Container.agda
+++ b/src/Data/Container.agda
@@ -37,22 +37,21 @@ open Container public
⟦_⟧ : ∀ {ℓ} → Container ℓ → Set ℓ → Set ℓ
⟦ C ⟧ X = Σ[ s ∶ Shape C ] (Position C s → X)
--- Equality (based on propositional equality).
+-- Equality, parametrised on an underlying relation.
-infix 4 _≈_
-
-_≈_ : ∀ {c} {C : Container c} {X : Set c} → ⟦ C ⟧ X → ⟦ C ⟧ X → Set c
-_≈_ {C = C} (s , f) (s′ , f′) =
- Σ[ eq ∶ s ≡ s′ ] (∀ p → f p ≡ f′ (P.subst (Position C) eq p))
+Eq : ∀ {c ℓ} {C : Container c} {X Y : Set c} →
+ (X → Y → Set ℓ) → ⟦ C ⟧ X → ⟦ C ⟧ Y → Set (c ⊔ ℓ)
+Eq {C = C} _≈_ (s , f) (s′ , f′) =
+ Σ[ eq ∶ s ≡ s′ ] (∀ p → f p ≈ f′ (P.subst (Position C) eq p))
private
- -- Note that, if propositional equality were extensional, then _≈_
- -- and _≡_ would coincide.
+ -- Note that, if propositional equality were extensional, then
+ -- Eq _≡_ and _≡_ would coincide.
- ≈⇒≡ : ∀ {c} {C : Container c} {X : Set c} {xs ys : ⟦ C ⟧ X} →
- P.Extensionality c c → xs ≈ ys → xs ≡ ys
- ≈⇒≡ {C = C} {X} ext (s≡s′ , f≈f′) = helper s≡s′ f≈f′
+ Eq⇒≡ : ∀ {c} {C : Container c} {X : Set c} {xs ys : ⟦ C ⟧ X} →
+ P.Extensionality c c → Eq _≡_ xs ys → xs ≡ ys
+ Eq⇒≡ {C = C} {X} ext (s≡s′ , f≈f′) = helper s≡s′ f≈f′
where
helper : {s s′ : Shape C} (eq : s ≡ s′) →
{f : Position C s → X}
@@ -61,38 +60,42 @@ private
_≡_ {A = ⟦ C ⟧ X} (s , f) (s′ , f′)
helper refl f≈f′ = P.cong (_,_ _) (ext f≈f′)
-setoid : ∀ {ℓ} → Container ℓ → Set ℓ → Setoid ℓ ℓ
+setoid : ∀ {ℓ} → Container ℓ → Setoid ℓ ℓ → Setoid ℓ ℓ
setoid C X = record
- { Carrier = ⟦ C ⟧ X
+ { Carrier = ⟦ C ⟧ X.Carrier
; _≈_ = _≈_
; isEquivalence = record
- { refl = (refl , λ _ → refl)
+ { refl = (refl , λ _ → X.refl)
; sym = sym
; trans = λ {_ _ zs} → trans zs
}
}
where
- sym : {xs ys : ⟦ C ⟧ X} → xs ≈ ys → ys ≈ xs
+ module X = Setoid X
+
+ _≈_ = Eq X._≈_
+
+ sym : {xs ys : ⟦ C ⟧ X.Carrier} → xs ≈ ys → ys ≈ xs
sym (eq , f) = helper eq f
where
helper : {s s′ : Shape C} (eq : s ≡ s′) →
- {f : Position C s → X}
- {f′ : Position C s′ → X} →
- (∀ p → f p ≡ f′ (P.subst (Position C) eq p)) →
- _≈_ {C = C} (s′ , f′) (s , f)
- helper refl eq = (refl , P.sym ⟨∘⟩ eq)
+ {f : Position C s → X.Carrier}
+ {f′ : Position C s′ → X.Carrier} →
+ (∀ p → X._≈_ (f p) (f′ $ P.subst (Position C) eq p)) →
+ (s′ , f′) ≈ (s , f)
+ helper refl eq = (refl , X.sym ⟨∘⟩ eq)
- trans : ∀ {xs ys : ⟦ C ⟧ X} zs → xs ≈ ys → ys ≈ zs → xs ≈ zs
+ trans : ∀ {xs ys : ⟦ C ⟧ X.Carrier} zs → xs ≈ ys → ys ≈ zs → xs ≈ zs
trans zs (eq₁ , f₁) (eq₂ , f₂) = helper eq₁ eq₂ (proj₂ zs) f₁ f₂
where
helper : {s s′ s″ : Shape C} (eq₁ : s ≡ s′) (eq₂ : s′ ≡ s″) →
- {f : Position C s → X}
- {f′ : Position C s′ → X} →
- (f″ : Position C s″ → X) →
- (∀ p → f p ≡ f′ (P.subst (Position C) eq₁ p)) →
- (∀ p → f′ p ≡ f″ (P.subst (Position C) eq₂ p)) →
- _≈_ {C = C} (s , f) (s″ , f″)
- helper refl refl _ eq₁ eq₂ = (refl , λ p → P.trans (eq₁ p) (eq₂ p))
+ {f : Position C s → X.Carrier}
+ {f′ : Position C s′ → X.Carrier} →
+ (f″ : Position C s″ → X.Carrier) →
+ (∀ p → X._≈_ (f p) (f′ $ P.subst (Position C) eq₁ p)) →
+ (∀ p → X._≈_ (f′ p) (f″ $ P.subst (Position C) eq₂ p)) →
+ (s , f) ≈ (s″ , f″)
+ helper refl refl _ eq₁ eq₂ = (refl , λ p → X.trans (eq₁ p) (eq₂ p))
------------------------------------------------------------------------
-- Functoriality
@@ -104,14 +107,16 @@ map f = Prod.map ⟨id⟩ (λ g → f ⟨∘⟩ g)
module Map where
- identity : ∀ {c} {C : Container c} {X}
- (xs : ⟦ C ⟧ X) → map ⟨id⟩ xs ≈ xs
- identity {C = C} xs = Setoid.refl (setoid C _)
+ identity : ∀ {c} {C : Container c} X →
+ let module X = Setoid X in
+ (xs : ⟦ C ⟧ X.Carrier) → Eq X._≈_ (map ⟨id⟩ xs) xs
+ identity {C = C} X xs = Setoid.refl (setoid C X)
- composition : ∀ {c} {C : Container c} {X Y Z}
- (f : Y → Z) (g : X → Y) (xs : ⟦ C ⟧ X) →
- map f (map g xs) ≈ map (f ⟨∘⟩ g) xs
- composition {C = C} f g xs = Setoid.refl (setoid C _)
+ composition : ∀ {c} {C : Container c} {X Y} Z →
+ let module Z = Setoid Z in
+ (f : Y → Z.Carrier) (g : X → Y) (xs : ⟦ C ⟧ X) →
+ Eq Z._≈_ (map f (map g xs)) (map (f ⟨∘⟩ g) xs)
+ composition {C = C} Z f g xs = Setoid.refl (setoid C Z)
------------------------------------------------------------------------
-- Container morphisms
@@ -137,9 +142,10 @@ module Morphism where
Natural : ∀ {c} {C₁ C₂ : Container c} →
(∀ {X} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X) → Set (suc c)
- Natural {C₁ = C₁} m =
- ∀ {X Y} (f : X → Y) (xs : ⟦ C₁ ⟧ X) →
- m (map f xs) ≈ map f (m xs)
+ Natural {c} {C₁} m =
+ ∀ {X} (Y : Setoid c c) → let module Y = Setoid Y in
+ (f : X → Y.Carrier) (xs : ⟦ C₁ ⟧ X) →
+ Eq Y._≈_ (m $ map f xs) (map f $ m xs)
-- Natural transformations.
@@ -150,15 +156,19 @@ module Morphism where
natural : ∀ {c} {C₁ C₂ : Container c}
(m : C₁ ⇒ C₂) → Natural ⟪ m ⟫
- natural {C₂ = C₂} m f xs = Setoid.refl (setoid C₂ _)
+ natural {C₂ = C₂} m Y f xs = Setoid.refl (setoid C₂ Y)
-- In fact, all natural functions of the right type are container
-- morphisms.
complete : ∀ {c} {C₁ C₂ : Container c} →
(nt : NT C₁ C₂) →
- ∃ λ m → ∀ {X} (xs : ⟦ C₁ ⟧ X) → proj₁ nt xs ≈ ⟪ m ⟫ xs
- complete (nt , nat) = (m , λ xs → nat (proj₂ xs) (proj₁ xs , ⟨id⟩))
+ ∃ λ m → (X : Setoid c c) →
+ let module X = Setoid X in
+ (xs : ⟦ C₁ ⟧ X.Carrier) →
+ Eq X._≈_ (proj₁ nt xs) (⟪ m ⟫ xs)
+ complete (nt , nat) =
+ (m , λ X xs → nat X (proj₂ xs) (proj₁ xs , ⟨id⟩))
where
m = record { shape = λ s → proj₁ (nt (s , ⟨id⟩))
; position = λ {s} → proj₂ (nt (s , ⟨id⟩))
diff --git a/src/Data/Container/AlternativeBagAndSetEquality.agda b/src/Data/Container/AlternativeBagAndSetEquality.agda
deleted file mode 100644
index 6ddcc2e..0000000
--- a/src/Data/Container/AlternativeBagAndSetEquality.agda
+++ /dev/null
@@ -1,145 +0,0 @@
-------------------------------------------------------------------------
--- Alternative definition of bag and set equality
-------------------------------------------------------------------------
-
-{-# OPTIONS --universe-polymorphism #-}
-
-module Data.Container.AlternativeBagAndSetEquality where
-
-open import Data.Container
-open import Data.Product as Prod
-open import Function
-open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence as Eq using (_⇔_; module Equivalent)
-open import Function.Inverse as Inv using (Isomorphism; module Inverse)
-open import Relation.Binary.HeterogeneousEquality as H using (_≅_; refl)
-open import Relation.Binary.PropositionalEquality as P using (_≡_; refl)
-
-open P.≡-Reasoning
-
--- Another way to define bag and set equality. Two containers xs and
--- ys are equal when viewed as bags if their sets of positions are
--- equipotent and the position functions map related positions to
--- equal values. For set equality the position sets do not need to be
--- equipotent, only equiinhabited.
-
-infix 4 _≈[_]′_
-
-_≈[_]′_ : ∀ {c} {C : Container c} {X : Set c} →
- ⟦ C ⟧ X → Kind → ⟦ C ⟧ X → Set c
-_≈[_]′_ {C = C} (s , f) k (s′ , f′) =
- ∃ λ (p₁≈p₂ : Isomorphism k (Position C s) (Position C s′)) →
- (∀ p → f p ≡ f′ (Equivalent.to (Inv.⇒⇔ p₁≈p₂) ⟨$⟩ p)) ×
- (∀ p → f′ p ≡ f (Equivalent.from (Inv.⇒⇔ p₁≈p₂) ⟨$⟩ p))
- -- The last component is unnecessary when k equals bag.
-
--- This definition is equivalent to the one in Data.Container.
-
-private
-
- ≈⇔≈′-set : ∀ {c} {C : Container c} {X : Set c} (xs ys : ⟦ C ⟧ X) →
- xs ≈[ set ] ys ⇔ xs ≈[ set ]′ ys
- ≈⇔≈′-set {c} xs ys = Eq.equivalent to from
- where
- to : xs ≈[ set ] ys → xs ≈[ set ]′ ys
- to xs≈ys =
- Eq.equivalent
- (λ p → proj₁ $ Equivalent.to xs≈ys ⟨$⟩ (p , refl))
- (λ p → proj₁ $ Equivalent.from xs≈ys ⟨$⟩ (p , refl)) ,
- (λ p → proj₂ (Equivalent.to xs≈ys ⟨$⟩ (p , refl))) ,
- (λ p → proj₂ (Equivalent.from xs≈ys ⟨$⟩ (p , refl)))
-
- from : xs ≈[ set ]′ ys → xs ≈[ set ] ys
- from (p₁≈p₂ , f₁≈f₂ , f₂≈f₁) {z} =
- Eq.equivalent
- (Prod.map (_⟨$⟩_ (Equivalent.to p₁≈p₂))
- (λ {x} eq → begin
- z ≡⟨ eq ⟩
- proj₂ xs x ≡⟨ f₁≈f₂ x ⟩
- proj₂ ys (Equivalent.to p₁≈p₂ ⟨$⟩ x) ∎))
- (Prod.map (_⟨$⟩_ (Equivalent.from p₁≈p₂))
- (λ {x} eq → begin
- z ≡⟨ eq ⟩
- proj₂ ys x ≡⟨ f₂≈f₁ x ⟩
- proj₂ xs (Equivalent.from p₁≈p₂ ⟨$⟩ x) ∎))
-
- ≈⇔≈′-bag : ∀ {c} {C : Container c} {X : Set c} (xs ys : ⟦ C ⟧ X) →
- xs ≈[ bag ] ys ⇔ xs ≈[ bag ]′ ys
- ≈⇔≈′-bag {c} {C} {X} xs ys = Eq.equivalent t f
- where
- open Inverse
-
- t : xs ≈[ bag ] ys → xs ≈[ bag ]′ ys
- t xs≈ys =
- record { to = Equivalent.to (proj₁ xs∼ys)
- ; from = Equivalent.from (proj₁ xs∼ys)
- ; inverse-of = record
- { left-inverse-of = from∘to
- ; right-inverse-of = to∘from
- }
- } ,
- proj₂ xs∼ys
- where
- xs∼ys = Equivalent.to (≈⇔≈′-set xs ys) ⟨$⟩ Inv.⇿⇒ xs≈ys
-
- from∘to : ∀ p → proj₁ (from xs≈ys ⟨$⟩
- (proj₁ (to xs≈ys ⟨$⟩ (p , refl)) ,
- refl)) ≡ p
- from∘to p = begin
- proj₁ (from xs≈ys ⟨$⟩ (proj₁ (to xs≈ys ⟨$⟩ (p , refl)) , refl)) ≡⟨ lemma (to xs≈ys ⟨$⟩ (p , refl)) ⟩
- proj₁ (from xs≈ys ⟨$⟩ (to xs≈ys ⟨$⟩ (p , refl)) ) ≡⟨ P.cong proj₁ $
- left-inverse-of xs≈ys (p , refl) ⟩
- p ∎
- where
- lemma : ∀ {y} (x : ∃ λ p′ → y ≡ proj₂ ys p′) →
- proj₁ (from xs≈ys ⟨$⟩ (proj₁ x , refl)) ≡
- proj₁ (from xs≈ys ⟨$⟩ x )
- lemma (p′ , refl) = refl
-
- to∘from : ∀ p → proj₁ (to xs≈ys ⟨$⟩
- (proj₁ (from xs≈ys ⟨$⟩ (p , refl)) ,
- refl)) ≡ p
- to∘from p = begin
- proj₁ (to xs≈ys ⟨$⟩ (proj₁ (from xs≈ys ⟨$⟩ (p , refl)) , refl)) ≡⟨ lemma (from xs≈ys ⟨$⟩ (p , refl)) ⟩
- proj₁ (to xs≈ys ⟨$⟩ (from xs≈ys ⟨$⟩ (p , refl)) ) ≡⟨ P.cong proj₁ $
- right-inverse-of xs≈ys (p , refl) ⟩
- p ∎
- where
- lemma : ∀ {y} (x : ∃ λ p′ → y ≡ proj₂ xs p′) →
- proj₁ (to xs≈ys ⟨$⟩ (proj₁ x , refl)) ≡
- proj₁ (to xs≈ys ⟨$⟩ x )
- lemma (p′ , refl) = refl
-
- f : xs ≈[ bag ]′ ys → xs ≈[ bag ] ys
- f (p₁≈p₂ , f₁≈f₂ , f₂≈f₁) {z} = record
- { to = Equivalent.to xs∼ys
- ; from = Equivalent.from xs∼ys
- ; inverse-of = record
- { left-inverse-of = λ x →
- let eq = P.trans (P.trans (proj₂ x) (P.trans (f₁≈f₂ (proj₁ x)) refl))
- (P.trans (f₂≈f₁ (to p₁≈p₂ ⟨$⟩ proj₁ x)) refl) in
- H.≅-to-≡ $
- H.cong₂ {B = λ x → z ≡ proj₂ xs x}
- _,_ (H.≡-to-≅ $ left-inverse-of p₁≈p₂ (proj₁ x))
- (proof-irrelevance eq (proj₂ x))
- ; right-inverse-of = λ x →
- let eq = P.trans (P.trans (proj₂ x) (P.trans (f₂≈f₁ (proj₁ x)) refl))
- (P.trans (f₁≈f₂ (from p₁≈p₂ ⟨$⟩ proj₁ x)) refl) in
- H.≅-to-≡ $
- H.cong₂ {B = λ x → z ≡ proj₂ ys x}
- _,_ (H.≡-to-≅ $ right-inverse-of p₁≈p₂ (proj₁ x))
- (proof-irrelevance eq (proj₂ x))
- }
- }
- where
- xs∼ys = Equivalent.from (≈⇔≈′-set xs ys) ⟨$⟩
- (Inv.⇿⇒ p₁≈p₂ , f₁≈f₂ , f₂≈f₁)
-
- proof-irrelevance : {A : Set c} {x y z : A}
- (p : x ≡ y) (q : x ≡ z) → p ≅ q
- proof-irrelevance refl refl = refl
-
-≈⇔≈′ : ∀ {k c} {C : Container c} {X : Set c} (xs ys : ⟦ C ⟧ X) →
- xs ≈[ k ] ys ⇔ xs ≈[ k ]′ ys
-≈⇔≈′ {set} = ≈⇔≈′-set
-≈⇔≈′ {bag} = ≈⇔≈′-bag
diff --git a/src/Data/Container/Any.agda b/src/Data/Container/Any.agda
index 612ee4e..ecde2a9 100644
--- a/src/Data/Container/Any.agda
+++ b/src/Data/Container/Any.agda
@@ -57,7 +57,7 @@ cong : ∀ {k c} {C : Container c}
Isomorphism k (◇ P₁ xs₁) (◇ P₂ xs₂)
cong {C = C} {P₁ = P₁} {P₂} {xs₁} {xs₂} P₁⇿P₂ xs₁≈xs₂ =
◇ P₁ xs₁ ⇿⟨ ⇿∈ C ⟩
- (∃ λ x → x ∈ xs₁ × P₁ x) ≈⟨ Σ.cong (xs₁≈xs₂ ⟨ ×⊎.*-cong ⟩ P₁⇿P₂ _) ⟩
+ (∃ λ x → x ∈ xs₁ × P₁ x) ≈⟨ Σ.cong Inv.id (xs₁≈xs₂ ⟨ ×⊎.*-cong ⟩ P₁⇿P₂ _) ⟩
(∃ λ x → x ∈ xs₂ × P₂ x) ⇿⟨ sym (⇿∈ C) ⟩
◇ P₂ xs₂ ∎
@@ -70,16 +70,16 @@ swap : ∀ {c} {C₁ C₂ : Container c} {X Y : Set c} {P : X → Y → Set c}
◈ xs (◈ ys ∘ P) ⇿ ◈ ys (◈ xs ∘ flip P)
swap {c} {C₁} {C₂} {P = P} {xs} {ys} =
◇ (λ x → ◇ (P x) ys) xs ⇿⟨ ⇿∈ C₁ ⟩
- (∃ λ x → x ∈ xs × ◇ (P x) ys) ⇿⟨ Σ.cong (λ {x} → Inv.id ⟨ ×⊎.*-cong {ℓ = c} ⟩ ⇿∈ C₂ {P = P x}) ⟩
- (∃ λ x → x ∈ xs × ∃ λ y → y ∈ ys × P x y) ⇿⟨ Σ.cong (λ {x} → ∃∃⇿∃∃ {A = x ∈ xs} (λ _ y → y ∈ ys × P x y)) ⟩
+ (∃ λ x → x ∈ xs × ◇ (P x) ys) ⇿⟨ Σ.cong Inv.id (λ {x} → Inv.id ⟨ ×⊎.*-cong {ℓ = c} ⟩ ⇿∈ C₂ {P = P x}) ⟩
+ (∃ λ x → x ∈ xs × ∃ λ y → y ∈ ys × P x y) ⇿⟨ Σ.cong Inv.id (λ {x} → ∃∃⇿∃∃ {A = x ∈ xs} (λ _ y → y ∈ ys × P x y)) ⟩
(∃₂ λ x y → x ∈ xs × y ∈ ys × P x y) ⇿⟨ ∃∃⇿∃∃ (λ x y → x ∈ xs × y ∈ ys × P x y) ⟩
- (∃₂ λ y x → x ∈ xs × y ∈ ys × P x y) ⇿⟨ Σ.cong (λ {y} → Σ.cong (λ {x} →
+ (∃₂ λ y x → x ∈ xs × y ∈ ys × P x y) ⇿⟨ Σ.cong Inv.id (λ {y} → Σ.cong Inv.id (λ {x} →
(x ∈ xs × y ∈ ys × P x y) ⇿⟨ sym $ ×⊎.*-assoc _ _ _ ⟩
((x ∈ xs × y ∈ ys) × P x y) ⇿⟨ ×⊎.*-comm _ _ ⟨ ×⊎.*-cong {ℓ = c} ⟩ Inv.id ⟩
((y ∈ ys × x ∈ xs) × P x y) ⇿⟨ ×⊎.*-assoc _ _ _ ⟩
(y ∈ ys × x ∈ xs × P x y) ∎)) ⟩
- (∃₂ λ y x → y ∈ ys × x ∈ xs × P x y) ⇿⟨ Σ.cong (λ {y} → ∃∃⇿∃∃ {B = y ∈ ys} (λ x _ → x ∈ xs × P x y)) ⟩
- (∃ λ y → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ⇿⟨ Σ.cong (λ {y} → Inv.id ⟨ ×⊎.*-cong {ℓ = c} ⟩ sym (⇿∈ C₁ {P = flip P y})) ⟩
+ (∃₂ λ y x → y ∈ ys × x ∈ xs × P x y) ⇿⟨ Σ.cong Inv.id (λ {y} → ∃∃⇿∃∃ {B = y ∈ ys} (λ x _ → x ∈ xs × P x y)) ⟩
+ (∃ λ y → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ⇿⟨ Σ.cong Inv.id (λ {y} → Inv.id ⟨ ×⊎.*-cong {ℓ = c} ⟩ sym (⇿∈ C₁ {P = flip P y})) ⟩
(∃ λ y → y ∈ ys × ◇ (flip P y) xs) ⇿⟨ sym (⇿∈ C₂) ⟩
◇ (λ y → ◇ (flip P y) xs) ys ∎
diff --git a/src/Data/Context.agda b/src/Data/Context.agda
new file mode 100644
index 0000000..3e924cd
--- /dev/null
+++ b/src/Data/Context.agda
@@ -0,0 +1,188 @@
+------------------------------------------------------------------------
+-- Contexts, variables, substitutions, etc.
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+-- Based on Conor McBride's "Outrageous but Meaningful Coincidences:
+-- Dependent type-safe syntax and evaluation".
+
+-- The contexts and variables are parametrised by a universe.
+
+open import Universe
+
+module Data.Context {u e} (Uni : Universe u e) where
+
+open Universe.Universe Uni
+
+open import Data.Product as Prod
+open import Data.Unit
+open import Function
+open import Level
+
+------------------------------------------------------------------------
+-- Contexts and "types"
+
+mutual
+
+ -- Contexts.
+
+ infixl 5 _▻_
+
+ data Ctxt : Set (u ⊔ e) where
+ ε : Ctxt
+ _▻_ : (Γ : Ctxt) (σ : Type Γ) → Ctxt
+
+ -- Semantic types: maps from environments to universe codes.
+
+ Type : Ctxt → Set (u ⊔ e)
+ Type Γ = Env Γ → U
+
+ -- Interpretation of contexts: environments.
+
+ Env : Ctxt → Set e
+ Env ε = Lift ⊤
+ Env (Γ ▻ σ) = Σ (Env Γ) λ γ → El (σ γ)
+
+-- Type signature for variables, terms, etc.
+
+Term-like : (ℓ : Level) → Set _
+Term-like ℓ = (Γ : Ctxt) → Type Γ → Set ℓ
+
+-- Semantic values: maps from environments to universe values.
+
+Value : Term-like _
+Value Γ σ = (γ : Env Γ) → El (σ γ)
+
+------------------------------------------------------------------------
+-- "Substitutions"
+
+-- Semantic substitutions or context morphisms: maps from environments
+-- to environments.
+
+infixr 4 _⇨_
+
+_⇨_ : Ctxt → Ctxt → Set _
+Γ ⇨ Δ = Env Δ → Env Γ
+
+-- Maps between types.
+
+infixr 4 _⇨₁_
+
+_⇨₁_ : ∀ {Γ} → Type Γ → Type Γ → Set _
+σ ⇨₁ τ = ∀ {γ} → El (τ γ) → El (σ γ)
+
+-- Application of substitutions to types.
+--
+-- Note that this operation is composition of functions. I have chosen
+-- not to give a separate name to the identity substitution, which is
+-- simply the identity function, nor to reverse composition of
+-- substitutions, which is composition of functions.
+
+infixl 8 _/_
+
+_/_ : ∀ {Γ Δ} → Type Γ → Γ ⇨ Δ → Type Δ
+σ / ρ = σ ∘ ρ
+
+-- Application of substitutions to values.
+
+infixl 8 _/v_
+
+_/v_ : ∀ {Γ Δ σ} → Value Γ σ → (ρ : Γ ⇨ Δ) → Value Δ (σ / ρ)
+v /v ρ = v ∘ ρ
+
+-- Weakening.
+
+wk : ∀ {Γ σ} → Γ ⇨ Γ ▻ σ
+wk = proj₁
+
+-- Extends a substitution with another value.
+
+infixl 5 _►_
+
+_►_ : ∀ {Γ Δ σ} (ρ : Γ ⇨ Δ) → Value Δ (σ / ρ) → Γ ▻ σ ⇨ Δ
+_►_ = <_,_>
+
+-- A substitution which only replaces the first "variable".
+
+sub : ∀ {Γ σ} → Value Γ σ → Γ ▻ σ ⇨ Γ
+sub v = id ► v
+
+-- One can construct a substitution between two non-empty contexts by
+-- supplying two functions, one which handles the last element and one
+-- which handles the rest.
+
+infixl 10 _↑_
+
+_↑_ : ∀ {Γ Δ σ τ} (ρ : Γ ⇨ Δ) → σ / ρ ⇨₁ τ → Γ ▻ σ ⇨ Δ ▻ τ
+_↑_ = Prod.map
+
+-- Lifting.
+
+infix 10 _↑
+
+_↑ : ∀ {Γ Δ σ} (ρ : Γ ⇨ Δ) → Γ ▻ σ ⇨ Δ ▻ σ / ρ
+ρ ↑ = ρ ↑ id
+
+------------------------------------------------------------------------
+-- Variables
+
+-- Variables (de Bruijn indices).
+
+infix 4 _∋_
+
+data _∋_ : Term-like (u ⊔ e) where
+ zero : ∀ {Γ σ} → Γ ▻ σ ∋ σ / wk
+ suc : ∀ {Γ σ τ} (x : Γ ∋ σ) → Γ ▻ τ ∋ σ / wk
+
+-- Interpretation of variables: a lookup function.
+
+lookup : ∀ {Γ σ} → Γ ∋ σ → Value Γ σ
+lookup zero (γ , v) = v
+lookup (suc x) (γ , v) = lookup x γ
+
+-- Application of substitutions to variables.
+
+infixl 8 _/∋_
+
+_/∋_ : ∀ {Γ Δ σ} → Γ ∋ σ → (ρ : Γ ⇨ Δ) → Value Δ (σ / ρ)
+x /∋ ρ = lookup x /v ρ
+
+------------------------------------------------------------------------
+-- Context extensions
+
+mutual
+
+ -- Context extensions.
+
+ infixl 5 _▻_
+
+ data Ctxt⁺ (Γ : Ctxt) : Set (u ⊔ e) where
+ ε : Ctxt⁺ Γ
+ _▻_ : (Γ⁺ : Ctxt⁺ Γ) (σ : Type (Γ ++ Γ⁺)) → Ctxt⁺ Γ
+
+ -- Appends a context extension to a context.
+
+ infixl 5 _++_
+
+ _++_ : (Γ : Ctxt) → Ctxt⁺ Γ → Ctxt
+ Γ ++ ε = Γ
+ Γ ++ (Γ⁺ ▻ σ) = (Γ ++ Γ⁺) ▻ σ
+
+mutual
+
+ -- Application of substitutions to context extensions.
+
+ infixl 8 _/⁺_
+
+ _/⁺_ : ∀ {Γ Δ} → Ctxt⁺ Γ → Γ ⇨ Δ → Ctxt⁺ Δ
+ ε /⁺ ρ = ε
+ (Γ⁺ ▻ σ) /⁺ ρ = Γ⁺ /⁺ ρ ▻ σ / ρ ↑⋆ Γ⁺
+
+ -- N-ary lifting of substitutions.
+
+ infixl 10 _↑⋆_
+
+ _↑⋆_ : ∀ {Γ Δ} (ρ : Γ ⇨ Δ) → ∀ Γ⁺ → Γ ++ Γ⁺ ⇨ Δ ++ Γ⁺ /⁺ ρ
+ ρ ↑⋆ ε = ρ
+ ρ ↑⋆ (Γ⁺ ▻ σ) = (ρ ↑⋆ Γ⁺) ↑
diff --git a/src/Data/Fin.agda b/src/Data/Fin.agda
index 9fba6ea..fbdd2f8 100644
--- a/src/Data/Fin.agda
+++ b/src/Data/Fin.agda
@@ -11,7 +11,7 @@ module Data.Fin where
open import Data.Nat as Nat
using (ℕ; zero; suc; z≤n; s≤s)
renaming ( _+_ to _N+_; _∸_ to _N∸_
- ; _≤_ to _N≤_; _<_ to _N<_; _≤?_ to _N≤?_)
+ ; _≤_ to _N≤_; _≥_ to _N≥_; _<_ to _N<_; _≤?_ to _N≤?_)
open import Function
open import Level hiding (lift)
open import Relation.Nullary.Decidable
@@ -65,6 +65,13 @@ raise : ∀ {m} n → Fin m → Fin (n N+ m)
raise zero i = i
raise (suc n) i = suc (raise n i)
+-- reduce≥ "m + n" _ = "n".
+
+reduce≥ : ∀ {m n} (i : Fin (m N+ n)) (i≥m : toℕ i N≥ m) → Fin n
+reduce≥ {zero} i i≥m = i
+reduce≥ {suc m} zero ()
+reduce≥ {suc m} (suc i) (s≤s i≥m) = reduce≥ i i≥m
+
-- inject⋆ m "n" = "n".
inject : ∀ {n} {i : Fin n} → Fin′ i → Fin n
diff --git a/src/Data/Fin/Dec.agda b/src/Data/Fin/Dec.agda
index 7b50302..d54ca0c 100644
--- a/src/Data/Fin/Dec.agda
+++ b/src/Data/Fin/Dec.agda
@@ -7,15 +7,23 @@
module Data.Fin.Dec where
open import Function
+import Data.Bool as Bool
open import Data.Nat hiding (_<_)
open import Data.Vec hiding (_∈_)
+open import Data.Vec.Equality as VecEq
+ using () renaming (module HeterogeneousEquality to HetVecEq)
open import Data.Fin
open import Data.Fin.Subset
open import Data.Fin.Subset.Props
open import Data.Product as Prod
open import Data.Empty
+open import Function
+import Function.Equivalence as Eq
open import Level hiding (Lift)
+open import Relation.Binary
+import Relation.Binary.HeterogeneousEquality as H
open import Relation.Nullary
+import Relation.Nullary.Decidable as Dec
open import Relation.Unary using (Pred)
infix 4 _∈?_
@@ -92,8 +100,8 @@ all∈? {q = q} dec = decFinSubset (λ f → f ∈? q) dec
all? : ∀ {n} {P : Fin n → Set} →
(∀ f → Dec (P f)) →
Dec (∀ f → P f)
-all? dec with all∈? {q = all inside} (λ {f} _ → dec f)
-... | yes ∀p = yes (λ f → ∀p (allInside f))
+all? dec with all∈? {q = ⊤} (λ {f} _ → dec f)
+... | yes ∀p = yes (λ f → ∀p ∈⊤)
... | no ¬∀p = no (λ ∀p → ¬∀p (λ {f} _ → ∀p f))
decLift : ∀ {n} {P : Fin n → Set} →
@@ -153,3 +161,15 @@ anySubset? {suc n} {P} dec with anySubset? (restrictS inside dec)
((j : Fin′ (suc i)) → P (inject j))
extend′ g zero = P0
extend′ g (suc j) = g j
+
+-- Decision procedure for _⊆_ (obtained via the natural lattice
+-- order).
+
+infix 4 _⊆?_
+
+_⊆?_ : ∀ {n} → Decidable (_⊆_ {n = n})
+p₁ ⊆? p₂ =
+ Dec.map (Eq.sym NaturalPoset.orders-equivalent) $
+ Dec.map′ H.≅-to-≡ H.≡-to-≅ $
+ Dec.map′ HetVecEq.to-≅ HetVecEq.from-≅ $
+ VecEq.DecidableEquality._≟_ Bool.decSetoid p₁ (p₁ ∩ p₂)
diff --git a/src/Data/Fin/Props.agda b/src/Data/Fin/Props.agda
index d52926e..7422d90 100644
--- a/src/Data/Fin/Props.agda
+++ b/src/Data/Fin/Props.agda
@@ -95,9 +95,9 @@ inject-lemma {i = zero} ()
inject-lemma {i = suc i} zero = refl
inject-lemma {i = suc i} (suc j) = cong suc (inject-lemma j)
-inject+-lemma : ∀ m k → m ≡ toℕ (inject+ k (fromℕ m))
-inject+-lemma zero k = refl
-inject+-lemma (suc m) k = cong suc (inject+-lemma m k)
+inject+-lemma : ∀ {m} n (i : Fin m) → toℕ i ≡ toℕ (inject+ n i)
+inject+-lemma n zero = refl
+inject+-lemma n (suc i) = cong suc (inject+-lemma n i)
inject₁-lemma : ∀ {m} (i : Fin m) → toℕ (inject₁ i) ≡ toℕ i
inject₁-lemma zero = refl
@@ -118,6 +118,18 @@ inject≤-lemma (suc i) (N.s≤s le) = cong suc (inject≤-lemma i le)
<′⇒≺ (N.≤′-step m≤′n) | n ≻toℕ i =
subst (λ i → i ≺ suc n) (inject₁-lemma i) (suc n ≻toℕ (inject₁ i))
+toℕ-raise : ∀ {m} n (i : Fin m) → toℕ (raise n i) ≡ n ℕ+ toℕ i
+toℕ-raise zero i = refl
+toℕ-raise (suc n) i = cong suc (toℕ-raise n i)
+
+fromℕ≤-toℕ : ∀ {m} (i : Fin m) (i<m : toℕ i ℕ< m) → fromℕ≤ i<m ≡ i
+fromℕ≤-toℕ zero (s≤s z≤n) = refl
+fromℕ≤-toℕ (suc i) (s≤s (s≤s m≤n)) = cong suc (fromℕ≤-toℕ i (s≤s m≤n))
+
+toℕ-fromℕ≤ : ∀ {m n} (m<n : m ℕ< n) → toℕ (fromℕ≤ m<n) ≡ m
+toℕ-fromℕ≤ (s≤s z≤n) = refl
+toℕ-fromℕ≤ (s≤s (s≤s m<n)) = cong suc (toℕ-fromℕ≤ (s≤s m<n))
+
------------------------------------------------------------------------
-- Operations
diff --git a/src/Data/Fin/Subset.agda b/src/Data/Fin/Subset.agda
index 3cb707e..484b275 100644
--- a/src/Data/Fin/Subset.agda
+++ b/src/Data/Fin/Subset.agda
@@ -4,21 +4,27 @@
module Data.Fin.Subset where
-open import Data.Nat
-open import Data.Vec hiding (_∈_)
+open import Algebra
+import Algebra.Props.BooleanAlgebra as BoolAlgProp
+import Algebra.Props.BooleanAlgebra.Expression as BAExpr
+import Data.Bool.Properties as BoolProp
open import Data.Fin
+open import Data.List as List using (List)
+open import Data.Nat
open import Data.Product
+open import Data.Vec using (Vec; _∷_; _[_]=_)
+import Relation.Binary.Vec.Pointwise as Pointwise
open import Relation.Nullary
-open import Relation.Binary.PropositionalEquality
-infixr 2 _∈_ _∉_ _⊆_ _⊈_
+infix 4 _∈_ _∉_ _⊆_ _⊈_
------------------------------------------------------------------------
-- Definitions
-data Side : Set where
- inside : Side
- outside : Side
+-- Sides.
+
+open import Data.Bool public
+ using () renaming (Bool to Side; true to inside; false to outside)
-- Partitions a finite set into two parts, the inside and the outside.
@@ -41,11 +47,47 @@ _⊈_ : ∀ {n} → Subset n → Subset n → Set
p₁ ⊈ p₂ = ¬ (p₁ ⊆ p₂)
------------------------------------------------------------------------
--- Some specific subsets
-
-all : ∀ {n} → Side → Subset n
-all {zero} _ = []
-all {suc n} s = s ∷ all s
+-- Set operations
+
+-- Pointwise lifting of the usual boolean algebra for booleans gives
+-- us a boolean algebra for subsets.
+--
+-- The underlying equality of the returned boolean algebra is
+-- propositional equality.
+
+booleanAlgebra : ℕ → BooleanAlgebra _ _
+booleanAlgebra n =
+ BoolAlgProp.replace-equality
+ (BAExpr.lift BoolProp.booleanAlgebra n)
+ Pointwise.Pointwise-≡
+
+private
+ open module BA {n} = BooleanAlgebra (booleanAlgebra n) public
+ using
+ ( ⊥ -- The empty subset.
+ ; ⊤ -- The subset containing all elements.
+ )
+ renaming
+ ( _∨_ to _∪_ -- Binary union.
+ ; _∧_ to _∩_ -- Binary intersection.
+ ; ¬_ to ∁ -- Complement.
+ )
+
+-- A singleton subset, containing just the given element.
+
+⁅_⁆ : ∀ {n} → Fin n → Subset n
+⁅ zero ⁆ = inside ∷ ⊥
+⁅ suc i ⁆ = outside ∷ ⁅ i ⁆
+
+-- N-ary union.
+
+⋃ : ∀ {n} → List (Subset n) → Subset n
+⋃ = List.foldr _∪_ ⊥
+
+-- N-ary intersection.
+
+⋂ : ∀ {n} → List (Subset n) → Subset n
+⋂ = List.foldr _∩_ ⊤
------------------------------------------------------------------------
-- Properties
diff --git a/src/Data/Fin/Subset/Props.agda b/src/Data/Fin/Subset/Props.agda
index 7f8022a..f6dc52b 100644
--- a/src/Data/Fin/Subset/Props.agda
+++ b/src/Data/Fin/Subset/Props.agda
@@ -4,14 +4,21 @@
module Data.Fin.Subset.Props where
-open import Data.Nat
-open import Data.Vec hiding (_∈_)
-open import Data.Empty
-open import Function
-open import Data.Fin
+open import Algebra
+import Algebra.Props.BooleanAlgebra as BoolProp
+open import Data.Empty using (⊥-elim)
+open import Data.Fin using (Fin); open Data.Fin.Fin
open import Data.Fin.Subset
+open import Data.Nat using (ℕ)
open import Data.Product
-open import Relation.Binary.PropositionalEquality
+open import Data.Sum as Sum
+open import Data.Vec hiding (_∈_)
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence
+ using (_⇔_; equivalent; module Equivalent)
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
------------------------------------------------------------------------
-- Constructor mangling
@@ -23,39 +30,125 @@ drop-∷-⊆ : ∀ {n s₁ s₂} {p₁ p₂ : Subset n} → s₁ ∷ p₁ ⊆ s
drop-∷-⊆ p₁s₁⊆p₂s₂ x∈p₁ = drop-there $ p₁s₁⊆p₂s₂ (there x∈p₁)
drop-∷-Empty : ∀ {n s} {p : Subset n} → Empty (s ∷ p) → Empty p
-drop-∷-Empty ¬¬∅ (x , x∈p) = ¬¬∅ (suc x , there x∈p)
+drop-∷-Empty ¬∃∈ (x , x∈p) = ¬∃∈ (suc x , there x∈p)
+
+------------------------------------------------------------------------
+-- Properties involving ⊥
+
+∉⊥ : ∀ {n} {x : Fin n} → x ∉ ⊥
+∉⊥ (there p) = ∉⊥ p
+
+⊥⊆ : ∀ {n} {p : Subset n} → ⊥ ⊆ p
+⊥⊆ x∈⊥ with ∉⊥ x∈⊥
+... | ()
+
+Empty-unique : ∀ {n} {p : Subset n} →
+ Empty p → p ≡ ⊥
+Empty-unique {p = []} ¬∃∈ = P.refl
+Empty-unique {p = s ∷ p} ¬∃∈ with Empty-unique (drop-∷-Empty ¬∃∈)
+Empty-unique {p = outside ∷ .⊥} ¬∃∈ | P.refl = P.refl
+Empty-unique {p = inside ∷ .⊥} ¬∃∈ | P.refl =
+ ⊥-elim (¬∃∈ (zero , here))
+
+------------------------------------------------------------------------
+-- Properties involving ⊤
+
+∈⊤ : ∀ {n} {x : Fin n} → x ∈ ⊤
+∈⊤ {x = zero} = here
+∈⊤ {x = suc x} = there ∈⊤
+
+⊆⊤ : ∀ {n} {p : Subset n} → p ⊆ ⊤
+⊆⊤ = const ∈⊤
+
+------------------------------------------------------------------------
+-- A property involving ⁅_⁆
+
+x∈⁅y⁆⇔x≡y : ∀ {n} {x y : Fin n} → x ∈ ⁅ y ⁆ ⇔ x ≡ y
+x∈⁅y⁆⇔x≡y {x = x} {y} =
+ equivalent (to y) (λ x≡y → P.subst (λ y → x ∈ ⁅ y ⁆) x≡y (x∈⁅x⁆ x))
+ where
+
+ to : ∀ {n x} (y : Fin n) → x ∈ ⁅ y ⁆ → x ≡ y
+ to (suc y) (there p) = P.cong suc (to y p)
+ to zero here = P.refl
+ to zero (there p) with ∉⊥ p
+ ... | ()
+
+ x∈⁅x⁆ : ∀ {n} (x : Fin n) → x ∈ ⁅ x ⁆
+ x∈⁅x⁆ zero = here
+ x∈⁅x⁆ (suc x) = there (x∈⁅x⁆ x)
------------------------------------------------------------------------
--- More interesting properties
+-- A property involving _∪_
+
+∪⇿⊎ : ∀ {n} {p₁ p₂ : Subset n} {x} → x ∈ p₁ ∪ p₂ ⇔ (x ∈ p₁ ⊎ x ∈ p₂)
+∪⇿⊎ = equivalent (to _ _) from
+ where
+ to : ∀ {n} (p₁ p₂ : Subset n) {x} → x ∈ p₁ ∪ p₂ → x ∈ p₁ ⊎ x ∈ p₂
+ to [] [] ()
+ to (inside ∷ p₁) (s₂ ∷ p₂) here = inj₁ here
+ to (outside ∷ p₁) (inside ∷ p₂) here = inj₂ here
+ to (s₁ ∷ p₁) (s₂ ∷ p₂) (there x∈p₁∪p₂) =
+ Sum.map there there (to p₁ p₂ x∈p₁∪p₂)
+
+ ⊆∪ˡ : ∀ {n p₁} (p₂ : Subset n) → p₁ ⊆ p₁ ∪ p₂
+ ⊆∪ˡ [] ()
+ ⊆∪ˡ (s ∷ p₂) here = here
+ ⊆∪ˡ (s ∷ p₂) (there x∈p₁) = there (⊆∪ˡ p₂ x∈p₁)
+
+ ⊆∪ʳ : ∀ {n} (p₁ p₂ : Subset n) → p₂ ⊆ p₁ ∪ p₂
+ ⊆∪ʳ p₁ p₂ rewrite BooleanAlgebra.∨-comm (booleanAlgebra _) p₁ p₂
+ = ⊆∪ˡ p₁
+
+ from : ∀ {n} {p₁ p₂ : Subset n} {x} → x ∈ p₁ ⊎ x ∈ p₂ → x ∈ p₁ ∪ p₂
+ from (inj₁ x∈p₁) = ⊆∪ˡ _ x∈p₁
+ from (inj₂ x∈p₂) = ⊆∪ʳ _ _ x∈p₂
+
+------------------------------------------------------------------------
+-- _⊆_ is a partial order
+
+-- The "natural poset" associated with the boolean algebra.
+
+module NaturalPoset where
+ private
+ open module BA {n} = BoolProp (booleanAlgebra n) public
+ using (poset)
+ open module Po {n} = Poset (poset {n = n}) public
+
+ -- _⊆_ is equivalent to the natural lattice order.
+
+ orders-equivalent : ∀ {n} {p₁ p₂ : Subset n} → p₁ ⊆ p₂ ⇔ p₁ ≤ p₂
+ orders-equivalent = equivalent (to _ _) (from _ _)
+ where
+ to : ∀ {n} (p₁ p₂ : Subset n) → p₁ ⊆ p₂ → p₁ ≤ p₂
+ to [] [] p₁⊆p₂ = P.refl
+ to (inside ∷ p₁) (_ ∷ p₂) p₁⊆p₂ with p₁⊆p₂ here
+ to (inside ∷ p₁) (.inside ∷ p₂) p₁⊆p₂ | here = P.cong (_∷_ inside) (to p₁ p₂ (drop-∷-⊆ p₁⊆p₂))
+ to (outside ∷ p₁) (_ ∷ p₂) p₁⊆p₂ = P.cong (_∷_ outside) (to p₁ p₂ (drop-∷-⊆ p₁⊆p₂))
-allInside : ∀ {n} (x : Fin n) → x ∈ all inside
-allInside zero = here
-allInside (suc x) = there (allInside x)
+ from : ∀ {n} (p₁ p₂ : Subset n) → p₁ ≤ p₂ → p₁ ⊆ p₂
+ from [] [] p₁≤p₂ x = x
+ from (.inside ∷ _) (_ ∷ _) p₁≤p₂ here rewrite P.cong head p₁≤p₂ = here
+ from (_ ∷ p₁) (_ ∷ p₂) p₁≤p₂ (there xs[i]=x) =
+ there (from p₁ p₂ (P.cong tail p₁≤p₂) xs[i]=x)
-allOutside : ∀ {n} (x : Fin n) → x ∉ all outside
-allOutside zero ()
-allOutside (suc x) (there x∈) = allOutside x x∈
+-- _⊆_ is a partial order.
-⊆⊇⟶≡ : ∀ {n} {p₁ p₂ : Subset n} →
- p₁ ⊆ p₂ → p₂ ⊆ p₁ → p₁ ≡ p₂
-⊆⊇⟶≡ = helper _ _
+poset : ℕ → Poset _ _ _
+poset n = record
+ { Carrier = Subset n
+ ; _≈_ = _≡_
+ ; _≤_ = _⊆_
+ ; isPartialOrder = record
+ { isPreorder = record
+ { isEquivalence = P.isEquivalence
+ ; reflexive = λ i≡j → from ⟨$⟩ reflexive i≡j
+ ; trans = λ x⊆y y⊆z → from ⟨$⟩ trans (to ⟨$⟩ x⊆y) (to ⟨$⟩ y⊆z)
+ }
+ ; antisym = λ x⊆y y⊆x → antisym (to ⟨$⟩ x⊆y) (to ⟨$⟩ y⊆x)
+ }
+ }
where
- helper : ∀ {n} (p₁ p₂ : Subset n) →
- p₁ ⊆ p₂ → p₂ ⊆ p₁ → p₁ ≡ p₂
- helper [] [] _ _ = refl
- helper (s₁ ∷ p₁) (s₂ ∷ p₂) ₁⊆₂ ₂⊆₁ with ⊆⊇⟶≡ (drop-∷-⊆ ₁⊆₂)
- (drop-∷-⊆ ₂⊆₁)
- helper (outside ∷ p) (outside ∷ .p) ₁⊆₂ ₂⊆₁ | refl = refl
- helper (inside ∷ p) (inside ∷ .p) ₁⊆₂ ₂⊆₁ | refl = refl
- helper (outside ∷ p) (inside ∷ .p) ₁⊆₂ ₂⊆₁ | refl with ₂⊆₁ here
- ... | ()
- helper (inside ∷ p) (outside ∷ .p) ₁⊆₂ ₂⊆₁ | refl with ₁⊆₂ here
- ... | ()
-
-∅⟶allOutside : ∀ {n} {p : Subset n} →
- Empty p → p ≡ all outside
-∅⟶allOutside {p = []} ¬¬∅ = refl
-∅⟶allOutside {p = s ∷ ps} ¬¬∅ with ∅⟶allOutside (drop-∷-Empty ¬¬∅)
-∅⟶allOutside {p = outside ∷ .(all outside)} ¬¬∅ | refl = refl
-∅⟶allOutside {p = inside ∷ .(all outside)} ¬¬∅ | refl =
- ⊥-elim (¬¬∅ (zero , here))
+ open NaturalPoset
+ open module E {p₁ p₂} =
+ Equivalent (orders-equivalent {n = n} {p₁ = p₁} {p₂ = p₂})
diff --git a/src/Data/Fin/Substitution/Lemmas.agda b/src/Data/Fin/Substitution/Lemmas.agda
index 67ee5d1..3f5baa0 100644
--- a/src/Data/Fin/Substitution/Lemmas.agda
+++ b/src/Data/Fin/Substitution/Lemmas.agda
@@ -4,12 +4,13 @@
module Data.Fin.Substitution.Lemmas where
+import Category.Applicative.Indexed as Applicative
open import Data.Fin.Substitution
open import Data.Nat
open import Data.Fin using (Fin; zero; suc; lift)
open import Data.Vec
import Data.Vec.Properties as VecProp
-open import Function as Fun using (_∘_)
+open import Function as Fun using (_∘_; _$_)
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; refl; sym; cong; cong₂)
open PropEq.≡-Reasoning
@@ -52,9 +53,10 @@ record Lemmas₀ (T : ℕ → Set) : Set where
lookup-map-weaken-↑⋆ zero x = refl
lookup-map-weaken-↑⋆ (suc k) zero = refl
lookup-map-weaken-↑⋆ (suc k) (suc x) {ρ} = begin
- lookup x (map weaken (map weaken ρ ↑⋆ k)) ≡⟨ VecProp.lookup-natural weaken x _ ⟩
+ lookup x (map weaken (map weaken ρ ↑⋆ k)) ≡⟨ Applicative.Morphism.op-<$> (VecProp.lookup-morphism x) weaken _ ⟩
weaken (lookup x (map weaken ρ ↑⋆ k)) ≡⟨ cong weaken (lookup-map-weaken-↑⋆ k x) ⟩
- weaken (lookup (lift k suc x) ((ρ ↑) ↑⋆ k)) ≡⟨ sym (VecProp.lookup-natural weaken (lift k suc x) _) ⟩
+ weaken (lookup (lift k suc x) ((ρ ↑) ↑⋆ k)) ≡⟨ sym $
+ Applicative.Morphism.op-<$> (VecProp.lookup-morphism (lift k suc x)) weaken _ ⟩
lookup (lift k suc x) (map weaken ((ρ ↑) ↑⋆ k)) ∎
record Lemmas₁ (T : ℕ → Set) : Set where
@@ -69,7 +71,7 @@ record Lemmas₁ (T : ℕ → Set) : Set where
lookup x ρ ≡ var y →
lookup x (map weaken ρ) ≡ var (suc y)
lookup-map-weaken x {y} {ρ} hyp = begin
- lookup x (map weaken ρ) ≡⟨ VecProp.lookup-natural weaken x ρ ⟩
+ lookup x (map weaken ρ) ≡⟨ Applicative.Morphism.op-<$> (VecProp.lookup-morphism x) weaken ρ ⟩
weaken (lookup x ρ) ≡⟨ cong weaken hyp ⟩
weaken (var y) ≡⟨ weaken-var ⟩
var (suc y) ∎
@@ -137,7 +139,7 @@ record Lemmas₂ (T : ℕ → Set) : Set where
lookup-⊙ : ∀ {m n k} x {ρ₁ : Sub T m n} {ρ₂ : Sub T n k} →
lookup x (ρ₁ ⊙ ρ₂) ≡ lookup x ρ₁ / ρ₂
- lookup-⊙ x = VecProp.lookup-natural _ x _
+ lookup-⊙ x = Applicative.Morphism.op-<$> (VecProp.lookup-morphism x) _ _
lookup-⨀ : ∀ {m n} x (ρs : Subs T m n) →
lookup x (⨀ ρs) ≡ var x /✶ ρs
diff --git a/src/Data/List.agda b/src/Data/List.agda
index 8f4c7ce..c76b482 100644
--- a/src/Data/List.agda
+++ b/src/Data/List.agda
@@ -57,13 +57,6 @@ map : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → List A → List B
map f [] = []
map f (x ∷ xs) = f x ∷ map f xs
-reverse : ∀ {a} {A : Set a} → List A → List A
-reverse xs = rev xs []
- where
- rev : ∀ {a} {A : Set a} → List A → List A → List A
- rev [] ys = ys
- rev (x ∷ xs) ys = rev xs (x ∷ ys)
-
replicate : ∀ {a} {A : Set a} → (n : ℕ) → A → List A
replicate zero x = []
replicate (suc n) x = x ∷ replicate n x
@@ -121,6 +114,9 @@ product = foldr _*_ 1
length : ∀ {a} {A : Set a} → List A → ℕ
length = foldr (λ _ → suc) 0
+reverse : ∀ {a} {A : Set a} → List A → List A
+reverse = foldl (λ rev x → x ∷ rev) []
+
-- * Building lists
-- ** Scans
@@ -284,19 +280,19 @@ monoid A = record
open import Category.Monad
-monad : RawMonad List
+monad : ∀ {ℓ} → RawMonad (List {ℓ})
monad = record
{ return = λ x → x ∷ []
; _>>=_ = λ xs f → concat (map f xs)
}
-monadZero : RawMonadZero List
+monadZero : ∀ {ℓ} → RawMonadZero (List {ℓ})
monadZero = record
{ monad = monad
; ∅ = []
}
-monadPlus : RawMonadPlus List
+monadPlus : ∀ {ℓ} → RawMonadPlus (List {ℓ})
monadPlus = record
{ monadZero = monadZero
; _∣_ = _++_
@@ -306,11 +302,11 @@ monadPlus = record
-- Monadic functions
private
- module Monadic {M} (Mon : RawMonad M) where
+ module Monadic {m} {M : Set m → Set m} (Mon : RawMonad M) where
open RawMonad Mon
- sequence : ∀ {A : Set} → List (M A) → M (List A)
+ sequence : ∀ {A} → List (M A) → M (List A)
sequence [] = return []
sequence (x ∷ xs) = _∷_ <$> x ⊛ sequence xs
diff --git a/src/Data/List/All.agda b/src/Data/List/All.agda
index 5caedfd..a48c608 100644
--- a/src/Data/List/All.agda
+++ b/src/Data/List/All.agda
@@ -2,12 +2,15 @@
-- Lists where all elements satisfy a given property
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Data.List.All where
-open import Function
open import Data.List as List hiding (map; all)
open import Data.List.Any as Any using (here; there)
open Any.Membership-≡ using (_∈_; _⊆_)
+open import Function
+open import Level
open import Relation.Nullary
import Relation.Nullary.Decidable as Dec
open import Relation.Unary using () renaming (_⊆_ to _⋐_)
@@ -17,30 +20,36 @@ open import Relation.Binary.PropositionalEquality
infixr 5 _∷_
-data All {A} (P : A → Set) : List A → Set where
+data All {a p} {A : Set a}
+ (P : A → Set p) : List A → Set (p ⊔ a) where
[] : All P []
_∷_ : ∀ {x xs} (px : P x) (pxs : All P xs) → All P (x ∷ xs)
-head : ∀ {A} {P : A → Set} {x xs} → All P (x ∷ xs) → P x
+head : ∀ {a p} {A : Set a} {P : A → Set p} {x xs} →
+ All P (x ∷ xs) → P x
head (px ∷ pxs) = px
-tail : ∀ {A} {P : A → Set} {x xs} → All P (x ∷ xs) → All P xs
+tail : ∀ {a p} {A : Set a} {P : A → Set p} {x xs} →
+ All P (x ∷ xs) → All P xs
tail (px ∷ pxs) = pxs
-lookup : ∀ {A} {P : A → Set} {xs} → All P xs → (∀ {x} → x ∈ xs → P x)
+lookup : ∀ {a p} {A : Set a} {P : A → Set p} {xs : List A} →
+ All P xs → (∀ {x : A} → x ∈ xs → P x)
lookup [] ()
lookup (px ∷ pxs) (here refl) = px
lookup (px ∷ pxs) (there x∈xs) = lookup pxs x∈xs
-tabulate : ∀ {A} {P : A → Set} {xs} → (∀ {x} → x ∈ xs → P x) → All P xs
+tabulate : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
+ (∀ {x} → x ∈ xs → P x) → All P xs
tabulate {xs = []} hyp = []
tabulate {xs = x ∷ xs} hyp = hyp (here refl) ∷ tabulate (hyp ∘ there)
-map : ∀ {A} {P Q : A → Set} → P ⋐ Q → All P ⋐ All Q
+map : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} →
+ P ⋐ Q → All P ⋐ All Q
map g [] = []
map g (px ∷ pxs) = g px ∷ map g pxs
-all : ∀ {A} {P : A → Set} →
+all : ∀ {a p} {A : Set a} {P : A → Set p} →
(∀ x → Dec (P x)) → (xs : List A) → Dec (All P xs)
all p [] = yes []
all p (x ∷ xs) with p x
diff --git a/src/Data/List/All/Properties.agda b/src/Data/List/All/Properties.agda
index accf2cb..2ed7a24 100644
--- a/src/Data/List/All/Properties.agda
+++ b/src/Data/List/All/Properties.agda
@@ -2,6 +2,8 @@
-- Properties relating All to various list functions
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Data.List.All.Properties where
open import Data.Bool
@@ -18,30 +20,34 @@ open import Relation.Unary using () renaming (_⊆_ to _⋐_)
-- Functions can be shifted between the predicate and the list.
-All-map : ∀ {A B} {P : B → Set} {f : A → B} {xs} →
+All-map : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
+ {f : A → B} {xs} →
All (P ∘ f) xs → All P (List.map f xs)
All-map [] = []
All-map (p ∷ ps) = p ∷ All-map ps
-map-All : ∀ {A B} {P : B → Set} {f : A → B} {xs} →
+map-All : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
+ {f : A → B} {xs} →
All P (List.map f xs) → All (P ∘ f) xs
map-All {xs = []} [] = []
map-All {xs = _ ∷ _} (p ∷ ps) = p ∷ map-All ps
-- A variant of All.map.
-gmap : ∀ {A B} {P : A → Set} {Q : B → Set} {f : A → B} →
+gmap : ∀ {a b p q}
+ {A : Set a} {B : Set b} {P : A → Set p} {Q : B → Set q}
+ {f : A → B} →
P ⋐ Q ∘ f → All P ⋐ All Q ∘ List.map f
gmap g = All-map ∘ All.map g
-- All and all are related via T.
-All-all : ∀ {A} (p : A → Bool) {xs} →
+All-all : ∀ {a} {A : Set a} (p : A → Bool) {xs} →
All (T ∘ p) xs → T (all p xs)
All-all p [] = _
All-all p (px ∷ pxs) = Equivalent.from T-∧ ⟨$⟩ (px , All-all p pxs)
-all-All : ∀ {A} (p : A → Bool) xs →
+all-All : ∀ {a} {A : Set a} (p : A → Bool) xs →
T (all p xs) → All (T ∘ p) xs
all-All p [] _ = []
all-All p (x ∷ xs) px∷xs with Equivalent.to (T-∧ {p x}) ⟨$⟩ px∷xs
@@ -49,12 +55,12 @@ all-All p (x ∷ xs) px∷xs | (px , pxs) = px ∷ all-All p xs pxs
-- All is anti-monotone.
-anti-mono : ∀ {A} {P : A → Set} {xs ys} →
+anti-mono : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
xs ⊆ ys → All P ys → All P xs
anti-mono xs⊆ys pys = All.tabulate (All.lookup pys ∘ xs⊆ys)
-- all is anti-monotone.
-all-anti-mono : ∀ {A} (p : A → Bool) {xs ys} →
+all-anti-mono : ∀ {a} {A : Set a} (p : A → Bool) {xs ys} →
xs ⊆ ys → T (all p ys) → T (all p xs)
all-anti-mono p xs⊆ys = All-all p ∘ anti-mono xs⊆ys ∘ all-All p _
diff --git a/src/Data/List/Any.agda b/src/Data/List/Any.agda
index 60f3237..da577fb 100644
--- a/src/Data/List/Any.agda
+++ b/src/Data/List/Any.agda
@@ -2,6 +2,8 @@
-- Lists where at least one element satisfies a given property
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Data.List.Any where
open import Data.Empty
@@ -25,25 +27,28 @@ open import Relation.Binary.PropositionalEquality as PropEq
-- Any P xs means that at least one element in xs satisfies P.
-data Any {A} (P : A → Set) : List A → Set where
+data Any {a p} {A : Set a}
+ (P : A → Set p) : List A → Set (a ⊔ p) where
here : ∀ {x xs} (px : P x) → Any P (x ∷ xs)
there : ∀ {x xs} (pxs : Any P xs) → Any P (x ∷ xs)
-- Map.
-map : ∀ {A} {P Q : A → Set} → P ⋐ Q → Any P ⋐ Any Q
+map : ∀ {a p q} {A : Set a} {P : A → Set p} → {Q : A → Set q} →
+ P ⋐ Q → Any P ⋐ Any Q
map g (here px) = here (g px)
map g (there pxs) = there (map g pxs)
-- If the head does not satisfy the predicate, then the tail will.
-tail : ∀ {A x xs} {P : A → Set} → ¬ P x → Any P (x ∷ xs) → Any P xs
+tail : ∀ {a p} {A : Set a} {x : A} {xs : List A} {P : A → Set p} →
+ ¬ P x → Any P (x ∷ xs) → Any P xs
tail ¬px (here px) = ⊥-elim (¬px px)
tail ¬px (there pxs) = pxs
-- Decides Any.
-any : {A : Set} {P : A → Set} →
+any : ∀ {a p} {A : Set a} {P : A → Set p} →
(∀ x → Dec (P x)) → (xs : List A) → Dec (Any P xs)
any p [] = no λ()
any p (x ∷ xs) with p x
@@ -52,14 +57,15 @@ any p (x ∷ xs) | no ¬px = Dec.map′ there (tail ¬px) (any p xs)
-- index x∈xs is the list position (zero-based) which x∈xs points to.
-index : ∀ {A} {P : A → Set} {xs} → Any P xs → Fin (List.length xs)
+index : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
+ Any P xs → Fin (List.length xs)
index (here px) = zero
index (there pxs) = suc (index pxs)
------------------------------------------------------------------------
-- List membership and some related definitions
-module Membership (S : Setoid zero zero) where
+module Membership {c ℓ : Level} (S : Setoid c ℓ) where
private
open module S = Setoid S using (_≈_) renaming (Carrier to A)
@@ -69,7 +75,8 @@ module Membership (S : Setoid zero zero) where
-- If a predicate P respects the underlying equality then Any P
-- respects the list equality.
- lift-resp : ∀ {P} → P Respects _≈_ → Any P Respects _≋_
+ lift-resp : ∀ {p} {P : A → Set 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) =
@@ -79,20 +86,20 @@ module Membership (S : Setoid zero zero) where
infix 4 _∈_ _∉_
- _∈_ : A → List A → Set
+ _∈_ : A → List A → Set _
x ∈ xs = Any (_≈_ x) xs
- _∉_ : A → List A → Set
+ _∉_ : A → List A → Set _
x ∉ xs = ¬ x ∈ xs
-- Subsets.
infix 4 _⊆_ _⊈_
- _⊆_ : List A → List A → Set
+ _⊆_ : List A → List A → Set _
xs ⊆ ys = ∀ {x} → x ∈ xs → x ∈ ys
- _⊈_ : List A → List A → Set
+ _⊈_ : List A → List A → Set _
xs ⊈ ys = ¬ xs ⊆ ys
-- Equality is respected by the predicate which is used to define
@@ -123,17 +130,20 @@ module Membership (S : Setoid zero zero) where
-- A variant of List.map.
- map-with-∈ : ∀ {B : Set} (xs : List A) → (∀ {x} → x ∈ xs → B) → List B
+ map-with-∈ : ∀ {b} {B : Set b}
+ (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)
-- Finds an element satisfying the predicate.
- find : ∀ {P : A → Set} {xs} → Any P xs → ∃ λ x → x ∈ xs × P x
+ find : ∀ {p} {P : A → Set p} {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 : ∀ {p} {P : A → Set p} {x xs} →
+ P Respects _≈_ → x ∈ xs → P x → Any P xs
lose resp x∈xs px = map (flip resp px) x∈xs
-- The code above instantiated (and slightly changed) for
@@ -142,15 +152,16 @@ module Membership (S : Setoid zero zero) where
module Membership-≡ where
private
- open module M {A : Set} = Membership (PropEq.setoid A) public
+ open module M {a} {A : Set a} = 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 : ∀ {a p} {A : Set a} {P : A → Set p} {x xs} →
+ x ∈ xs → P x → Any P xs
lose {P = P} = M.lose (PropEq.subst P)
-- _⊆_ is a preorder.
- ⊆-preorder : Set → Preorder _ _ _
+ ⊆-preorder : ∀ {a} → Set a → Preorder _ _ _
⊆-preorder A = Ind.InducedPreorder₂ (PropEq.setoid (List A)) _∈_
(PropEq.subst (_∈_ _))
@@ -159,17 +170,17 @@ module Membership-≡ where
open Inv public
using (Kind) renaming (equivalent to set; inverse to bag)
- [_]-Equality : Kind → Set → Setoid _ _
+ [_]-Equality : Kind → ∀ {a} → Set a → 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
+ _≈[_]_ : ∀ {a} {A : Set a} → List A → Kind → List A → Set _
+ _≈[_]_ {A = A} xs k ys = Setoid._≈_ ([ k ]-Equality A) xs ys
-- Bag equality implies set equality.
- bag-=⇒set-= : {A : Set} {xs ys : List A} →
+ bag-=⇒set-= : ∀ {a} {A : Set a} {xs ys : List A} →
xs ≈[ bag ] ys → xs ≈[ set ] ys
bag-=⇒set-= xs≈ys = Inv.⇿⇒ xs≈ys
@@ -178,17 +189,17 @@ module Membership-≡ where
module ⊆-Reasoning where
import Relation.Binary.PreorderReasoning as PreR
private
- open module PR {A : Set} = PreR (⊆-preorder A) public
+ open module PR {a} {A : Set a} = PreR (⊆-preorder A) public
renaming (_∼⟨_⟩_ to _⊆⟨_⟩_; _≈⟨_⟩_ to _≡⟨_⟩_)
infixr 2 _≈⟨_⟩_
infix 1 _∈⟨_⟩_
- _∈⟨_⟩_ : ∀ {A : Set} x {xs ys : List A} →
+ _∈⟨_⟩_ : ∀ {a} {A : Set a} 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} →
+ _≈⟨_⟩_ : ∀ {k a} {A : Set a} 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
@@ -198,5 +209,6 @@ module Membership-≡ where
-- If any element satisfies P, then P is satisfied.
-satisfied : ∀ {A} {P : A → Set} {xs} → Any P xs → ∃ P
+satisfied : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
+ Any P xs → ∃ P
satisfied = Prod.map id Prod.proj₂ ∘ Membership-≡.find
diff --git a/src/Data/List/Any/BagAndSetEquality.agda b/src/Data/List/Any/BagAndSetEquality.agda
index 9a1a916..00c0e34 100644
--- a/src/Data/List/Any/BagAndSetEquality.agda
+++ b/src/Data/List/Any/BagAndSetEquality.agda
@@ -2,6 +2,8 @@
-- Properties related to bag and set equality
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- Bag and set equality are defined in Data.List.Any.
module Data.List.Any.BagAndSetEquality where
@@ -9,34 +11,39 @@ module Data.List.Any.BagAndSetEquality where
open import Algebra
open import Algebra.FunctionProperties
open import Category.Monad
+open import Data.Empty
open import Data.List as List
import Data.List.Properties as LP
-open import Data.List.Any as Any using (Any)
+open import Data.List.Any as Any using (Any); open Any.Any
open import Data.List.Any.Properties
open import Data.Product
open import Data.Sum
+open import Data.Unit
open import Function
open import Function.Equality using (_⟨$⟩_)
import Function.Equivalence as FE
open import Function.Inverse as Inv using (_⇿_; module Inverse)
open import Function.Inverse.TypeIsomorphisms
+open import Level
open import Relation.Binary
import Relation.Binary.EqReasoning as EqR
+import Relation.Binary.Sigma.Pointwise as Σ
open import Relation.Binary.PropositionalEquality as P
- using (_≡_; _≗_)
+ using (_≡_; _≗_; _with-≡_)
+open import Relation.Nullary
open Any.Membership-≡
-open RawMonad List.monad
private
- module ListMonoid {A : Set} = Monoid (List.monoid A)
- module Eq {k} {A : Set} = Setoid ([ k ]-Equality A)
+ module Eq {k a} {A : Set a} = Setoid ([ k ]-Equality A)
module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
+ open module ListMonad {ℓ} = RawMonad (List.monad {ℓ = ℓ})
+ module ListMonoid {a} {A : Set a} = Monoid (List.monoid A)
-- _++_ and [] form a commutative monoid, with either bag or set
-- equality as the underlying equality.
-commutativeMonoid : Kind → Set → CommutativeMonoid _ _
-commutativeMonoid k A = record
+commutativeMonoid : ∀ {a} → Kind → Set a → CommutativeMonoid _ _
+commutativeMonoid {a} k A = record
{ Carrier = List A
; _≈_ = λ xs ys → xs ≈[ k ] ys
; _∙_ = _++_
@@ -45,40 +52,50 @@ commutativeMonoid k A = record
{ isSemigroup = record
{ isEquivalence = Eq.isEquivalence
; assoc = λ xs ys zs →
- Eq.reflexive $ ListMonoid.assoc xs ys zs
+ Eq.reflexive (ListMonoid.assoc xs ys zs)
; ∙-cong = λ {xs₁ xs₂ xs₃ xs₄} xs₁≈xs₂ xs₃≈xs₄ {x} →
- x ∈ xs₁ ++ xs₃ ⇿⟨ sym ++⇿ ⟩
+ x ∈ xs₁ ++ xs₃ ⇿⟨ sym $ ++⇿ {a = a} {p = a} ⟩
(x ∈ xs₁ ⊎ x ∈ xs₃) ≈⟨ xs₁≈xs₂ ⟨ ×⊎.+-cong ⟩ xs₃≈xs₄ ⟩
- (x ∈ xs₂ ⊎ x ∈ xs₄) ⇿⟨ ++⇿ ⟩
+ (x ∈ xs₂ ⊎ x ∈ xs₄) ⇿⟨ ++⇿ {a = a} {p = a} ⟩
x ∈ xs₂ ++ xs₄ ∎
}
; identityˡ = λ xs {x} → x ∈ xs ∎
; comm = λ xs ys {x} →
- x ∈ xs ++ ys ⇿⟨ ++⇿++ xs ys ⟩
+ x ∈ xs ++ ys ⇿⟨ ++⇿++ {a = a} {p = a} xs ys ⟩
x ∈ ys ++ xs ∎
}
}
where open Inv.EquationalReasoning
+-- The only list which is bag or set equal to the empty list is the
+-- empty list itself.
+
+empty-unique : ∀ {k a} {A : Set a} {xs : List A} →
+ xs ≈[ k ] [] → xs ≡ []
+empty-unique {xs = []} _ = P.refl
+empty-unique {xs = _ ∷ _} ∷≈[]
+ with FE.Equivalent.to (Inv.⇒⇔ ∷≈[]) ⟨$⟩ here P.refl
+... | ()
+
-- _++_ is idempotent (under set equality).
-++-idempotent : {A : Set} →
+++-idempotent : ∀ {a} {A : Set a} →
Idempotent (λ (xs ys : List A) → xs ≈[ set ] ys) _++_
-++-idempotent xs {x} =
- x ∈ xs ++ xs ≈⟨ FE.equivalent ([ id , id ]′ ∘ _⟨$⟩_ (Inverse.from ++⇿))
- (_⟨$⟩_ (Inverse.to ++⇿) ∘ inj₁) ⟩
+++-idempotent {a} xs {x} =
+ x ∈ xs ++ xs ≈⟨ FE.equivalent ([ id , id ]′ ∘ _⟨$⟩_ (Inverse.from $ ++⇿ {a = a} {p = a}))
+ (_⟨$⟩_ (Inverse.to $ ++⇿ {a = a} {p = a}) ∘ inj₁) ⟩
x ∈ xs ∎
where open Inv.EquationalReasoning
-- List.map is a congruence.
-map-cong : ∀ {k} {A B : Set} {f₁ f₂ : A → B} {xs₁ xs₂} →
+map-cong : ∀ {ℓ k} {A B : Set ℓ} {f₁ f₂ : A → B} {xs₁ xs₂} →
f₁ ≗ f₂ → xs₁ ≈[ k ] xs₂ →
List.map f₁ xs₁ ≈[ k ] List.map f₂ xs₂
-map-cong {f₁ = f₁} {f₂} {xs₁} {xs₂} f₁≗f₂ xs₁≈xs₂ {x} =
- x ∈ List.map f₁ xs₁ ⇿⟨ sym map⇿ ⟩
+map-cong {ℓ} {f₁ = f₁} {f₂} {xs₁} {xs₂} f₁≗f₂ xs₁≈xs₂ {x} =
+ x ∈ List.map f₁ xs₁ ⇿⟨ sym $ map⇿ {a = ℓ} {b = ℓ} {p = ℓ} ⟩
Any (λ y → x ≡ f₁ y) xs₁ ≈⟨ Any-cong (Inv.⇿⇒ ∘ helper) xs₁≈xs₂ ⟩
- Any (λ y → x ≡ f₂ y) xs₂ ⇿⟨ map⇿ ⟩
+ Any (λ y → x ≡ f₂ y) xs₂ ⇿⟨ map⇿ {a = ℓ} {b = ℓ} {p = ℓ} ⟩
x ∈ List.map f₂ xs₂ ∎
where
open Inv.EquationalReasoning
@@ -95,30 +112,30 @@ map-cong {f₁ = f₁} {f₂} {xs₁} {xs₂} f₁≗f₂ xs₁≈xs₂ {x} =
-- concat is a congruence.
-concat-cong : ∀ {k} {A : Set} {xss₁ xss₂ : List (List A)} →
+concat-cong : ∀ {a k} {A : Set a} {xss₁ xss₂ : List (List A)} →
xss₁ ≈[ k ] xss₂ → concat xss₁ ≈[ k ] concat xss₂
-concat-cong {xss₁ = xss₁} {xss₂} xss₁≈xss₂ {x} =
- x ∈ concat xss₁ ⇿⟨ sym concat⇿ ⟩
+concat-cong {a} {xss₁ = xss₁} {xss₂} xss₁≈xss₂ {x} =
+ x ∈ concat xss₁ ⇿⟨ sym $ concat⇿ {a = a} {p = a} ⟩
Any (Any (_≡_ x)) xss₁ ≈⟨ Any-cong (λ _ → _ ∎) xss₁≈xss₂ ⟩
- Any (Any (_≡_ x)) xss₂ ⇿⟨ concat⇿ ⟩
+ Any (Any (_≡_ x)) xss₂ ⇿⟨ concat⇿ {a = a} {p = a} ⟩
x ∈ concat xss₂ ∎
where open Inv.EquationalReasoning
-- The list monad's bind is a congruence.
->>=-cong : ∀ {k} {A B : Set} {xs₁ xs₂} {f₁ f₂ : A → List B} →
+>>=-cong : ∀ {ℓ k} {A B : Set ℓ} {xs₁ xs₂} {f₁ f₂ : A → List B} →
xs₁ ≈[ k ] xs₂ → (∀ x → f₁ x ≈[ k ] f₂ x) →
(xs₁ >>= f₁) ≈[ k ] (xs₂ >>= f₂)
->>=-cong {xs₁ = xs₁} {xs₂} {f₁} {f₂} xs₁≈xs₂ f₁≈f₂ {x} =
- x ∈ (xs₁ >>= f₁) ⇿⟨ sym >>=⇿ ⟩
+>>=-cong {ℓ} {xs₁ = xs₁} {xs₂} {f₁} {f₂} xs₁≈xs₂ f₁≈f₂ {x} =
+ x ∈ (xs₁ >>= f₁) ⇿⟨ sym $ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟩
Any (λ y → x ∈ f₁ y) xs₁ ≈⟨ Any-cong (λ x → f₁≈f₂ x) xs₁≈xs₂ ⟩
- Any (λ y → x ∈ f₂ y) xs₂ ⇿⟨ >>=⇿ ⟩
+ Any (λ y → x ∈ f₂ y) xs₂ ⇿⟨ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟩
x ∈ (xs₂ >>= f₂) ∎
where open Inv.EquationalReasoning
-- _⊛_ is a congruence.
-⊛-cong : ∀ {k A B} {fs₁ fs₂ : List (A → B)} {xs₁ xs₂} →
+⊛-cong : ∀ {ℓ k} {A B : Set ℓ} {fs₁ fs₂ : List (A → B)} {xs₁ xs₂} →
fs₁ ≈[ k ] fs₂ → xs₁ ≈[ k ] xs₂ → fs₁ ⊛ xs₁ ≈[ k ] fs₂ ⊛ xs₂
⊛-cong fs₁≈fs₂ xs₁≈xs₂ =
>>=-cong fs₁≈fs₂ λ f →
@@ -128,32 +145,34 @@ concat-cong {xss₁ = xss₁} {xss₂} xss₁≈xss₂ {x} =
-- _⊗_ is a congruence.
-⊗-cong : ∀ {k A B} {xs₁ xs₂ : List A} {ys₁ ys₂ : List B} →
+⊗-cong : ∀ {ℓ k} {A B : Set ℓ} {xs₁ xs₂ : List A} {ys₁ ys₂ : List B} →
xs₁ ≈[ k ] xs₂ → ys₁ ≈[ k ] ys₂ →
(xs₁ ⊗ ys₁) ≈[ k ] (xs₂ ⊗ ys₂)
-⊗-cong xs₁≈xs₂ ys₁≈ys₂ =
- ⊛-cong (⊛-cong (Eq.refl {x = [ _,_ ]}) xs₁≈xs₂) ys₁≈ys₂
+⊗-cong {ℓ} xs₁≈xs₂ ys₁≈ys₂ =
+ ⊛-cong (⊛-cong (Eq.refl {x = [ _,_ {a = ℓ} {b = ℓ} ]})
+ xs₁≈xs₂)
+ ys₁≈ys₂
-- The list monad's bind distributes from the left over _++_.
>>=-left-distributive :
- ∀ {A B : Set} (xs : List A) {f g : A → List B} →
+ ∀ {ℓ} {A B : Set ℓ} (xs : List A) {f g : A → List B} →
(xs >>= λ x → f x ++ g x) ≈[ bag ] (xs >>= f) ++ (xs >>= g)
->>=-left-distributive xs {f} {g} {y} =
- y ∈ (xs >>= λ x → f x ++ g x) ⇿⟨ sym >>=⇿ ⟩
- Any (λ x → y ∈ f x ++ g x) xs ⇿⟨ sym (Any-cong (λ _ → ++⇿) (_ ∎)) ⟩
- Any (λ x → y ∈ f x ⊎ y ∈ g x) xs ⇿⟨ sym ⊎⇿ ⟩
- (Any (λ x → y ∈ f x) xs ⊎ Any (λ x → y ∈ g x) xs) ⇿⟨ >>=⇿ ⟨ ×⊎.+-cong ⟩ >>=⇿ ⟩
- (y ∈ (xs >>= f) ⊎ y ∈ (xs >>= g)) ⇿⟨ ++⇿ ⟩
+>>=-left-distributive {ℓ} xs {f} {g} {y} =
+ y ∈ (xs >>= λ x → f x ++ g x) ⇿⟨ sym $ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟩
+ Any (λ x → y ∈ f x ++ g x) xs ⇿⟨ sym (Any-cong (λ _ → ++⇿ {a = ℓ} {p = ℓ}) (_ ∎)) ⟩
+ Any (λ x → y ∈ f x ⊎ y ∈ g x) xs ⇿⟨ sym $ ⊎⇿ {a = ℓ} {p = ℓ} {q = ℓ} ⟩
+ (Any (λ x → y ∈ f x) xs ⊎ Any (λ x → y ∈ g x) xs) ⇿⟨ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟨ ×⊎.+-cong {ℓ = ℓ} ⟩ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟩
+ (y ∈ (xs >>= f) ⊎ y ∈ (xs >>= g)) ⇿⟨ ++⇿ {a = ℓ} {p = ℓ} ⟩
y ∈ (xs >>= f) ++ (xs >>= g) ∎
where open Inv.EquationalReasoning
-- The same applies to _⊛_.
⊛-left-distributive :
- ∀ {A B} (fs : List (A → B)) xs₁ xs₂ →
+ ∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) xs₁ xs₂ →
fs ⊛ (xs₁ ++ xs₂) ≈[ bag ] (fs ⊛ xs₁) ++ (fs ⊛ xs₂)
-⊛-left-distributive fs xs₁ xs₂ = begin
+⊛-left-distributive {B = B} fs xs₁ xs₂ = begin
fs ⊛ (xs₁ ++ xs₂) ≡⟨ P.refl ⟩
(fs >>= λ f → xs₁ ++ xs₂ >>= return ∘ f) ≡⟨ (LP.Monad.cong (P.refl {x = fs}) λ f →
LP.Monad.right-distributive xs₁ xs₂ (return ∘ f)) ⟩
@@ -164,4 +183,142 @@ concat-cong {xss₁ = xss₁} {xss₂} xss₁≈xss₂ {x} =
(fs >>= λ f → xs₂ >>= return ∘ f) ≡⟨ P.refl ⟩
(fs ⊛ xs₁) ++ (fs ⊛ xs₂) ∎
- where open EqR ([ bag ]-Equality _)
+ where open EqR ([ bag ]-Equality B)
+
+private
+
+ -- If x ∷ xs is set equal to x ∷ ys, then xs and ys are not
+ -- necessarily set equal.
+
+ ¬-drop-cons :
+ ∀ {a} {A : Set a} {x : A} →
+ ¬ (∀ {xs ys} → x ∷ xs ≈[ set ] x ∷ ys → xs ≈[ set ] ys)
+ ¬-drop-cons {x = x} drop-cons
+ with FE.Equivalent.to x≈[] ⟨$⟩ here P.refl
+ where
+ x,x≈x : (x ∷ x ∷ []) ≈[ set ] [ x ]
+ x,x≈x = ++-idempotent [ x ]
+
+ x≈[] : [ x ] ≈[ set ] []
+ x≈[] = drop-cons x,x≈x
+ ... | ()
+
+-- However, the corresponding property does hold for bag equality.
+
+drop-cons : ∀ {a} {A : Set a} {x : A} {xs ys} →
+ x ∷ xs ≈[ bag ] x ∷ ys → xs ≈[ bag ] ys
+drop-cons {A = A} {x} {xs} {ys} x∷xs≈x∷ys {z} =
+ z ∈ xs ⇿⟨ lemma xs ⟩
+ (∃ λ (z∈x∷xs : z ∈ x ∷ xs) → There z∈x∷xs) ⇿⟨ Σ.⇿ x∷xs≈x∷ys′ x∷xs≈x∷ys′-preserves-There ⟩
+ (∃ λ (z∈x∷ys : z ∈ x ∷ ys) → There z∈x∷ys) ⇿⟨ sym $ lemma ys ⟩
+ z ∈ ys ∎
+ where
+ open Inv.EquationalReasoning
+
+ -- Inhabited for there but not here.
+
+ There : ∀ {z : A} {xs} → z ∈ xs → Set
+ There (here _) = ⊥
+ There (there _) = ⊤
+
+ -- There is proof irrelevant.
+
+ proof-irrelevance : ∀ {z xs} {z∈xs : z ∈ xs}
+ (p q : There z∈xs) → p ≡ q
+ proof-irrelevance {z∈xs = here _} () ()
+ proof-irrelevance {z∈xs = there _} _ _ = P.refl
+
+ -- An isomorphism.
+
+ lemma : ∀ xs → z ∈ xs ⇿ (∃ λ (z∈x∷xs : z ∈ x ∷ xs) → There z∈x∷xs)
+ lemma xs = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.refl
+ ; right-inverse-of = to∘from
+ }
+ }
+ where
+ to : z ∈ xs → (∃ λ (z∈x∷xs : z ∈ x ∷ xs) → There z∈x∷xs)
+ to z∈xs = (there z∈xs , _)
+
+ from : (∃ λ (z∈x∷xs : z ∈ x ∷ xs) → There z∈x∷xs) → z ∈ xs
+ from (here z≡x , ())
+ from (there z∈xs , _) = z∈xs
+
+ to∘from : ∀ p → to (from p) ≡ p
+ to∘from (here z≡x , ())
+ to∘from (there z∈xs , _) = P.refl
+
+ -- The isomorphisms x∷xs≈x∷ys may not preserve There, because they
+ -- could map here P.refl to something other than here P.refl.
+ -- However, we can construct isomorphisms which do preserve There:
+
+ x∷xs≈x∷ys′ : x ∷ xs ≈[ bag ] x ∷ ys
+ x∷xs≈x∷ys′ = record
+ { to = P.→-to-⟶ $ f x∷xs≈x∷ys
+ ; from = P.→-to-⟶ $ f $ Inv.sym x∷xs≈x∷ys
+ ; inverse-of = record
+ { left-inverse-of = f∘f x∷xs≈x∷ys
+ ; right-inverse-of = f∘f $ Inv.sym x∷xs≈x∷ys
+ }
+ }
+ where
+ f : ∀ {xs ys} → x ∷ xs ≈[ bag ] x ∷ ys →
+ ∀ {z} → z ∈ x ∷ xs → z ∈ x ∷ ys
+ f x∷xs≈x∷ys (here P.refl) = here P.refl
+ f x∷xs≈x∷ys (there z∈xs) with Inverse.to x∷xs≈x∷ys ⟨$⟩ there z∈xs
+ ... | there z∈ys = there z∈ys
+ ... | here P.refl = Inverse.to x∷xs≈x∷ys ⟨$⟩ here P.refl
+
+ f∘f : ∀ {xs ys}
+ (x∷xs≈x∷ys : x ∷ xs ≈[ bag ] x ∷ ys) →
+ ∀ {z} (p : z ∈ x ∷ xs) →
+ f (Inv.sym x∷xs≈x∷ys) (f x∷xs≈x∷ys p) ≡ p
+ f∘f x∷xs≈x∷ys (here P.refl) = P.refl
+ f∘f x∷xs≈x∷ys (there z∈xs)
+ with Inverse.to x∷xs≈x∷ys ⟨$⟩ there z∈xs
+ | Inverse.left-inverse-of x∷xs≈x∷ys (there z∈xs)
+ ... | there z∈ys | left rewrite left = P.refl
+ ... | here P.refl | left
+ with Inverse.to x∷xs≈x∷ys ⟨$⟩ here P.refl
+ | Inverse.left-inverse-of x∷xs≈x∷ys (here P.refl)
+ ... | there z∈xs′ | left′ rewrite left′ = left
+ ... | here P.refl | left′ rewrite left′ = left
+
+ x∷xs≈x∷ys′-preserves-There :
+ ∀ {z} {z∈x∷xs : z ∈ x ∷ xs} →
+ There z∈x∷xs ⇿ There (Inverse.to x∷xs≈x∷ys′ ⟨$⟩ z∈x∷xs)
+ x∷xs≈x∷ys′-preserves-There {z∈x∷xs = z∈x∷xs} = record
+ { to = P.→-to-⟶ $ to z∈x∷xs
+ ; from = P.→-to-⟶ $ from z∈x∷xs
+ ; inverse-of = record
+ { left-inverse-of = λ _ → proof-irrelevance _ _
+ ; right-inverse-of = λ _ → proof-irrelevance _ _
+ }
+ }
+ where
+ open P.≡-Reasoning renaming (_∎ to _□)
+
+ to : ∀ {z} (z∈x∷xs : z ∈ x ∷ xs) →
+ There z∈x∷xs → There (Inverse.to x∷xs≈x∷ys′ ⟨$⟩ z∈x∷xs)
+ to (here _) ()
+ to (there z∈) _ with P.inspect (Inverse.to x∷xs≈x∷ys ⟨$⟩ there z∈)
+ ... | there _ with-≡ eq rewrite eq = _
+ ... | here P.refl with-≡ eq rewrite eq
+ with P.inspect (Inverse.to x∷xs≈x∷ys ⟨$⟩ here P.refl)
+ ... | there _ with-≡ eq′ rewrite eq′ = _
+ ... | here P.refl with-≡ eq′ rewrite eq′
+ with begin
+ there z∈ ≡⟨ P.sym $ Inverse.left-inverse-of x∷xs≈x∷ys _ ⟩
+ Inverse.from x∷xs≈x∷ys ⟨$⟩ (Inverse.to x∷xs≈x∷ys ⟨$⟩ there z∈) ≡⟨ P.cong (_⟨$⟩_ (Inverse.from x∷xs≈x∷ys)) eq ⟩
+ Inverse.from x∷xs≈x∷ys ⟨$⟩ here P.refl ≡⟨ P.cong (_⟨$⟩_ (Inverse.from x∷xs≈x∷ys)) (P.sym eq′) ⟩
+ Inverse.from x∷xs≈x∷ys ⟨$⟩ (Inverse.to x∷xs≈x∷ys ⟨$⟩ here P.refl) ≡⟨ Inverse.left-inverse-of x∷xs≈x∷ys _ ⟩
+ Any.here {xs = xs} (P.refl {x = x}) □
+ ... | ()
+
+ from : ∀ {z} (z∈x∷xs : z ∈ x ∷ xs) →
+ There (Inverse.to x∷xs≈x∷ys′ ⟨$⟩ z∈x∷xs) → There z∈x∷xs
+ from (here P.refl) ()
+ from (there z∈xs) _ = _
diff --git a/src/Data/List/Any/Membership.agda b/src/Data/List/Any/Membership.agda
index a7c5d19..c98235d 100644
--- a/src/Data/List/Any/Membership.agda
+++ b/src/Data/List/Any/Membership.agda
@@ -2,6 +2,8 @@
-- Properties related to list membership
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- List membership is defined in Data.List.Any. This module does not
-- treat the general variant of list membership, parametrised on a
-- setoid, only the variant where the equality is fixed to be
@@ -37,44 +39,45 @@ open import Relation.Nullary
open import Relation.Nullary.Negation
open Any.Membership-≡
-open RawMonad List.monad
private
module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
+ open module ListMonad {ℓ} = RawMonad (List.monad {ℓ = ℓ})
------------------------------------------------------------------------
-- Properties relating _∈_ to various list functions
-map-∈⇿ : ∀ {A B : Set} {f : A → B} {y xs} →
+map-∈⇿ : ∀ {a b} {A : Set a} {B : Set b} {f : A → B} {y xs} →
(∃ λ x → x ∈ xs × y ≡ f x) ⇿ y ∈ List.map f xs
-map-∈⇿ {f = f} {y} {xs} =
- (∃ λ x → x ∈ xs × y ≡ f x) ⇿⟨ Any⇿ ⟩
- Any (λ x → y ≡ f x) xs ⇿⟨ map⇿ ⟩
+map-∈⇿ {a} {b} {f = f} {y} {xs} =
+ (∃ λ x → x ∈ xs × y ≡ f x) ⇿⟨ Any⇿ {a = a} {p = b} ⟩
+ Any (λ x → y ≡ f x) xs ⇿⟨ map⇿ {a = a} {b = b} {p = b} ⟩
y ∈ List.map f xs ∎
where open Inv.EquationalReasoning
-concat-∈⇿ : ∀ {A : Set} {x : A} {xss} →
+concat-∈⇿ : ∀ {a} {A : Set a} {x : A} {xss} →
(∃ λ xs → x ∈ xs × xs ∈ xss) ⇿ x ∈ concat xss
-concat-∈⇿ {x = x} {xss} =
- (∃ λ xs → x ∈ xs × xs ∈ xss) ⇿⟨ Σ.cong $ ×⊎.*-comm _ _ ⟩
- (∃ λ xs → xs ∈ xss × x ∈ xs) ⇿⟨ Any⇿ ⟩
- Any (Any (_≡_ x)) xss ⇿⟨ concat⇿ ⟩
+concat-∈⇿ {a} {x = x} {xss} =
+ (∃ λ xs → x ∈ xs × xs ∈ xss) ⇿⟨ Σ.cong {a₁ = a} {b₁ = a} {b₂ = a} Inv.id $ ×⊎.*-comm _ _ ⟩
+ (∃ λ xs → xs ∈ xss × x ∈ xs) ⇿⟨ Any⇿ {a = a} {p = a} ⟩
+ Any (Any (_≡_ x)) xss ⇿⟨ concat⇿ {a = a} {p = a} ⟩
x ∈ concat xss ∎
where open Inv.EquationalReasoning
->>=-∈⇿ : ∀ {A B : Set} {xs} {f : A → List B} {y} →
+>>=-∈⇿ : ∀ {ℓ} {A B : Set ℓ} {xs} {f : A → List B} {y} →
(∃ λ x → x ∈ xs × y ∈ f x) ⇿ y ∈ (xs >>= f)
->>=-∈⇿ {xs = xs} {f} {y} =
- (∃ λ x → x ∈ xs × y ∈ f x) ⇿⟨ Any⇿ ⟩
- Any (Any (_≡_ y) ∘ f) xs ⇿⟨ >>=⇿ ⟩
+>>=-∈⇿ {ℓ} {xs = xs} {f} {y} =
+ (∃ λ x → x ∈ xs × y ∈ f x) ⇿⟨ Any⇿ {a = ℓ} {p = ℓ} ⟩
+ Any (Any (_≡_ y) ∘ f) xs ⇿⟨ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟩
y ∈ (xs >>= f) ∎
where open Inv.EquationalReasoning
-⊛-∈⇿ : ∀ {A B : Set} (fs : List (A → B)) {xs y} →
+⊛-∈⇿ : ∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) {xs y} →
(∃₂ λ f x → f ∈ fs × x ∈ xs × y ≡ f x) ⇿ y ∈ fs ⊛ xs
-⊛-∈⇿ fs {xs} {y} =
- (∃₂ λ f x → f ∈ fs × x ∈ xs × y ≡ f x) ⇿⟨ Σ.cong (∃∃⇿∃∃ _) ⟩
- (∃ λ f → f ∈ fs × ∃ λ x → x ∈ xs × y ≡ f x) ⇿⟨ Σ.cong ((_ ∎) ⟨ ×⊎.*-cong ⟩ Any⇿) ⟩
- (∃ λ f → f ∈ fs × Any (_≡_ y ∘ f) xs) ⇿⟨ Any⇿ ⟩
+⊛-∈⇿ {ℓ} fs {xs} {y} =
+ (∃₂ λ f x → f ∈ fs × x ∈ xs × y ≡ f x) ⇿⟨ Σ.cong {a₁ = ℓ} {b₁ = ℓ} {b₂ = ℓ} Inv.id (∃∃⇿∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
+ (∃ λ f → f ∈ fs × ∃ λ x → x ∈ xs × y ≡ f x) ⇿⟨ Σ.cong {a₁ = ℓ} {b₁ = ℓ} {b₂ = ℓ}
+ Inv.id ((_ ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any⇿ {a = ℓ} {p = ℓ}) ⟩
+ (∃ λ f → f ∈ fs × Any (_≡_ y ∘ f) xs) ⇿⟨ Any⇿ {a = ℓ} {p = ℓ} ⟩
Any (λ f → Any (_≡_ y ∘ f) xs) fs ⇿⟨ ⊛⇿ ⟩
y ∈ fs ⊛ xs ∎
where open Inv.EquationalReasoning
@@ -102,43 +105,44 @@ concat-∈⇿ {x = x} {xss} =
------------------------------------------------------------------------
-- Various functions are monotone
-mono : ∀ {A : Set} {P : A → Set} {xs ys} →
+mono : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
xs ⊆ ys → Any P xs → Any P ys
mono xs⊆ys =
- _⟨$⟩_ (Inverse.to Any⇿) ∘
+ _⟨$⟩_ (Inverse.to Any⇿) ∘′
Prod.map id (Prod.map xs⊆ys id) ∘
_⟨$⟩_ (Inverse.from Any⇿)
-map-mono : ∀ {A B : Set} (f : A → B) {xs ys} →
+map-mono : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) {xs ys} →
xs ⊆ ys → List.map f xs ⊆ List.map f ys
map-mono f xs⊆ys =
_⟨$⟩_ (Inverse.to map-∈⇿) ∘
Prod.map id (Prod.map xs⊆ys id) ∘
_⟨$⟩_ (Inverse.from map-∈⇿)
-_++-mono_ : ∀ {A : Set} {xs₁ xs₂ ys₁ ys₂ : List A} →
+_++-mono_ : ∀ {a} {A : Set a} {xs₁ xs₂ ys₁ ys₂ : List A} →
xs₁ ⊆ ys₁ → xs₂ ⊆ ys₂ → xs₁ ++ xs₂ ⊆ ys₁ ++ ys₂
_++-mono_ xs₁⊆ys₁ xs₂⊆ys₂ =
_⟨$⟩_ (Inverse.to ++⇿) ∘
Sum.map xs₁⊆ys₁ xs₂⊆ys₂ ∘
_⟨$⟩_ (Inverse.from ++⇿)
-concat-mono : ∀ {A : Set} {xss yss : List (List A)} →
+concat-mono : ∀ {a} {A : Set a} {xss yss : List (List A)} →
xss ⊆ yss → concat xss ⊆ concat yss
-concat-mono xss⊆yss =
- _⟨$⟩_ (Inverse.to concat-∈⇿) ∘
+concat-mono {a} xss⊆yss =
+ _⟨$⟩_ (Inverse.to $ concat-∈⇿ {a = a}) ∘
Prod.map id (Prod.map id xss⊆yss) ∘
- _⟨$⟩_ (Inverse.from concat-∈⇿)
+ _⟨$⟩_ (Inverse.from $ concat-∈⇿ {a = a})
->>=-mono : ∀ {A B : Set} (f g : A → List B) {xs ys} →
+>>=-mono : ∀ {ℓ} {A B : Set ℓ} (f g : A → List B) {xs ys} →
xs ⊆ ys → (∀ {x} → f x ⊆ g x) →
(xs >>= f) ⊆ (ys >>= g)
->>=-mono f g xs⊆ys f⊆g =
- _⟨$⟩_ (Inverse.to >>=-∈⇿) ∘
+>>=-mono {ℓ} f g xs⊆ys f⊆g =
+ _⟨$⟩_ (Inverse.to $ >>=-∈⇿ {ℓ = ℓ}) ∘
Prod.map id (Prod.map xs⊆ys f⊆g) ∘
- _⟨$⟩_ (Inverse.from >>=-∈⇿)
+ _⟨$⟩_ (Inverse.from $ >>=-∈⇿ {ℓ = ℓ})
-_⊛-mono_ : ∀ {A B : Set} {fs gs : List (A → B)} {xs ys} →
+_⊛-mono_ : ∀ {ℓ} {A B : Set ℓ}
+ {fs gs : List (A → B)} {xs ys : List A} →
fs ⊆ gs → xs ⊆ ys → (fs ⊛ xs) ⊆ (gs ⊛ ys)
_⊛-mono_ {fs = fs} {gs} fs⊆gs xs⊆ys =
_⟨$⟩_ (Inverse.to $ ⊛-∈⇿ gs) ∘
@@ -152,16 +156,17 @@ xs₁⊆ys₁ ⊗-mono xs₂⊆ys₂ =
Prod.map xs₁⊆ys₁ xs₂⊆ys₂ ∘
_⟨$⟩_ (Inverse.from ⊗-∈⇿)
-any-mono : ∀ {A : Set} (p : A → Bool) →
+any-mono : ∀ {a} {A : Set a} (p : A → Bool) →
∀ {xs ys} → xs ⊆ ys → T (any p xs) → T (any p ys)
-any-mono p xs⊆ys =
- _⟨$⟩_ (Equivalent.to any⇔) ∘
+any-mono {a} p xs⊆ys =
+ _⟨$⟩_ (Equivalent.to $ any⇔ {a = a}) ∘
mono xs⊆ys ∘
- _⟨$⟩_ (Equivalent.from any⇔)
+ _⟨$⟩_ (Equivalent.from $ any⇔ {a = a})
map-with-∈-mono :
- {A B : Set} {xs : List A} {f : ∀ {x} → x ∈ xs → B}
- {ys : List A} {g : ∀ {x} → x ∈ ys → B}
+ ∀ {a b} {A : Set a} {B : Set b}
+ {xs : List A} {f : ∀ {x} → x ∈ xs → B}
+ {ys : List A} {g : ∀ {x} → x ∈ ys → B}
(xs⊆ys : xs ⊆ ys) → (∀ {x} → f {x} ≗ g ∘ xs⊆ys) →
map-with-∈ xs f ⊆ map-with-∈ ys g
map-with-∈-mono {f = f} {g = g} xs⊆ys f≈g {x} =
@@ -179,11 +184,12 @@ map-with-∈-mono {f = f} {g = g} xs⊆ys f≈g {x} =
-- Only a finite number of distinct elements can be members of a
-- given list.
-finite : {A : Set} (f : Inj.Injection (P.setoid ℕ) (P.setoid A)) →
+finite : ∀ {a} {A : Set a}
+ (f : Inj.Injection (P.setoid ℕ) (P.setoid A)) →
∀ xs → ¬ (∀ i → Inj.Injection.to f ⟨$⟩ i ∈ xs)
-finite inj [] ∈[] with ∈[] zero
+finite inj [] ∈[] with ∈[] zero
... | ()
-finite {A} inj (x ∷ xs) ∈x∷xs = excluded-middle helper
+finite {A = A} inj (x ∷ xs) ∈x∷xs = excluded-middle helper
where
open Inj.Injection inj
@@ -235,4 +241,7 @@ finite {A} inj (x ∷ xs) ∈x∷xs = excluded-middle helper
... | tri> _ _ (j≤i , _) | tri≈ _ i≡k _ = ⊥-elim (lemma (subst (_≤_ j) i≡k j≤i) $ sym $ injective eq)
... | tri> _ _ (j≤i , _) | tri> _ _ (k≤i , _) = injective eq
- inj′ = record { to = →-to-⟶ f; injective = injective′ }
+ inj′ = record
+ { to = →-to-⟶ {B = P.setoid A} f
+ ; injective = injective′
+ }
diff --git a/src/Data/List/Any/Properties.agda b/src/Data/List/Any/Properties.agda
index f4b5d9b..8ff99e6 100644
--- a/src/Data/List/Any/Properties.agda
+++ b/src/Data/List/Any/Properties.agda
@@ -2,16 +2,23 @@
-- Properties related to Any
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- The other modules under Data.List.Any also contain properties
-- related to Any.
module Data.List.Any.Properties where
open import Algebra
+import Algebra.FunctionProperties as FP
open import Category.Monad
open import Data.Bool
open import Data.Bool.Properties
open import Data.Empty
+open import Data.List as List
+open import Data.List.Any as Any using (Any; here; there)
+open import Data.Product as Prod hiding (swap)
+open import Data.Sum as Sum using (_⊎_; inj₁; inj₂; [_,_]′)
open import Function
open import Function.Equality using (_⟨$⟩_)
open import Function.Equivalence
@@ -19,10 +26,7 @@ open import Function.Equivalence
open import Function.Inverse as Inv
using (Isomorphism; _⇿_; module Inverse)
open import Function.Inverse.TypeIsomorphisms
-open import Data.List as List
-open import Data.List.Any as Any using (Any; here; there)
-open import Data.Product as Prod hiding (swap)
-open import Data.Sum as Sum using (_⊎_; inj₁; inj₂; [_,_]′)
+open import Level
open import Relation.Binary
import Relation.Binary.HeterogeneousEquality as H
open import Relation.Binary.PropositionalEquality as P
@@ -32,22 +36,24 @@ import Relation.Binary.Sigma.Pointwise as Σ
open Any.Membership-≡
open Inv.EquationalReasoning
-open RawMonad List.monad
private
module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
+ open module ListMonad {ℓ} = RawMonad (List.monad {ℓ = ℓ})
------------------------------------------------------------------------
-- Some lemmas related to map, find and lose
-- Any.map is functorial.
-map-id : ∀ {A} {P : A → Set} (f : P ⋐ P) {xs} →
+map-id : ∀ {a p} {A : Set a} {P : A → Set p} (f : P ⋐ P) {xs} →
(∀ {x} (p : P x) → f p ≡ p) →
(p : Any P xs) → Any.map f p ≡ p
map-id f hyp (here p) = P.cong here (hyp p)
map-id f hyp (there p) = P.cong there $ map-id f hyp p
-map-∘ : ∀ {A} {P Q R : A → Set} (f : Q ⋐ R) (g : P ⋐ Q)
+map-∘ : ∀ {a p q r}
+ {A : Set a} {P : A → Set p} {Q : A → Set q} {R : A → Set r}
+ (f : Q ⋐ R) (g : P ⋐ Q)
{xs} (p : Any P xs) →
Any.map (f ∘ g) p ≡ Any.map f (Any.map g p)
map-∘ f g (here p) = refl
@@ -55,7 +61,7 @@ map-∘ f g (there p) = P.cong there $ map-∘ f g p
-- Lemmas relating map and find.
-map∘find : ∀ {A : Set} {P : A → Set} {xs}
+map∘find : ∀ {a p} {A : Set a} {P : A → Set p} {xs}
(p : Any P xs) → let p′ = find p in
{f : _≡_ (proj₁ p′) ⋐ P} →
f refl ≡ proj₂ (proj₂ p′) →
@@ -63,8 +69,8 @@ map∘find : ∀ {A : Set} {P : A → Set} {xs}
map∘find (here p) hyp = P.cong here hyp
map∘find (there p) hyp = P.cong there (map∘find p hyp)
-find∘map : ∀ {A : Set} {P Q : A → Set} {xs}
- (p : Any P xs) (f : P ⋐ Q) →
+find∘map : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q}
+ {xs : List A} (p : Any P xs) (f : P ⋐ Q) →
find (Any.map f p) ≡ Prod.map id (Prod.map id f) (find p)
find∘map (here p) f = refl
find∘map (there p) f rewrite find∘map p f = refl
@@ -72,7 +78,7 @@ find∘map (there p) f rewrite find∘map p f = refl
-- find satisfies a simple equality when the predicate is a
-- propositional equality.
-find-∈ : ∀ {A : Set} {x : A} {xs} (x∈xs : x ∈ xs) →
+find-∈ : ∀ {a} {A : Set a} {x : A} {xs : List A} (x∈xs : x ∈ xs) →
find x∈xs ≡ (x , x∈xs , refl)
find-∈ (here refl) = refl
find-∈ (there x∈xs) rewrite find-∈ x∈xs = refl
@@ -81,13 +87,14 @@ private
-- find and lose are inverses (more or less).
- lose∘find : ∀ {A : Set} {P : A → Set} {xs} (p : Any P xs) →
- uncurry′ lose (proj₂ $ find p) ≡ p
+ lose∘find : ∀ {a p} {A : Set a} {P : A → Set p} {xs : List A}
+ (p : Any P xs) →
+ uncurry′ lose (proj₂ (find p)) ≡ p
lose∘find p = map∘find p P.refl
- find∘lose : ∀ {A : Set} (P : A → Set) {x xs}
- (x∈xs : x ∈ xs) (p : P x) →
- find {P = P} (lose x∈xs p) ≡ (x , x∈xs , p)
+ find∘lose : ∀ {a p} {A : Set a} (P : A → Set 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 (P.subst P) p)
| find-∈ x∈xs
@@ -95,28 +102,31 @@ private
-- Any can be expressed using _∈_.
-Any⇿ : ∀ {A : Set} {P : A → Set} {xs} →
+Any⇿ : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
(∃ λ x → x ∈ xs × P x) ⇿ Any P xs
-Any⇿ {P = P} = record
- { to = P.→-to-⟶ (uncurry′ lose ∘ proj₂)
- ; from = P.→-to-⟶ find
+Any⇿ {P = P} {xs} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ (find {P = P})
; inverse-of = record
{ left-inverse-of = λ p →
find∘lose P (proj₁ (proj₂ p)) (proj₂ (proj₂ p))
; right-inverse-of = lose∘find
}
}
+ where
+ to : (∃ λ x → x ∈ xs × P x) → Any P xs
+ to = uncurry′ lose ∘ proj₂
------------------------------------------------------------------------
-- Any is a congruence
-Any-cong : ∀ {k} {A : Set} {P₁ P₂ : A → Set} {xs₁ xs₂ : List A} →
+Any-cong : ∀ {k ℓ} {A : Set ℓ} {P₁ P₂ : A → Set ℓ} {xs₁ xs₂ : List A} →
(∀ x → Isomorphism k (P₁ x) (P₂ x)) → xs₁ ≈[ k ] xs₂ →
Isomorphism k (Any P₁ xs₁) (Any P₂ xs₂)
Any-cong {P₁ = P₁} {P₂} {xs₁} {xs₂} P₁⇿P₂ xs₁≈xs₂ =
- Any P₁ xs₁ ⇿⟨ sym Any⇿ ⟩
- (∃ λ x → x ∈ xs₁ × P₁ x) ≈⟨ Σ.cong (xs₁≈xs₂ ⟨ ×⊎.*-cong ⟩ P₁⇿P₂ _) ⟩
- (∃ λ x → x ∈ xs₂ × P₂ x) ⇿⟨ Any⇿ ⟩
+ Any P₁ xs₁ ⇿⟨ sym $ Any⇿ {P = P₁} ⟩
+ (∃ λ x → x ∈ xs₁ × P₁ x) ≈⟨ Σ.cong Inv.id (xs₁≈xs₂ ⟨ ×⊎.*-cong ⟩ P₁⇿P₂ _) ⟩
+ (∃ λ x → x ∈ xs₂ × P₂ x) ⇿⟨ Any⇿ {P = P₂} ⟩
Any P₂ xs₂ ∎
------------------------------------------------------------------------
@@ -124,28 +134,28 @@ Any-cong {P₁ = P₁} {P₂} {xs₁} {xs₂} P₁⇿P₂ xs₁≈xs₂ =
-- Nested occurrences of Any can sometimes be swapped. See also ×⇿.
-swap : ∀ {A B : Set} {P : A → B → Set} {xs ys} →
+swap : ∀ {ℓ} {A B : Set ℓ} {P : A → B → Set ℓ} {xs ys} →
Any (λ x → Any (P x) ys) xs ⇿ Any (λ y → Any (flip P y) xs) ys
-swap {P = P} {xs} {ys} =
- Any (λ x → Any (P x) ys) xs ⇿⟨ sym Any⇿ ⟩
- (∃ λ x → x ∈ xs × Any (P x) ys) ⇿⟨ sym $ Σ.cong (Inv.id ⟨ ×⊎.*-cong ⟩ Any⇿) ⟩
- (∃ λ x → x ∈ xs × ∃ λ y → y ∈ ys × P x y) ⇿⟨ Σ.cong (∃∃⇿∃∃ _) ⟩
- (∃₂ λ x y → x ∈ xs × y ∈ ys × P x y) ⇿⟨ ∃∃⇿∃∃ _ ⟩
- (∃₂ λ y x → x ∈ xs × y ∈ ys × P x y) ⇿⟨ Σ.cong (λ {y} → Σ.cong (λ {x} →
+swap {ℓ} {P = P} {xs} {ys} =
+ Any (λ x → Any (P x) ys) xs ⇿⟨ sym $ Any⇿ {a = ℓ} {p = ℓ} ⟩
+ (∃ λ x → x ∈ xs × Any (P x) ys) ⇿⟨ sym $ Σ.cong Inv.id (λ {x} → (x ∈ xs ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any⇿ {a = ℓ} {p = ℓ}) ⟩
+ (∃ λ x → x ∈ xs × ∃ λ y → y ∈ ys × P x y) ⇿⟨ Σ.cong {a₁ = ℓ} Inv.id (∃∃⇿∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
+ (∃₂ λ x y → x ∈ xs × y ∈ ys × P x y) ⇿⟨ ∃∃⇿∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _ ⟩
+ (∃₂ λ y x → x ∈ xs × y ∈ ys × P x y) ⇿⟨ Σ.cong Inv.id (λ {y} → Σ.cong Inv.id (λ {x} →
(x ∈ xs × y ∈ ys × P x y) ⇿⟨ sym $ ×⊎.*-assoc _ _ _ ⟩
- ((x ∈ xs × y ∈ ys) × P x y) ⇿⟨ ×⊎.*-comm _ _ ⟨ ×⊎.*-cong ⟩ Inv.id ⟩
+ ((x ∈ xs × y ∈ ys) × P x y) ⇿⟨ ×⊎.*-comm (x ∈ xs) (y ∈ ys) ⟨ ×⊎.*-cong ⟩ (P x y ∎) ⟩
((y ∈ ys × x ∈ xs) × P x y) ⇿⟨ ×⊎.*-assoc _ _ _ ⟩
(y ∈ ys × x ∈ xs × P x y) ∎)) ⟩
- (∃₂ λ y x → y ∈ ys × x ∈ xs × P x y) ⇿⟨ Σ.cong (∃∃⇿∃∃ _) ⟩
- (∃ λ y → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ⇿⟨ Σ.cong (Inv.id ⟨ ×⊎.*-cong ⟩ Any⇿) ⟩
- (∃ λ y → y ∈ ys × Any (flip P y) xs) ⇿⟨ Any⇿ ⟩
+ (∃₂ λ y x → y ∈ ys × x ∈ xs × P x y) ⇿⟨ Σ.cong {a₁ = ℓ} Inv.id (∃∃⇿∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
+ (∃ λ y → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ⇿⟨ Σ.cong Inv.id (λ {y} → (y ∈ ys ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any⇿ {a = ℓ} {p = ℓ}) ⟩
+ (∃ λ y → y ∈ ys × Any (flip P y) xs) ⇿⟨ Any⇿ {a = ℓ} {p = ℓ} ⟩
Any (λ y → Any (flip P y) xs) ys ∎
------------------------------------------------------------------------
-- Lemmas relating Any to ⊥
-⊥⇿Any⊥ : {A : Set} {xs : List A} → ⊥ ⇿ Any (const ⊥) xs
-⊥⇿Any⊥ {A} = record
+⊥⇿Any⊥ : ∀ {a} {A : Set a} {xs : List A} → ⊥ ⇿ Any (const ⊥) xs
+⊥⇿Any⊥ {A = A} = record
{ to = P.→-to-⟶ (λ ())
; from = P.→-to-⟶ (λ p → from p)
; inverse-of = record
@@ -154,11 +164,11 @@ swap {P = P} {xs} {ys} =
}
}
where
- from : {xs : List A} → Any (const ⊥) xs → {B : Set} → B
+ from : {xs : List A} → Any (const ⊥) xs → ∀ {b} {B : Set b} → B
from (here ())
from (there p) = from p
-⊥⇿Any[] : {A : Set} {P : A → Set} → ⊥ ⇿ Any P []
+⊥⇿Any[] : ∀ {a} {A : Set a} {P : A → Set} → ⊥ ⇿ Any P []
⊥⇿Any[] = record
{ to = P.→-to-⟶ (λ ())
; from = P.→-to-⟶ (λ ())
@@ -173,7 +183,7 @@ swap {P = P} {xs} {ys} =
-- Sums commute with Any (for a fixed list).
-⊎⇿ : ∀ {A : Set} {P Q : A → Set} {xs} →
+⊎⇿ : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} {xs} →
(Any P xs ⊎ Any Q xs) ⇿ Any (λ x → P x ⊎ Q x) xs
⊎⇿ {P = P} {Q} = record
{ to = P.→-to-⟶ to
@@ -208,7 +218,8 @@ swap {P = P} {xs} {ys} =
-- Products "commute" with Any.
-×⇿ : ∀ {A B} {P : A → Set} {Q : B → Set} {xs} {ys} →
+×⇿ : {A B : Set} {P : A → Set} {Q : B → Set}
+ {xs : List A} {ys : List B} →
(Any P xs × Any Q ys) ⇿ Any (λ x → Any (λ y → P x × Q y) ys) xs
×⇿ {P = P} {Q} {xs} {ys} = record
{ to = P.→-to-⟶ to
@@ -223,7 +234,7 @@ swap {P = P} {xs} {ys} =
to (p , q) = Any.map (λ p → Any.map (λ q → (p , q)) q) p
from : Any (λ x → Any (λ y → P x × Q y) ys) xs → Any P xs × Any Q ys
- from pq with Prod.map id (Prod.map id find) $ find pq
+ from pq with Prod.map id (Prod.map id find) (find pq)
... | (x , x∈xs , y , y∈ys , p , q) = (lose x∈xs p , lose y∈ys q)
from∘to : ∀ pq → from (to pq) ≡ pq
@@ -265,35 +276,40 @@ swap {P = P} {xs} {ys} =
private
- map⁺ : ∀ {A B} {P : B → Set} {f : A → B} {xs} →
+ map⁺ : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
+ {f : A → B} {xs} →
Any (P ∘ f) xs → Any P (List.map f xs)
map⁺ (here p) = here p
map⁺ (there p) = there $ map⁺ p
- map⁻ : ∀ {A B} {P : B → Set} {f : A → B} {xs} →
+ map⁻ : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
+ {f : A → B} {xs} →
Any P (List.map f xs) → Any (P ∘ f) xs
map⁻ {xs = []} ()
map⁻ {xs = x ∷ xs} (here p) = here p
map⁻ {xs = x ∷ xs} (there p) = there $ map⁻ p
- map⁺∘map⁻ : ∀ {A B : Set} {P : B → Set} {f : A → B} {xs} →
+ map⁺∘map⁻ : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
+ {f : A → B} {xs} →
(p : Any P (List.map f xs)) →
map⁺ (map⁻ p) ≡ p
map⁺∘map⁻ {xs = []} ()
map⁺∘map⁻ {xs = x ∷ xs} (here p) = refl
map⁺∘map⁻ {xs = x ∷ xs} (there p) = P.cong there (map⁺∘map⁻ p)
- map⁻∘map⁺ : ∀ {A B} (P : B → Set) {f : A → B} {xs} →
+ map⁻∘map⁺ : ∀ {a b p} {A : Set a} {B : Set b} (P : B → Set p)
+ {f : A → B} {xs} →
(p : Any (P ∘ f) xs) →
map⁻ {P = P} (map⁺ p) ≡ p
map⁻∘map⁺ P (here p) = refl
map⁻∘map⁺ P (there p) = P.cong there (map⁻∘map⁺ P p)
-map⇿ : ∀ {A B} {P : B → Set} {f : A → B} {xs} →
+map⇿ : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
+ {f : A → B} {xs} →
Any (P ∘ f) xs ⇿ Any P (List.map f xs)
-map⇿ {P = P} = record
- { to = P.→-to-⟶ $ map⁺ {P = P}
- ; from = P.→-to-⟶ map⁻
+map⇿ {P = P} {f = f} = record
+ { to = P.→-to-⟶ $ map⁺ {P = P} {f = f}
+ ; from = P.→-to-⟶ $ map⁻ {P = P} {f = f}
; inverse-of = record
{ left-inverse-of = map⁻∘map⁺ P
; right-inverse-of = map⁺∘map⁻
@@ -304,23 +320,23 @@ map⇿ {P = P} = record
private
- ++⁺ˡ : ∀ {A} {P : A → Set} {xs ys} →
+ ++⁺ˡ : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
Any P xs → Any P (xs ++ ys)
++⁺ˡ (here p) = here p
++⁺ˡ (there p) = there (++⁺ˡ p)
- ++⁺ʳ : ∀ {A} {P : A → Set} xs {ys} →
+ ++⁺ʳ : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys} →
Any P ys → Any P (xs ++ ys)
++⁺ʳ [] p = p
++⁺ʳ (x ∷ xs) p = there (++⁺ʳ xs p)
- ++⁻ : ∀ {A} {P : A → Set} xs {ys} →
+ ++⁻ : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys} →
Any P (xs ++ ys) → Any P xs ⊎ Any P ys
++⁻ [] p = inj₂ p
++⁻ (x ∷ xs) (here p) = inj₁ (here p)
++⁻ (x ∷ xs) (there p) = Sum.map there id (++⁻ xs p)
- ++⁺∘++⁻ : ∀ {A} {P : A → Set} xs {ys}
+ ++⁺∘++⁻ : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys}
(p : Any P (xs ++ ys)) →
[ ++⁺ˡ , ++⁺ʳ xs ]′ (++⁻ xs p) ≡ p
++⁺∘++⁻ [] p = refl
@@ -329,7 +345,8 @@ private
++⁺∘++⁻ (x ∷ xs) (there p) | inj₁ p′ | ih = P.cong there ih
++⁺∘++⁻ (x ∷ xs) (there p) | inj₂ p′ | ih = P.cong there ih
- ++⁻∘++⁺ : ∀ {A} {P : A → Set} xs {ys} (p : Any P xs ⊎ Any P ys) →
+ ++⁻∘++⁺ : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys}
+ (p : Any P xs ⊎ Any P ys) →
++⁻ xs ([ ++⁺ˡ , ++⁺ʳ xs ]′ p) ≡ p
++⁻∘++⁺ [] (inj₁ ())
++⁻∘++⁺ [] (inj₂ p) = refl
@@ -337,11 +354,11 @@ private
++⁻∘++⁺ (x ∷ xs) {ys} (inj₁ (there p)) rewrite ++⁻∘++⁺ xs {ys} (inj₁ p) = refl
++⁻∘++⁺ (x ∷ xs) (inj₂ p) rewrite ++⁻∘++⁺ xs (inj₂ p) = refl
-++⇿ : ∀ {A} {P : A → Set} {xs ys} →
+++⇿ : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
(Any P xs ⊎ Any P ys) ⇿ Any P (xs ++ ys)
-++⇿ {xs = xs} = record
- { to = P.→-to-⟶ [ ++⁺ˡ , ++⁺ʳ xs ]′
- ; from = P.→-to-⟶ $ ++⁻ xs
+++⇿ {P = P} {xs = xs} = record
+ { to = P.→-to-⟶ [ ++⁺ˡ {P = P}, ++⁺ʳ {P = P} xs ]′
+ ; from = P.→-to-⟶ $ ++⁻ {P = P} xs
; inverse-of = record
{ left-inverse-of = ++⁻∘++⁺ xs
; right-inverse-of = ++⁺∘++⁻ xs
@@ -352,29 +369,30 @@ private
private
- return⁺ : ∀ {A} {P : A → Set} {x} →
+ return⁺ : ∀ {a p} {A : Set a} {P : A → Set p} {x} →
P x → Any P (return x)
return⁺ = here
- return⁻ : ∀ {A} {P : A → Set} {x} →
+ return⁻ : ∀ {a p} {A : Set a} {P : A → Set p} {x} →
Any P (return x) → P x
return⁻ (here p) = p
return⁻ (there ())
- return⁺∘return⁻ : ∀ {A} {P : A → Set} {x} (p : Any P (return x)) →
+ return⁺∘return⁻ : ∀ {a p} {A : Set a} {P : A → Set p} {x}
+ (p : Any P (return x)) →
return⁺ (return⁻ p) ≡ p
return⁺∘return⁻ (here p) = refl
return⁺∘return⁻ (there ())
- return⁻∘return⁺ : ∀ {A} (P : A → Set) {x} (p : P x) →
+ return⁻∘return⁺ : ∀ {a p} {A : Set a} (P : A → Set p) {x} (p : P x) →
return⁻ {P = P} (return⁺ p) ≡ p
return⁻∘return⁺ P p = refl
-return⇿ : ∀ {A} {P : A → Set} {x} →
+return⇿ : ∀ {a p} {A : Set a} {P : A → Set p} {x} →
P x ⇿ Any P (return x)
return⇿ {P = P} = record
- { to = P.→-to-⟶ return⁺
- ; from = P.→-to-⟶ return⁻
+ { to = P.→-to-⟶ $ return⁺ {P = P}
+ ; from = P.→-to-⟶ $ return⁻ {P = P}
; inverse-of = record
{ left-inverse-of = return⁻∘return⁺ P
; right-inverse-of = return⁺∘return⁻
@@ -385,32 +403,32 @@ return⇿ {P = P} = record
private
- concat⁺ : ∀ {A} {P : A → Set} {xss} →
+ concat⁺ : ∀ {a p} {A : Set a} {P : A → Set p} {xss} →
Any (Any P) xss → Any P (concat xss)
concat⁺ (here p) = ++⁺ˡ p
concat⁺ (there {x = xs} p) = ++⁺ʳ xs (concat⁺ p)
- concat⁻ : ∀ {A} {P : A → Set} xss →
+ concat⁻ : ∀ {a p} {A : Set a} {P : A → Set p} xss →
Any P (concat xss) → Any (Any P) xss
concat⁻ [] ()
concat⁻ ([] ∷ xss) p = there $ concat⁻ xss p
- concat⁻ ((x ∷ xs) ∷ xss) (here p) = here (here p)
+ concat⁻ ((x ∷ xs) ∷ xss) (here p) = here (here p)
concat⁻ ((x ∷ xs) ∷ xss) (there p)
with concat⁻ (xs ∷ xss) p
... | here p′ = here (there p′)
... | there p′ = there p′
- concat⁻∘++⁺ˡ : ∀ {A} {P : A → Set} {xs} xss (p : Any P xs) →
+ concat⁻∘++⁺ˡ : ∀ {a p} {A : Set a} {P : A → Set p} {xs} xss (p : Any P xs) →
concat⁻ (xs ∷ xss) (++⁺ˡ p) ≡ here p
concat⁻∘++⁺ˡ xss (here p) = refl
concat⁻∘++⁺ˡ xss (there p) rewrite concat⁻∘++⁺ˡ xss p = refl
- concat⁻∘++⁺ʳ : ∀ {A} {P : A → Set} xs xss (p : Any P (concat xss)) →
+ concat⁻∘++⁺ʳ : ∀ {a p} {A : Set a} {P : A → Set p} xs xss (p : Any P (concat xss)) →
concat⁻ (xs ∷ xss) (++⁺ʳ xs p) ≡ there (concat⁻ xss p)
concat⁻∘++⁺ʳ [] xss p = refl
concat⁻∘++⁺ʳ (x ∷ xs) xss p rewrite concat⁻∘++⁺ʳ xs xss p = refl
- concat⁺∘concat⁻ : ∀ {A} {P : A → Set} xss (p : Any P (concat xss)) →
+ concat⁺∘concat⁻ : ∀ {a p} {A : Set a} {P : A → Set p} xss (p : Any P (concat xss)) →
concat⁺ (concat⁻ xss p) ≡ p
concat⁺∘concat⁻ [] ()
concat⁺∘concat⁻ ([] ∷ xss) p = concat⁺∘concat⁻ xss p
@@ -420,18 +438,18 @@ private
concat⁺∘concat⁻ ((x ∷ xs) ∷ xss) (there .(++⁺ˡ p′)) | here p′ | refl = refl
concat⁺∘concat⁻ ((x ∷ xs) ∷ xss) (there .(++⁺ʳ xs (concat⁺ p′))) | there p′ | refl = refl
- concat⁻∘concat⁺ : ∀ {A} {P : A → Set} {xss} (p : Any (Any P) xss) →
+ concat⁻∘concat⁺ : ∀ {a p} {A : Set a} {P : A → Set p} {xss} (p : Any (Any P) xss) →
concat⁻ xss (concat⁺ p) ≡ p
concat⁻∘concat⁺ (here p) = concat⁻∘++⁺ˡ _ p
concat⁻∘concat⁺ (there {x = xs} {xs = xss} p)
rewrite concat⁻∘++⁺ʳ xs xss (concat⁺ p) =
P.cong there $ concat⁻∘concat⁺ p
-concat⇿ : ∀ {A} {P : A → Set} {xss} →
+concat⇿ : ∀ {a p} {A : Set a} {P : A → Set p} {xss} →
Any (Any P) xss ⇿ Any P (concat xss)
-concat⇿ {xss = xss} = record
- { to = P.→-to-⟶ concat⁺
- ; from = P.→-to-⟶ $ concat⁻ xss
+concat⇿ {P = P} {xss = xss} = record
+ { to = P.→-to-⟶ $ concat⁺ {P = P}
+ ; from = P.→-to-⟶ $ concat⁻ {P = P} xss
; inverse-of = record
{ left-inverse-of = concat⁻∘concat⁺
; right-inverse-of = concat⁺∘concat⁻ xss
@@ -440,41 +458,46 @@ concat⇿ {xss = xss} = record
-- _>>=_.
->>=⇿ : ∀ {A B P xs} {f : A → List B} →
+>>=⇿ : ∀ {ℓ p} {A B : Set ℓ} {P : B → Set p} {xs} {f : A → List B} →
Any (Any P ∘ f) xs ⇿ Any P (xs >>= f)
>>=⇿ {P = P} {xs} {f} =
- Any (Any P ∘ f) xs ⇿⟨ map⇿ ⟩
- Any (Any P) (List.map f xs) ⇿⟨ concat⇿ ⟩
+ Any (Any P ∘ f) xs ⇿⟨ map⇿ {P = Any P} {f = f} ⟩
+ Any (Any P) (List.map f xs) ⇿⟨ concat⇿ {P = P} ⟩
Any P (xs >>= f) ∎
-- _⊛_.
-⊛⇿ : ∀ {A B P} {fs : List (A → B)} {xs} →
+⊛⇿ : ∀ {ℓ} {A B : Set ℓ} {P : B → Set ℓ}
+ {fs : List (A → B)} {xs : List A} →
Any (λ f → Any (P ∘ f) xs) fs ⇿ Any P (fs ⊛ xs)
-⊛⇿ {P = P} {fs} {xs} =
- Any (λ f → Any (P ∘ f) xs) fs ⇿⟨ Any-cong (λ _ → Any-cong (λ _ → return⇿) (_ ∎)) (_ ∎) ⟩
- Any (λ f → Any (Any P ∘ return ∘ f) xs) fs ⇿⟨ Any-cong (λ _ → >>=⇿) (_ ∎) ⟩
- Any (λ f → Any P (xs >>= return ∘ f)) fs ⇿⟨ >>=⇿ ⟩
+⊛⇿ {ℓ} {P = P} {fs} {xs} =
+ Any (λ f → Any (P ∘ f) xs) fs ⇿⟨ Any-cong (λ _ → Any-cong (λ _ → return⇿ {a = ℓ} {p = ℓ}) (_ ∎)) (_ ∎) ⟩
+ Any (λ f → Any (Any P ∘ return ∘ f) xs) fs ⇿⟨ Any-cong (λ _ → >>=⇿ {ℓ = ℓ} {p = ℓ}) (_ ∎) ⟩
+ Any (λ f → Any P (xs >>= return ∘ f)) fs ⇿⟨ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟩
Any P (fs ⊛ xs) ∎
-- An alternative introduction rule for _⊛_.
-⊛⁺′ : ∀ {A B P Q} {fs : List (A → B)} {xs} →
+⊛⁺′ : ∀ {ℓ} {A B : Set ℓ} {P : A → Set ℓ} {Q : B → Set ℓ}
+ {fs : List (A → B)} {xs} →
Any (P ⟨→⟩ Q) fs → Any P xs → Any Q (fs ⊛ xs)
-⊛⁺′ pq p =
- Inverse.to ⊛⇿ ⟨$⟩ Any.map (λ pq → Any.map (λ {x} → pq {x}) p) pq
+⊛⁺′ {ℓ} pq p =
+ Inverse.to (⊛⇿ {ℓ = ℓ}) ⟨$⟩
+ Any.map (λ pq → Any.map (λ {x} → pq {x}) p) pq
-- _⊗_.
-⊗⇿ : ∀ {A B P} {xs : List A} {ys : List B} →
+⊗⇿ : ∀ {ℓ} {A B : Set ℓ} {P : A × B → Set ℓ}
+ {xs : List A} {ys : List B} →
Any (λ x → Any (λ y → P (x , y)) ys) xs ⇿ Any P (xs ⊗ ys)
-⊗⇿ {P = P} {xs} {ys} =
- Any (λ x → Any (λ y → P (x , y)) ys) xs ⇿⟨ return⇿ ⟩
+⊗⇿ {ℓ} {P = P} {xs} {ys} =
+ Any (λ x → Any (λ y → P (x , y)) ys) xs ⇿⟨ return⇿ {a = ℓ} {p = ℓ} ⟩
Any (λ _,_ → Any (λ x → Any (λ y → P (x , y)) ys) xs) (return _,_) ⇿⟨ ⊛⇿ ⟩
Any (λ x, → Any (P ∘ x,) ys) (_,_ <$> xs) ⇿⟨ ⊛⇿ ⟩
Any P (xs ⊗ ys) ∎
-⊗⇿′ : ∀ {A B P Q} {xs : List A} {ys : List B} →
+⊗⇿′ : {A B : Set} {P : A → Set} {Q : B → Set}
+ {xs : List A} {ys : List B} →
(Any P xs × Any Q ys) ⇿ Any (P ⟨×⟩ Q) (xs ⊗ ys)
⊗⇿′ {P = P} {Q} {xs} {ys} =
(Any P xs × Any Q ys) ⇿⟨ ×⇿ ⟩
@@ -484,9 +507,10 @@ concat⇿ {xss = xss} = record
-- map-with-∈.
map-with-∈⇿ :
- ∀ {A B : Set} {P : B → Set} {xs : List A} {f : ∀ {x} → x ∈ xs → B} →
+ ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p} {xs : List A}
+ {f : ∀ {x} → x ∈ xs → B} →
(∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)) ⇿ Any P (map-with-∈ xs f)
-map-with-∈⇿ {A} {B} {P} = record
+map-with-∈⇿ {A = A} {B} {P} = record
{ to = P.→-to-⟶ (map-with-∈⁺ _)
; from = P.→-to-⟶ (map-with-∈⁻ _ _)
; inverse-of = record
@@ -534,14 +558,14 @@ map-with-∈⇿ {A} {B} {P} = record
private
- any⁺ : ∀ {A} (p : A → Bool) {xs} →
+ any⁺ : ∀ {a} {A : Set a} (p : A → Bool) {xs} →
Any (T ∘ p) xs → T (any p xs)
any⁺ p (here px) = Equivalent.from T-∨ ⟨$⟩ inj₁ px
any⁺ p (there {x = x} pxs) with p x
... | true = _
... | false = any⁺ p pxs
- any⁻ : ∀ {A} (p : A → Bool) xs →
+ any⁻ : ∀ {a} {A : Set a} (p : A → Bool) xs →
T (any p xs) → Any (T ∘ p) xs
any⁻ p [] ()
any⁻ p (x ∷ xs) px∷xs with inspect (p x)
@@ -550,7 +574,7 @@ private
any⁻ p (x ∷ xs) pxs | false with-≡ refl | .false =
there (any⁻ p xs pxs)
-any⇔ : ∀ {A} {p : A → Bool} {xs} →
+any⇔ : ∀ {a} {A : Set a} {p : A → Bool} {xs} →
Any (T ∘ p) xs ⇔ T (any p xs)
any⇔ = equivalent (any⁺ _) (any⁻ _ _)
@@ -559,14 +583,15 @@ any⇔ = equivalent (any⁺ _) (any⁻ _ _)
private
- ++-comm : ∀ {A} {P : A → Set} xs ys →
+ ++-comm : ∀ {a p} {A : Set a} {P : A → Set p} xs ys →
Any P (xs ++ ys) → Any P (ys ++ xs)
++-comm xs ys = [ ++⁺ʳ ys , ++⁺ˡ ]′ ∘ ++⁻ xs
- ++-comm∘++-comm : ∀ {A} {P : A → Set} xs {ys} (p : Any P (xs ++ ys)) →
+ ++-comm∘++-comm : ∀ {a p} {A : Set a} {P : A → Set p}
+ xs {ys} (p : Any P (xs ++ ys)) →
++-comm ys xs (++-comm xs ys p) ≡ p
++-comm∘++-comm [] {ys} p
- rewrite ++⁻∘++⁺ ys {ys = []} (inj₁ p) = P.refl
+ rewrite ++⁻∘++⁺ ys {ys = []} (inj₁ p) = P.refl
++-comm∘++-comm {P = P} (x ∷ xs) {ys} (here p)
rewrite ++⁻∘++⁺ {P = P} ys {ys = x ∷ xs} (inj₂ (here p)) = P.refl
++-comm∘++-comm (x ∷ xs) (there p) with ++⁻ xs p | ++-comm∘++-comm xs p
@@ -579,11 +604,11 @@ private
rewrite ++⁻∘++⁺ ys {ys = xs} (inj₁ p)
| ++⁻∘++⁺ ys {ys = x ∷ xs} (inj₁ p) = P.refl
-++⇿++ : ∀ {A} {P : A → Set} xs ys →
+++⇿++ : ∀ {a p} {A : Set a} {P : A → Set p} xs ys →
Any P (xs ++ ys) ⇿ Any P (ys ++ xs)
-++⇿++ xs ys = record
- { to = P.→-to-⟶ $ ++-comm xs ys
- ; from = P.→-to-⟶ $ ++-comm ys xs
+++⇿++ {P = P} xs ys = record
+ { to = P.→-to-⟶ $ ++-comm {P = P} xs ys
+ ; from = P.→-to-⟶ $ ++-comm {P = P} ys xs
; inverse-of = record
{ left-inverse-of = ++-comm∘++-comm xs
; right-inverse-of = ++-comm∘++-comm ys
diff --git a/src/Data/List/NonEmpty.agda b/src/Data/List/NonEmpty.agda
index 4b8fdf5..d21827c 100644
--- a/src/Data/List/NonEmpty.agda
+++ b/src/Data/List/NonEmpty.agda
@@ -2,6 +2,8 @@
-- Non-empty lists
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Data.List.NonEmpty where
open import Data.Product hiding (map)
@@ -14,100 +16,102 @@ open import Relation.Binary.PropositionalEquality
infixr 5 _∷_ _∷ʳ_ _⁺++⁺_ _++⁺_ _⁺++_
-data List⁺ (A : Set) : Set where
+data List⁺ {a} (A : Set a) : Set a where
[_] : (x : A) → List⁺ A
_∷_ : (x : A) (xs : List⁺ A) → List⁺ A
-length_-1 : ∀ {A} → List⁺ A → ℕ
+length_-1 : ∀ {a} {A : Set a} → List⁺ A → ℕ
length [ x ] -1 = 0
length x ∷ xs -1 = 1 + length xs -1
------------------------------------------------------------------------
-- Conversion
-fromVec : ∀ {n A} → Vec A (suc n) → List⁺ A
+fromVec : ∀ {n a} {A : Set a} → Vec A (suc n) → List⁺ A
fromVec {zero} (x ∷ _) = [ x ]
fromVec {suc n} (x ∷ xs) = x ∷ fromVec xs
-toVec : ∀ {A} (xs : List⁺ A) → Vec A (suc (length xs -1))
+toVec : ∀ {a} {A : Set a} (xs : List⁺ A) → Vec A (suc (length xs -1))
toVec [ x ] = Vec.[_] x
toVec (x ∷ xs) = x ∷ toVec xs
-lift : ∀ {A B} →
+lift : ∀ {a b} {A : Set a} {B : Set b} →
(∀ {m} → Vec A (suc m) → ∃ λ n → Vec B (suc n)) →
List⁺ A → List⁺ B
lift f xs = fromVec (proj₂ (f (toVec xs)))
-fromList : ∀ {A} → A → List A → List⁺ A
+fromList : ∀ {a} {A : Set a} → A → List A → List⁺ A
fromList x xs = fromVec (Vec.fromList (x ∷ xs))
-toList : ∀ {A} → List⁺ A → List A
-toList = Vec.toList ∘ toVec
+toList : ∀ {a} {A : Set a} → List⁺ A → List A
+toList {a} = Vec.toList {a} ∘ toVec
------------------------------------------------------------------------
-- Other operations
-head : ∀ {A} → List⁺ A → A
-head = Vec.head ∘ toVec
+head : ∀ {a} {A : Set a} → List⁺ A → A
+head {a} = Vec.head {a} ∘ toVec
-tail : ∀ {A} → List⁺ A → List A
-tail = Vec.toList ∘ Vec.tail ∘ toVec
+tail : ∀ {a} {A : Set a} → List⁺ A → List A
+tail {a} = Vec.toList {a} ∘ Vec.tail {a} ∘ toVec
-map : ∀ {A B} → (A → B) → List⁺ A → List⁺ B
+map : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → List⁺ A → List⁺ B
map f = lift (λ xs → (, Vec.map f xs))
-- Right fold. Note that s is only applied to the last element (see
-- the examples below).
-foldr : {A B : Set} → (A → B → B) → (A → B) → List⁺ A → B
+foldr : ∀ {a b} {A : Set a} {B : Set b} →
+ (A → B → B) → (A → B) → List⁺ A → B
foldr c s [ x ] = s x
foldr c s (x ∷ xs) = c x (foldr c s xs)
-- Left fold. Note that s is only applied to the first element (see
-- the examples below).
-foldl : {A B : Set} → (B → A → B) → (A → B) → List⁺ A → B
+foldl : ∀ {a b} {A : Set a} {B : Set b} →
+ (B → A → B) → (A → B) → List⁺ A → B
foldl c s [ x ] = s x
foldl c s (x ∷ xs) = foldl c (c (s x)) xs
-- Append (several variants).
-_⁺++⁺_ : ∀ {A} → List⁺ A → List⁺ A → List⁺ A
+_⁺++⁺_ : ∀ {a} {A : Set a} → List⁺ A → List⁺ A → List⁺ A
xs ⁺++⁺ ys = foldr _∷_ (λ x → x ∷ ys) xs
-_⁺++_ : ∀ {A} → List⁺ A → List A → List⁺ A
+_⁺++_ : ∀ {a} {A : Set a} → List⁺ A → List A → List⁺ A
xs ⁺++ ys = foldr _∷_ (λ x → fromList x ys) xs
-_++⁺_ : ∀ {A} → List A → List⁺ A → List⁺ A
+_++⁺_ : ∀ {a} {A : Set a} → List A → List⁺ A → List⁺ A
xs ++⁺ ys = List.foldr _∷_ ys xs
-concat : ∀ {A} → List⁺ (List⁺ A) → List⁺ A
+concat : ∀ {a} {A : Set a} → List⁺ (List⁺ A) → List⁺ A
concat [ xs ] = xs
concat (xs ∷ xss) = xs ⁺++⁺ concat xss
-monad : RawMonad List⁺
+monad : ∀ {f} → RawMonad (List⁺ {a = f})
monad = record
{ return = [_]
; _>>=_ = λ xs f → concat (map f xs)
}
-reverse : ∀ {A} → List⁺ A → List⁺ A
+reverse : ∀ {a} {A : Set a} → List⁺ A → List⁺ A
reverse = lift (,_ ∘′ Vec.reverse)
-- Snoc.
-_∷ʳ_ : ∀ {A} → List⁺ A → A → List⁺ A
+_∷ʳ_ : ∀ {a} {A : Set a} → List⁺ A → A → List⁺ A
xs ∷ʳ x = foldr _∷_ (λ y → y ∷ [ x ]) xs
-- A snoc-view of non-empty lists.
infixl 5 _∷ʳ′_
-data SnocView {A} : List⁺ A → Set where
+data SnocView {a} {A : Set a} : List⁺ A → Set a where
[_] : (x : A) → SnocView [ x ]
_∷ʳ′_ : (xs : List⁺ A) (x : A) → SnocView (xs ∷ʳ x)
-snocView : ∀ {A} (xs : List⁺ A) → SnocView xs
+snocView : ∀ {a} {A : Set a} (xs : List⁺ A) → SnocView xs
snocView [ x ] = [ x ]
snocView (x ∷ xs) with snocView xs
snocView (x ∷ .([ y ])) | [ y ] = [ x ] ∷ʳ′ y
@@ -115,7 +119,7 @@ snocView (x ∷ .(ys ∷ʳ y)) | ys ∷ʳ′ y = (x ∷ ys) ∷ʳ′ y
-- The last element in the list.
-last : ∀ {A} → List⁺ A → A
+last : ∀ {a} {A : Set a} → List⁺ A → A
last xs with snocView xs
last .([ y ]) | [ y ] = y
last .(ys ∷ʳ y) | ys ∷ʳ′ y = y
diff --git a/src/Data/List/NonEmpty/Properties.agda b/src/Data/List/NonEmpty/Properties.agda
index 1a8b9ae..78d04cc 100644
--- a/src/Data/List/NonEmpty/Properties.agda
+++ b/src/Data/List/NonEmpty/Properties.agda
@@ -8,41 +8,49 @@ module Data.List.NonEmpty.Properties where
open import Algebra
open import Category.Monad
-open import Function
-open import Data.Product
open import Data.List as List using (List; []; _∷_; _++_)
-open RawMonad List.monad using () renaming (_>>=_ to _⋆>>=_)
-private module LM {a} {A : Set a} = Monoid (List.monoid A)
open import Data.List.NonEmpty as List⁺
-open RawMonad List⁺.monad
+open import Data.Product
+open import Function
open import Relation.Binary.PropositionalEquality
-open ≡-Reasoning
-η : ∀ {A} (xs : List⁺ A) → head xs ∷ tail xs ≡ List⁺.toList xs
+open ≡-Reasoning
+private
+ module LM {a} {A : Set a} = Monoid (List.monoid A)
+ open module LMo {a} =
+ RawMonad {f = a} List.monad
+ using () renaming (_>>=_ to _⋆>>=_)
+ open module L⁺Mo {a} =
+ RawMonad {f = a} List⁺.monad
+
+η : ∀ {a} {A : Set a}
+ (xs : List⁺ A) → head xs ∷ tail xs ≡ List⁺.toList xs
η [ x ] = refl
η (x ∷ xs) = refl
-toList-fromList : ∀ {A} x (xs : List A) →
+toList-fromList : ∀ {a} {A : Set a} x (xs : List A) →
x ∷ xs ≡ List⁺.toList (List⁺.fromList x xs)
toList-fromList x [] = refl
toList-fromList x (y ∷ xs) = cong (_∷_ x) (toList-fromList y xs)
-toList-⁺++ : ∀ {A} (xs : List⁺ A) ys →
+toList-⁺++ : ∀ {a} {A : Set a} (xs : List⁺ A) ys →
List⁺.toList xs ++ ys ≡
List⁺.toList (xs ⁺++ ys)
toList-⁺++ [ x ] ys = toList-fromList x ys
toList-⁺++ (x ∷ xs) ys = cong (_∷_ x) (toList-⁺++ xs ys)
-toList-⁺++⁺ : ∀ {A} (xs ys : List⁺ A) →
+toList-⁺++⁺ : ∀ {a} {A : Set a} (xs ys : List⁺ A) →
List⁺.toList xs ++ List⁺.toList ys ≡
List⁺.toList (xs ⁺++⁺ ys)
toList-⁺++⁺ [ x ] ys = refl
toList-⁺++⁺ (x ∷ xs) ys = cong (_∷_ x) (toList-⁺++⁺ xs ys)
-toList->>= : ∀ {A B} (f : A → List⁺ B) (xs : List⁺ A) →
+toList->>= : ∀ {ℓ} {A B : Set ℓ}
+ (f : A → List⁺ B) (xs : List⁺ A) →
(List⁺.toList xs ⋆>>= List⁺.toList ∘ f) ≡
(List⁺.toList (xs >>= f))
-toList->>= f [ x ] = proj₂ LM.identity _
+toList->>= {ℓ} f [ x ] =
+ proj₂ {a = ℓ} {b = ℓ} LM.identity _
toList->>= f (x ∷ xs) = begin
List⁺.toList (f x) ++ (List⁺.toList xs ⋆>>= List⁺.toList ∘ f) ≡⟨ cong (_++_ (List⁺.toList (f x))) (toList->>= f xs) ⟩
List⁺.toList (f x) ++ List⁺.toList (xs >>= f) ≡⟨ toList-⁺++⁺ (f x) (xs >>= f) ⟩
diff --git a/src/Data/List/Properties.agda b/src/Data/List/Properties.agda
index 724dfb5..d17428c 100644
--- a/src/Data/List/Properties.agda
+++ b/src/Data/List/Properties.agda
@@ -22,8 +22,8 @@ open import Relation.Binary.PropositionalEquality as P
using (_≡_; _≢_; _≗_; refl)
import Relation.Binary.EqReasoning as EqR
-open RawMonadPlus List.monadPlus
private
+ open module LMP {ℓ} = RawMonadPlus (List.monadPlus {ℓ = ℓ})
module LM {a} {A : Set a} = Monoid (List.monoid A)
∷-injective : ∀ {a} {A : Set a} {x y xs ys} →
@@ -48,7 +48,7 @@ right-identity-unique [] refl = refl
right-identity-unique (x ∷ xs) eq =
right-identity-unique xs (proj₂ (∷-injective eq))
-left-identity-unique : ∀ {A : Set} {xs} (ys : List A) →
+left-identity-unique : ∀ {a} {A : Set a} {xs} (ys : List A) →
xs ≡ ys ++ xs → ys ≡ []
left-identity-unique [] _ = refl
left-identity-unique {xs = []} (y ∷ ys) ()
@@ -292,14 +292,49 @@ length-gfilter p (x ∷ xs) with p x
length-gfilter p (x ∷ xs) | just y = s≤s (length-gfilter p xs)
length-gfilter p (x ∷ xs) | nothing = ≤-step (length-gfilter p xs)
+-- Reverse.
+
+unfold-reverse : ∀ {a} {A : Set a} (x : A) (xs : List A) →
+ reverse (x ∷ xs) ≡ reverse xs ∷ʳ x
+unfold-reverse x xs = begin
+ foldl (flip _∷_) [ x ] xs ≡⟨ helper [ x ] xs ⟩
+ reverse xs ∷ʳ x ∎
+ where
+ open P.≡-Reasoning
+
+ helper : ∀ {a} {A : Set a} (xs ys : List A) →
+ foldl (flip _∷_) xs ys ≡ reverse ys ++ xs
+ helper xs [] = P.refl
+ helper xs (y ∷ ys) = begin
+ foldl (flip _∷_) (y ∷ xs) ys ≡⟨ helper (y ∷ xs) ys ⟩
+ reverse ys ++ y ∷ xs ≡⟨ P.sym $ LM.assoc (reverse ys) _ _ ⟩
+ (reverse ys ∷ʳ y) ++ xs ≡⟨ P.sym $ P.cong (λ zs → zs ++ xs) (unfold-reverse y ys) ⟩
+ reverse (y ∷ ys) ++ xs ∎
+
+reverse-++-commute :
+ ∀ {a} {A : Set a} (xs ys : List A) →
+ reverse (xs ++ ys) ≡ reverse ys ++ reverse xs
+reverse-++-commute {a} [] ys = begin
+ reverse ys ≡⟨ P.sym $ proj₂ {a = a} {b = a} LM.identity _ ⟩
+ reverse ys ++ [] ≡⟨ P.refl ⟩
+ reverse ys ++ reverse [] ∎
+ where open P.≡-Reasoning
+reverse-++-commute (x ∷ xs) ys = begin
+ reverse (x ∷ xs ++ ys) ≡⟨ unfold-reverse x (xs ++ ys) ⟩
+ reverse (xs ++ ys) ++ [ x ] ≡⟨ P.cong (λ zs → zs ++ [ x ]) (reverse-++-commute xs ys) ⟩
+ (reverse ys ++ reverse xs) ++ [ x ] ≡⟨ LM.assoc (reverse ys) _ _ ⟩
+ reverse ys ++ (reverse xs ++ [ x ]) ≡⟨ P.sym $ P.cong (λ zs → reverse ys ++ zs) (unfold-reverse x xs) ⟩
+ reverse ys ++ reverse (x ∷ xs) ∎
+ where open P.≡-Reasoning
+
-- The list monad.
module Monad where
- left-zero : {A B : Set} (f : A → List B) → (∅ >>= f) ≡ ∅
+ left-zero : ∀ {ℓ} {A B : Set ℓ} (f : A → List B) → (∅ >>= f) ≡ ∅
left-zero f = refl
- right-zero : {A B : Set} (xs : List A) →
+ right-zero : ∀ {ℓ} {A B : Set ℓ} (xs : List A) →
(xs >>= const ∅) ≡ (∅ ∶ List B)
right-zero [] = refl
right-zero (x ∷ xs) = right-zero xs
@@ -311,7 +346,8 @@ module Monad where
(xs >>= λ x → f x ∣ g x) ≢ ((xs >>= f) ∣ (xs >>= g))
not-left-distributive ()
- right-distributive : {A B : Set} (xs ys : List A) (f : A → List B) →
+ right-distributive : ∀ {ℓ} {A B : Set ℓ}
+ (xs ys : List A) (f : A → List B) →
(xs ∣ ys >>= f) ≡ ((xs >>= f) ∣ (ys >>= f))
right-distributive [] ys f = refl
right-distributive (x ∷ xs) ys f = begin
@@ -320,16 +356,16 @@ module Monad where
(f x ∣ (xs >>= f)) ∣ (ys >>= f) ∎
where open P.≡-Reasoning
- left-identity : {A B : Set} (x : A) (f : A → List B) →
+ left-identity : ∀ {ℓ} {A B : Set ℓ} (x : A) (f : A → List B) →
(return x >>= f) ≡ f x
- left-identity x f = proj₂ LM.identity (f x)
+ left-identity {ℓ} x f = proj₂ (LM.identity {a = ℓ}) (f x)
- right-identity : {A : Set} (xs : List A) →
+ right-identity : ∀ {a} {A : Set a} (xs : List A) →
(xs >>= return) ≡ xs
right-identity [] = refl
right-identity (x ∷ xs) = P.cong (_∷_ x) (right-identity xs)
- associative : {A B C : Set}
+ associative : ∀ {ℓ} {A B C : Set ℓ}
(xs : List A) (f : A → List B) (g : B → List C) →
(xs >>= λ x → f x >>= g) ≡ (xs >>= f >>= g)
associative [] f g = refl
@@ -339,7 +375,7 @@ module Monad where
(f x ∣ (xs >>= f) >>= g) ∎
where open P.≡-Reasoning
- cong : ∀ {A B : Set} {xs₁ xs₂} {f₁ f₂ : A → List B} →
+ cong : ∀ {ℓ} {A B : Set ℓ} {xs₁ xs₂} {f₁ f₂ : A → List B} →
xs₁ ≡ xs₂ → f₁ ≗ f₂ → (xs₁ >>= f₁) ≡ (xs₂ >>= f₂)
cong {xs₁ = xs} refl f₁≗f₂ = P.cong concat (map-cong f₁≗f₂ xs)
@@ -357,12 +393,12 @@ module Applicative where
-- A variant of flip map.
- pam : {A B : Set} → List A → (A → B) → List B
+ pam : ∀ {ℓ} {A B : Set ℓ} → List A → (A → B) → List B
pam xs f = xs >>= return ∘ f
-- ∅ is a left zero for _⊛_.
- left-zero : ∀ {A B} xs → (∅ ∶ List (A → B)) ⊛ xs ≡ ∅
+ left-zero : ∀ {ℓ} {A B : Set ℓ} xs → (∅ ∶ List (A → B)) ⊛ xs ≡ ∅
left-zero xs = begin
∅ ⊛ xs ≡⟨ refl ⟩
(∅ >>= pam xs) ≡⟨ Monad.left-zero (pam xs) ⟩
@@ -370,18 +406,18 @@ module Applicative where
-- ∅ is a right zero for _⊛_.
- right-zero : ∀ {A B} (fs : List (A → B)) → fs ⊛ ∅ ≡ ∅
- right-zero fs = begin
+ right-zero : ∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) → fs ⊛ ∅ ≡ ∅
+ right-zero {ℓ} fs = begin
fs ⊛ ∅ ≡⟨ refl ⟩
(fs >>= pam ∅) ≡⟨ (Monad.cong (refl {x = fs}) λ f →
- Monad.left-zero (return ∘ f)) ⟩
+ Monad.left-zero (return {ℓ = ℓ} ∘ f)) ⟩
(fs >>= λ _ → ∅) ≡⟨ Monad.right-zero fs ⟩
∅ ∎
-- _⊛_ distributes over _∣_ from the right.
right-distributive :
- ∀ {A B} (fs₁ fs₂ : List (A → B)) xs →
+ ∀ {ℓ} {A B : Set ℓ} (fs₁ fs₂ : List (A → B)) xs →
(fs₁ ∣ fs₂) ⊛ xs ≡ (fs₁ ⊛ xs ∣ fs₂ ⊛ xs)
right-distributive fs₁ fs₂ xs = begin
(fs₁ ∣ fs₂) ⊛ xs ≡⟨ refl ⟩
@@ -400,7 +436,7 @@ module Applicative where
-- Applicative functor laws.
- identity : ∀ {A} (xs : List A) → return id ⊛ xs ≡ xs
+ identity : ∀ {a} {A : Set a} (xs : List A) → return id ⊛ xs ≡ xs
identity xs = begin
return id ⊛ xs ≡⟨ refl ⟩
(return id >>= pam xs) ≡⟨ Monad.left-identity id (pam xs) ⟩
@@ -409,7 +445,7 @@ module Applicative where
private
- pam-lemma : {A B C : Set}
+ pam-lemma : ∀ {ℓ} {A B C : Set ℓ}
(xs : List A) (f : A → B) (fs : B → List C) →
(pam xs f >>= fs) ≡ (xs >>= λ x → fs (f x))
pam-lemma xs f fs = begin
@@ -418,13 +454,14 @@ module Applicative where
(xs >>= λ x → fs (f x)) ∎
composition :
- ∀ {A B C} (fs : List (B → C)) (gs : List (A → B)) xs →
+ ∀ {ℓ} {A B C : Set ℓ}
+ (fs : List (B → C)) (gs : List (A → B)) xs →
return _∘′_ ⊛ fs ⊛ gs ⊛ xs ≡ fs ⊛ (gs ⊛ xs)
- composition fs gs xs = begin
+ composition {ℓ} fs gs xs = begin
return _∘′_ ⊛ fs ⊛ gs ⊛ xs ≡⟨ refl ⟩
(return _∘′_ >>= pam fs >>= pam gs >>= pam xs) ≡⟨ Monad.cong (Monad.cong (Monad.left-identity _∘′_ (pam fs))
- (λ _ → refl))
- (λ _ → refl) ⟩
+ (λ f → refl {x = pam gs f}))
+ (λ fg → refl {x = pam xs fg}) ⟩
(pam fs _∘′_ >>= pam gs >>= pam xs) ≡⟨ Monad.cong (pam-lemma fs _∘′_ (pam gs)) (λ _ → refl) ⟩
((fs >>= λ f → pam gs (_∘′_ f)) >>= pam xs) ≡⟨ P.sym $ Monad.associative fs (λ f → pam gs (_∘′_ f)) (pam xs) ⟩
(fs >>= λ f → pam gs (_∘′_ f) >>= pam xs) ≡⟨ (Monad.cong (refl {x = fs}) λ f →
@@ -437,7 +474,7 @@ module Applicative where
(fs >>= pam (gs >>= pam xs)) ≡⟨ refl ⟩
fs ⊛ (gs ⊛ xs) ∎
- homomorphism : ∀ {A B : Set} (f : A → B) x →
+ homomorphism : ∀ {ℓ} {A B : Set ℓ} (f : A → B) x →
return f ⊛ return x ≡ return (f x)
homomorphism f x = begin
return f ⊛ return x ≡⟨ refl ⟩
@@ -445,7 +482,7 @@ module Applicative where
pam (return x) f ≡⟨ Monad.left-identity x (return ∘ f) ⟩
return (f x) ∎
- interchange : ∀ {A B} (fs : List (A → B)) {x} →
+ interchange : ∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) {x} →
fs ⊛ return x ≡ return (λ f → f x) ⊛ fs
interchange fs {x} = begin
fs ⊛ return x ≡⟨ refl ⟩
diff --git a/src/Data/Maybe.agda b/src/Data/Maybe.agda
index 2e7e8fc..5260a4e 100644
--- a/src/Data/Maybe.agda
+++ b/src/Data/Maybe.agda
@@ -27,12 +27,18 @@ maybeToBool : ∀ {a} {A : Set a} → Maybe A → Bool
maybeToBool (just _) = true
maybeToBool nothing = false
--- A non-dependent eliminator.
+-- A dependent eliminator.
-maybe : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → B → Maybe A → B
+maybe : ∀ {a b} {A : Set a} {B : Maybe A → Set b} →
+ ((x : A) → B (just x)) → B nothing → (x : Maybe A) → B x
maybe j n (just x) = j x
maybe j n nothing = n
+-- A non-dependent eliminator.
+
+maybe′ : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → B → Maybe A → B
+maybe′ = maybe
+
------------------------------------------------------------------------
-- Maybe monad
@@ -41,34 +47,35 @@ open import Category.Functor
open import Category.Monad
open import Category.Monad.Identity
-functor : RawFunctor Maybe
+functor : ∀ {f} → RawFunctor (Maybe {a = f})
functor = record
{ _<$>_ = λ f → maybe (just ∘ f) nothing
}
-monadT : ∀ {M} → RawMonad M → RawMonad (λ A → M (Maybe A))
+monadT : ∀ {f} {M : Set f → Set f} →
+ RawMonad M → RawMonad (λ A → M (Maybe A))
monadT M = record
{ return = M.return ∘ just
; _>>=_ = λ m f → M._>>=_ m (maybe f (M.return nothing))
}
where module M = RawMonad M
-monad : RawMonad Maybe
+monad : ∀ {f} → RawMonad (Maybe {a = f})
monad = monadT IdentityMonad
-monadZero : RawMonadZero Maybe
+monadZero : ∀ {f} → RawMonadZero (Maybe {a = f})
monadZero = record
{ monad = monad
; ∅ = nothing
}
-monadPlus : RawMonadPlus Maybe
-monadPlus = record
+monadPlus : ∀ {f} → RawMonadPlus (Maybe {a = f})
+monadPlus {f} = record
{ monadZero = monadZero
; _∣_ = _∣_
}
where
- _∣_ : ∀ {a : Set} → Maybe a → Maybe a → Maybe a
+ _∣_ : {A : Set f} → Maybe A → Maybe A → Maybe A
nothing ∣ y = y
just x ∣ y = just x
diff --git a/src/Data/Nat.agda b/src/Data/Nat.agda
index feacbd7..cfcd1e7 100644
--- a/src/Data/Nat.agda
+++ b/src/Data/Nat.agda
@@ -93,14 +93,14 @@ _+_ : ℕ → ℕ → ℕ
zero + n = n
suc m + n = suc (m + n)
+{-# BUILTIN NATPLUS _+_ #-}
+
-- Argument-swapping addition. Used by Data.Vec._⋎_.
_+⋎_ : ℕ → ℕ → ℕ
zero +⋎ n = n
suc m +⋎ n = suc (n +⋎ m)
-{-# BUILTIN NATPLUS _+_ #-}
-
_∸_ : ℕ → ℕ → ℕ
m ∸ zero = m
zero ∸ suc n = zero
diff --git a/src/Data/Nat/DivMod.agda b/src/Data/Nat/DivMod.agda
index 70a6c01..76d6d9a 100644
--- a/src/Data/Nat/DivMod.agda
+++ b/src/Data/Nat/DivMod.agda
@@ -23,7 +23,9 @@ private
lem₁ : ∀ m k → _
lem₁ m k = cong suc $ begin
m
- ≡⟨ Fin.inject+-lemma m k ⟩
+ ≡⟨ sym $ Fin.to-from m ⟩
+ toℕ (fromℕ m)
+ ≡⟨ Fin.inject+-lemma k (fromℕ m) ⟩
toℕ (Fin.inject+ k (fromℕ m))
≡⟨ solve 1 (λ x → x := x :+ con 0) refl _ ⟩
toℕ (Fin.inject+ k (fromℕ m)) + 0
diff --git a/src/Data/Plus.agda b/src/Data/Plus.agda
new file mode 100644
index 0000000..c2b1b95
--- /dev/null
+++ b/src/Data/Plus.agda
@@ -0,0 +1,84 @@
+------------------------------------------------------------------------
+-- Transitive closures
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Data.Plus where
+
+open import Function
+open import Function.Equivalence as Equiv using (_⇔_)
+open import Level
+open import Relation.Binary
+
+------------------------------------------------------------------------
+-- Transitive closure
+
+infix 4 Plus
+
+syntax Plus R x y = x [ R ]⁺ y
+
+data Plus {a ℓ} {A : Set a} (_∼_ : Rel A ℓ) : Rel A (a ⊔ ℓ) where
+ [_] : ∀ {x y} (x∼y : x ∼ y) → x [ _∼_ ]⁺ y
+ _∼⁺⟨_⟩_ : ∀ x {y z} (x∼⁺y : x [ _∼_ ]⁺ y) (y∼⁺z : y [ _∼_ ]⁺ z) →
+ x [ _∼_ ]⁺ z
+
+-- "Equational" reasoning notation. Example:
+--
+-- lemma =
+-- x ∼⁺⟨ [ lemma₁ ] ⟩
+-- y ∼⁺⟨ lemma₂ ⟩∎
+-- z ∎
+
+finally : ∀ {a ℓ} {A : Set a} {_∼_ : Rel A ℓ} x y →
+ x [ _∼_ ]⁺ y → x [ _∼_ ]⁺ y
+finally _ _ = id
+
+syntax finally x y x∼⁺y = x ∼⁺⟨ x∼⁺y ⟩∎ y ∎
+
+infixr 2 _∼⁺⟨_⟩_
+infix 2 finally
+
+-- Map.
+
+map : ∀ {a a′ ℓ ℓ′} {A : Set a} {A′ : Set a′}
+ {_R_ : Rel A ℓ} {_R′_ : Rel A′ ℓ′} {f : A → A′} →
+ _R_ =[ f ]⇒ _R′_ → Plus _R_ =[ f ]⇒ Plus _R′_
+map R⇒R′ [ xRy ] = [ R⇒R′ xRy ]
+map R⇒R′ (x ∼⁺⟨ xR⁺z ⟩ zR⁺y) =
+ _ ∼⁺⟨ map R⇒R′ xR⁺z ⟩ map R⇒R′ zR⁺y
+
+------------------------------------------------------------------------
+-- Alternative definition of transitive closure
+
+-- A generalisation of Data.List.Nonempty.List⁺.
+
+infixr 5 _∷_ _++_
+infix 4 Plus′
+
+syntax Plus′ R x y = x ⟨ R ⟩⁺ y
+
+data Plus′ {a ℓ} {A : Set a} (_∼_ : Rel A ℓ) : Rel A (a ⊔ ℓ) where
+ [_] : ∀ {x y} (x∼y : x ∼ y) → x ⟨ _∼_ ⟩⁺ y
+ _∷_ : ∀ {x y z} (x∼y : x ∼ y) (y∼⁺z : y ⟨ _∼_ ⟩⁺ z) → x ⟨ _∼_ ⟩⁺ z
+
+-- Transitivity.
+
+_++_ : ∀ {a ℓ} {A : Set a} {_∼_ : Rel A ℓ} {x y z} →
+ x ⟨ _∼_ ⟩⁺ y → y ⟨ _∼_ ⟩⁺ z → x ⟨ _∼_ ⟩⁺ z
+[ x∼y ] ++ y∼⁺z = x∼y ∷ y∼⁺z
+(x∼y ∷ y∼⁺z) ++ z∼⁺u = x∼y ∷ (y∼⁺z ++ z∼⁺u)
+
+-- Plus and Plus′ are equivalent.
+
+equivalent : ∀ {a ℓ} {A : Set a} {_∼_ : Rel A ℓ} {x y} →
+ Plus _∼_ x y ⇔ Plus′ _∼_ x y
+equivalent {_∼_ = _∼_} = Equiv.equivalent complete sound
+ where
+ complete : Plus _∼_ ⇒ Plus′ _∼_
+ complete [ x∼y ] = [ x∼y ]
+ complete (x ∼⁺⟨ x∼⁺y ⟩ y∼⁺z) = complete x∼⁺y ++ complete y∼⁺z
+
+ sound : Plus′ _∼_ ⇒ Plus _∼_
+ sound [ x∼y ] = [ x∼y ]
+ sound (x∼y ∷ y∼⁺z) = _ ∼⁺⟨ [ x∼y ] ⟩ sound y∼⁺z
diff --git a/src/Data/Product/N-ary.agda b/src/Data/Product/N-ary.agda
new file mode 100644
index 0000000..83c4dd7
--- /dev/null
+++ b/src/Data/Product/N-ary.agda
@@ -0,0 +1,66 @@
+------------------------------------------------------------------------
+-- N-ary products
+------------------------------------------------------------------------
+
+-- Vectors (as in Data.Vec) also represent n-ary products, so what is
+-- the point of this module? The n-ary products below are intended to
+-- be used with a fixed n, in which case the nil constructor can be
+-- avoided: pairs are represented as pairs (x , y), not as triples
+-- (x , y , unit).
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Data.Product.N-ary where
+
+open import Data.Nat
+open import Data.Product
+open import Data.Unit
+open import Data.Vec
+open import Function.Inverse
+open import Level
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+
+-- N-ary product.
+
+infix 8 _^_
+
+_^_ : ∀ {ℓ} → Set ℓ → ℕ → Set ℓ
+A ^ 0 = Lift ⊤
+A ^ 1 = A
+A ^ suc (suc n) = A × A ^ suc n
+
+-- Conversions.
+
+⇿Vec : ∀ {a} {A : Set a} n → A ^ n ⇿ Vec A n
+⇿Vec n = record
+ { to = P.→-to-⟶ (toVec n)
+ ; from = P.→-to-⟶ fromVec
+ ; inverse-of = record
+ { left-inverse-of = fromVec∘toVec n
+ ; right-inverse-of = toVec∘fromVec
+ }
+ }
+ where
+ toVec : ∀ {a} {A : Set a} n → A ^ n → Vec A n
+ toVec 0 (lift tt) = []
+ toVec 1 x = [ x ]
+ toVec (suc (suc n)) (x , xs) = x ∷ toVec _ xs
+
+ fromVec : ∀ {a} {A : Set a} {n} → Vec A n → A ^ n
+ fromVec [] = lift tt
+ fromVec (x ∷ []) = x
+ fromVec (x ∷ y ∷ xs) = (x , fromVec (y ∷ xs))
+
+ fromVec∘toVec : ∀ {a} {A : Set a} n (xs : A ^ n) →
+ fromVec (toVec n xs) ≡ xs
+ fromVec∘toVec 0 (lift tt) = P.refl
+ fromVec∘toVec 1 x = P.refl
+ fromVec∘toVec 2 (x , y) = P.refl
+ fromVec∘toVec (suc (suc (suc n))) (x , y , xs) =
+ P.cong (_,_ x) (fromVec∘toVec (suc (suc n)) (y , xs))
+
+ toVec∘fromVec : ∀ {a} {A : Set a} {n} (xs : Vec A n) →
+ toVec n (fromVec xs) ≡ xs
+ toVec∘fromVec [] = P.refl
+ toVec∘fromVec (x ∷ []) = P.refl
+ toVec∘fromVec (x ∷ y ∷ xs) = P.cong (_∷_ x) (toVec∘fromVec (y ∷ xs))
diff --git a/src/Data/ReflexiveClosure.agda b/src/Data/ReflexiveClosure.agda
new file mode 100644
index 0000000..1fabe8f
--- /dev/null
+++ b/src/Data/ReflexiveClosure.agda
@@ -0,0 +1,48 @@
+------------------------------------------------------------------------
+-- Reflexive closures
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Data.ReflexiveClosure where
+
+open import Data.Unit
+open import Level
+open import Relation.Binary
+open import Relation.Binary.Simple
+
+------------------------------------------------------------------------
+-- Reflexive closure
+
+data Refl {a ℓ} {A : Set a} (_∼_ : Rel A ℓ) : Rel A (a ⊔ ℓ) where
+ [_] : ∀ {x y} (x∼y : x ∼ y) → Refl _∼_ x y
+ refl : Reflexive (Refl _∼_)
+
+-- Map.
+
+map : ∀ {a a′ ℓ ℓ′} {A : Set a} {A′ : Set a′}
+ {_R_ : Rel A ℓ} {_R′_ : Rel A′ ℓ′} {f : A → A′} →
+ _R_ =[ f ]⇒ _R′_ → Refl _R_ =[ f ]⇒ Refl _R′_
+map R⇒R′ [ xRy ] = [ R⇒R′ xRy ]
+map R⇒R′ refl = refl
+
+-- The reflexive closure has no effect on reflexive relations.
+
+drop-refl : ∀ {a ℓ} {A : Set a} {_R_ : Rel A ℓ} →
+ Reflexive _R_ → Refl _R_ ⇒ _R_
+drop-refl rfl [ x∼y ] = x∼y
+drop-refl rfl refl = rfl
+
+------------------------------------------------------------------------
+-- Example: Maybe
+
+module Maybe where
+
+ Maybe : ∀ {ℓ} → Set ℓ → Set ℓ
+ Maybe A = Refl (Const A) tt tt
+
+ nothing : ∀ {a} {A : Set a} → Maybe A
+ nothing = refl
+
+ just : ∀ {a} {A : Set a} → A → Maybe A
+ just = [_]
diff --git a/src/Data/Star.agda b/src/Data/Star.agda
index f2d716a..a3b3deb 100644
--- a/src/Data/Star.agda
+++ b/src/Data/Star.agda
@@ -2,6 +2,8 @@
-- The reflexive transitive closures of McBride, Norell and Jansson
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- This module could be placed under Relation.Binary. However, since
-- its primary purpose is to be used for _data_ it has been placed
-- under Data instead.
@@ -16,7 +18,7 @@ infixr 5 _◅_
-- Reflexive transitive closure.
-data Star {I : Set} (T : Rel I zero) : Rel I zero where
+data Star {i t} {I : Set i} (T : Rel I t) : Rel I (i ⊔ t) where
ε : Reflexive (Star T)
_◅_ : ∀ {i j k} (x : T i j) (xs : Star T j k) → Star T i k
-- The type of _◅_ is Trans T (Star T) (Star T); I expanded
@@ -27,7 +29,7 @@ data Star {I : Set} (T : Rel I zero) : Rel I zero where
infixr 5 _◅◅_
-_◅◅_ : ∀ {I} {T : Rel I zero} → Transitive (Star T)
+_◅◅_ : ∀ {i t} {I : Set i} {T : Rel I t} → Transitive (Star T)
ε ◅◅ ys = ys
(x ◅ xs) ◅◅ ys = x ◅ (xs ◅◅ ys)
@@ -38,7 +40,7 @@ _◅◅_ : ∀ {I} {T : Rel I zero} → Transitive (Star T)
infixl 5 _▻_
-_▻_ : ∀ {I} {T : Rel I zero} {i j k} →
+_▻_ : ∀ {i t} {I : Set i} {T : Rel I t} {i j k} →
Star T j k → T i j → Star T i k
_▻_ = flip _◅_
@@ -46,77 +48,82 @@ _▻_ = flip _◅_
infixr 5 _▻▻_
-_▻▻_ : ∀ {I} {T : Rel I zero} {i j k} →
+_▻▻_ : ∀ {i t} {I : Set i} {T : Rel I t} {i j k} →
Star T j k → Star T i j → Star T i k
_▻▻_ = flip _◅◅_
-- A generalised variant of map which allows the index type to change.
-gmap : ∀ {I} {T : Rel I zero} {J} {U : Rel J zero} →
+gmap : ∀ {i j t u} {I : Set i} {T : Rel I t} {J : Set j} {U : Rel J u} →
(f : I → J) → T =[ f ]⇒ U → Star T =[ f ]⇒ Star U
gmap f g ε = ε
gmap f g (x ◅ xs) = g x ◅ gmap f g xs
-map : ∀ {I} {T U : Rel I zero} → T ⇒ U → Star T ⇒ Star U
+map : ∀ {i t u} {I : Set i} {T : Rel I t} {U : Rel I u} →
+ T ⇒ U → Star T ⇒ Star U
map = gmap id
-- A generalised variant of fold.
-gfold : ∀ {I J T} (f : I → J) P →
+gfold : ∀ {i j t p} {I : Set i} {J : Set j} {T : Rel I t}
+ (f : I → J) (P : Rel J p) →
Trans T (P on f) (P on f) →
TransFlip (Star T) (P on f) (P on f)
gfold f P _⊕_ ∅ ε = ∅
gfold f P _⊕_ ∅ (x ◅ xs) = x ⊕ gfold f P _⊕_ ∅ xs
-fold : ∀ {I T} (P : Rel I zero) →
+fold : ∀ {i t p} {I : Set i} {T : Rel I t} (P : Rel I p) →
Trans T P P → Reflexive P → Star T ⇒ P
fold P _⊕_ ∅ = gfold id P _⊕_ ∅
-gfoldl : ∀ {I J T} (f : I → J) P →
+gfoldl : ∀ {i j t p} {I : Set i} {J : Set j} {T : Rel I t}
+ (f : I → J) (P : Rel J p) →
Trans (P on f) T (P on f) →
Trans (P on f) (Star T) (P on f)
gfoldl f P _⊕_ ∅ ε = ∅
gfoldl f P _⊕_ ∅ (x ◅ xs) = gfoldl f P _⊕_ (∅ ⊕ x) xs
-foldl : ∀ {I T} (P : Rel I zero) →
+foldl : ∀ {i t p} {I : Set i} {T : Rel I t} (P : Rel I p) →
Trans P T P → Reflexive P → Star T ⇒ P
foldl P _⊕_ ∅ = gfoldl id P _⊕_ ∅
-concat : ∀ {I} {T : Rel I zero} → Star (Star T) ⇒ Star T
+concat : ∀ {i t} {I : Set i} {T : Rel I t} → Star (Star T) ⇒ Star T
concat {T = T} = fold (Star T) _◅◅_ ε
-- If the underlying relation is symmetric, then the reflexive
-- transitive closure is also symmetric.
-revApp : ∀ {I} {T U : Rel I zero} → Sym T U →
- ∀ {i j k} → Star T j i → Star U j k → Star U i k
+revApp : ∀ {i t u} {I : Set i} {T : Rel I t} {U : Rel I u} →
+ Sym T U → ∀ {i j k} → Star T j i → Star U j k → Star U i k
revApp rev ε ys = ys
revApp rev (x ◅ xs) ys = revApp rev xs (rev x ◅ ys)
-reverse : ∀ {I} {T U : Rel I zero} → Sym T U → Sym (Star T) (Star U)
+reverse : ∀ {i t u} {I : Set i} {T : Rel I t} {U : Rel I u} →
+ Sym T U → Sym (Star T) (Star U)
reverse rev xs = revApp rev xs ε
-- Reflexive transitive closures form a (generalised) monad.
-- return could also be called singleton.
-return : ∀ {I} {T : Rel I zero} → T ⇒ Star T
+return : ∀ {i t} {I : Set i} {T : Rel I t} → T ⇒ Star T
return x = x ◅ ε
-- A generalised variant of the Kleisli star (flip bind, or
-- concatMap).
-kleisliStar : ∀ {I J} {T : Rel I zero} {U : Rel J zero} (f : I → J) →
- T =[ f ]⇒ Star U → Star T =[ f ]⇒ Star U
-kleisliStar f g = concat ∘ gmap f g
+kleisliStar : ∀ {i j t u}
+ {I : Set i} {J : Set j} {T : Rel I t} {U : Rel J u}
+ (f : I → J) → T =[ f ]⇒ Star U → Star T =[ f ]⇒ Star U
+kleisliStar f g = concat ∘′ gmap f g
-_⋆ : ∀ {I} {T U : Rel I zero} →
+_⋆ : ∀ {i t u} {I : Set i} {T : Rel I t} {U : Rel I u} →
T ⇒ Star U → Star T ⇒ Star U
_⋆ = kleisliStar id
infixl 1 _>>=_
-_>>=_ : ∀ {I} {T U : Rel I zero} {i j} →
+_>>=_ : ∀ {i t u} {I : Set i} {T : Rel I t} {U : Rel I u} {i j} →
Star T i j → T ⇒ Star U → Star U i j
m >>= f = (f ⋆) m
@@ -124,7 +131,7 @@ m >>= f = (f ⋆) m
-- (as defined in Category.Monad.Indexed). If it were, then _>>=_
-- would have a type similar to
--
--- ∀ {I} {T U : Rel I zero} {i j k} →
+-- ∀ {I} {T U : Rel I t} {i j k} →
-- Star T i j → (T i j → Star U j k) → Star U i k.
-- ^^^^^
-- Note, however, that there is no scope for applying T to any indices
diff --git a/src/Data/Star/Nat.agda b/src/Data/Star/Nat.agda
index 57f0b0d..b012ce7 100644
--- a/src/Data/Star/Nat.agda
+++ b/src/Data/Star/Nat.agda
@@ -2,6 +2,8 @@
-- Natural numbers defined in terms of Data.Star
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Data.Star.Nat where
open import Data.Star
@@ -9,7 +11,6 @@ open import Data.Unit
open import Relation.Binary
open import Relation.Binary.Simple
open import Function
-import Level as L
-- Natural numbers.
@@ -26,7 +27,7 @@ suc = _◅_ tt
-- The length of a star-list.
-length : ∀ {I} {T : Rel I L.zero} {i j} → Star T i j → ℕ
+length : ∀ {i t} {I : Set i} {T : Rel I t} {i j} → Star T i j → ℕ
length = gmap (const tt) (const tt)
-- Arithmetic.
diff --git a/src/Data/Star/Properties.agda b/src/Data/Star/Properties.agda
index 52465ba..2294a53 100644
--- a/src/Data/Star/Properties.agda
+++ b/src/Data/Star/Properties.agda
@@ -2,31 +2,32 @@
-- Some properties related to Data.Star
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Data.Star.Properties where
open import Data.Star
open import Function
-open import Level
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; refl; sym; cong; cong₂)
import Relation.Binary.PreorderReasoning as PreR
-◅◅-assoc : ∀ {I} {T : Rel I zero} {i j k l}
- (xs : Star T i j) (ys : Star T j k)
- (zs : Star T k l) →
+◅◅-assoc : ∀ {i t} {I : Set i} {T : Rel I t} {i j k l}
+ (xs : Star T i j) (ys : Star T j k)
+ (zs : Star T k l) →
(xs ◅◅ ys) ◅◅ zs ≡ xs ◅◅ (ys ◅◅ zs)
◅◅-assoc ε ys zs = refl
◅◅-assoc (x ◅ xs) ys zs = cong (_◅_ x) (◅◅-assoc xs ys zs)
-gmap-id : ∀ {I} {T : Rel I zero} {i j} (xs : Star T i j) →
+gmap-id : ∀ {i t} {I : Set i} {T : Rel I t} {i j} (xs : Star T i j) →
gmap id id xs ≡ xs
gmap-id ε = refl
gmap-id (x ◅ xs) = cong (_◅_ x) (gmap-id xs)
-gmap-∘ : ∀ {I} {T : Rel I zero}
- {J} {U : Rel J zero}
- {K} {V : Rel K zero}
+gmap-∘ : ∀ {i t} {I : Set i} {T : Rel I t}
+ {j u} {J : Set j} {U : Rel J u}
+ {k v} {K : Set k} {V : Rel K v}
(f : J → K) (g : U =[ f ]⇒ V)
(f′ : I → J) (g′ : T =[ f′ ]⇒ U)
{i j} (xs : Star T i j) →
@@ -34,14 +35,16 @@ gmap-∘ : ∀ {I} {T : Rel I zero}
gmap-∘ f g f′ g′ ε = refl
gmap-∘ f g f′ g′ (x ◅ xs) = cong (_◅_ (g (g′ x))) (gmap-∘ f g f′ g′ xs)
-gmap-◅◅ : ∀ {I} {T : Rel I zero} {J} {U : Rel J zero}
+gmap-◅◅ : ∀ {i t j u}
+ {I : Set i} {T : Rel I t} {J : Set j} {U : Rel J u}
(f : I → J) (g : T =[ f ]⇒ U)
{i j k} (xs : Star T i j) (ys : Star T j k) →
gmap {U = U} f g (xs ◅◅ ys) ≡ gmap f g xs ◅◅ gmap f g ys
gmap-◅◅ f g ε ys = refl
gmap-◅◅ f g (x ◅ xs) ys = cong (_◅_ (g x)) (gmap-◅◅ f g xs ys)
-gmap-cong : ∀ {I} {T : Rel I zero} {J} {U : Rel J zero}
+gmap-cong : ∀ {i t j u}
+ {I : Set i} {T : Rel I t} {J : Set j} {U : Rel J u}
(f : I → J) (g : T =[ f ]⇒ U) (g′ : T =[ f ]⇒ U) →
(∀ {i j} (x : T i j) → g x ≡ g′ x) →
∀ {i j} (xs : Star T i j) →
@@ -49,7 +52,8 @@ gmap-cong : ∀ {I} {T : Rel I zero} {J} {U : Rel J zero}
gmap-cong f g g′ eq ε = refl
gmap-cong f g g′ eq (x ◅ xs) = cong₂ _◅_ (eq x) (gmap-cong f g g′ eq xs)
-fold-◅◅ : ∀ {I} (P : Rel I zero) (_⊕_ : Transitive P) (∅ : Reflexive P) →
+fold-◅◅ : ∀ {i p} {I : Set i}
+ (P : Rel I p) (_⊕_ : Transitive P) (∅ : Reflexive P) →
(∀ {i j} (x : P i j) → ∅ ⊕ x ≡ x) →
(∀ {i j k l} (x : P i j) (y : P j k) (z : P k l) →
(x ⊕ y) ⊕ z ≡ x ⊕ (y ⊕ z)) →
@@ -65,7 +69,7 @@ fold-◅◅ P _⊕_ ∅ left-unit assoc (x ◅ xs) ys = begin
-- Reflexive transitive closures are preorders.
-preorder : ∀ {I} (T : Rel I zero) → Preorder _ _ _
+preorder : ∀ {i t} {I : Set i} (T : Rel I t) → Preorder _ _ _
preorder T = record
{ _≈_ = _≡_
; _∼_ = Star T
@@ -81,7 +85,7 @@ preorder T = record
-- Preorder reasoning for Star.
-module StarReasoning {I : Set} (T : Rel I zero) where
+module StarReasoning {i t} {I : Set i} (T : Rel I t) where
open PreR (preorder T) public
renaming (_∼⟨_⟩_ to _⟶⋆⟨_⟩_; _≈⟨_⟩_ to _≡⟨_⟩_)
diff --git a/src/Data/Vec.agda b/src/Data/Vec.agda
index a87fed1..3181cdb 100644
--- a/src/Data/Vec.agda
+++ b/src/Data/Vec.agda
@@ -6,10 +6,12 @@
module Data.Vec where
+open import Category.Applicative
open import Data.Nat
open import Data.Fin using (Fin; zero; suc)
open import Data.List as List using (List)
open import Data.Product using (∃; ∃₂; _×_; _,_)
+open import Function
open import Level
open import Relation.Binary.PropositionalEquality
@@ -54,24 +56,35 @@ _++_ : ∀ {a m n} {A : Set a} → Vec A m → Vec A n → Vec A (m + n)
[] ++ ys = ys
(x ∷ xs) ++ ys = x ∷ (xs ++ ys)
+infixl 4 _⊛_
+
+_⊛_ : ∀ {a b n} {A : Set a} {B : Set b} →
+ Vec (A → B) n → Vec A n → Vec B n
+[] ⊛ [] = []
+(f ∷ fs) ⊛ (x ∷ xs) = f x ∷ (fs ⊛ xs)
+
+replicate : ∀ {a n} {A : Set a} → A → Vec A n
+replicate {n = zero} x = []
+replicate {n = suc n} x = x ∷ replicate x
+
+applicative : ∀ {a n} → RawApplicative (λ (A : Set a) → Vec A n)
+applicative = record
+ { pure = replicate
+ ; _⊛_ = _⊛_
+ }
+
map : ∀ {a b n} {A : Set a} {B : Set b} →
(A → B) → Vec A n → Vec B n
-map f [] = []
-map f (x ∷ xs) = f x ∷ map f xs
+map f xs = replicate f ⊛ xs
zipWith : ∀ {a b c n} {A : Set a} {B : Set b} {C : Set c} →
(A → B → C) → Vec A n → Vec B n → Vec C n
-zipWith _⊕_ [] [] = []
-zipWith _⊕_ (x ∷ xs) (y ∷ ys) = (x ⊕ y) ∷ zipWith _⊕_ xs ys
+zipWith _⊕_ xs ys = replicate _⊕_ ⊛ xs ⊛ ys
zip : ∀ {a b n} {A : Set a} {B : Set b} →
Vec A n → Vec B n → Vec (A × B) n
zip = zipWith _,_
-replicate : ∀ {a n} {A : Set a} → A → Vec A n
-replicate {n = zero} x = []
-replicate {n = suc n} x = x ∷ replicate x
-
foldr : ∀ {a b} {A : Set a} (B : ℕ → Set b) {m} →
(∀ {n} → A → B n → B (suc n)) →
B zero →
@@ -166,11 +179,11 @@ _>>=_ : ∀ {a b m n} {A : Set a} {B : Set b} →
Vec A m → (A → Vec B n) → Vec B (m * n)
xs >>= f = concat (map f xs)
-infixl 4 _⊛_
+infixl 4 _⊛*_
-_⊛_ : ∀ {a b m n} {A : Set a} {B : Set b} →
- Vec (A → B) m → Vec A n → Vec B (m * n)
-fs ⊛ xs = fs >>= λ f → map f xs
+_⊛*_ : ∀ {a b m n} {A : Set a} {B : Set b} →
+ Vec (A → B) m → Vec A n → Vec B (m * n)
+fs ⊛* xs = fs >>= λ f → map f xs
-- Interleaves the two vectors.
@@ -181,10 +194,18 @@ _⋎_ : ∀ {a m n} {A : Set a} →
[] ⋎ ys = ys
(x ∷ xs) ⋎ ys = x ∷ (ys ⋎ xs)
+-- A lookup function.
+
lookup : ∀ {a n} {A : Set a} → Fin n → Vec A n → A
lookup zero (x ∷ xs) = x
lookup (suc i) (x ∷ xs) = lookup i xs
+-- An inverse of flip lookup.
+
+tabulate : ∀ {n a} {A : Set a} → (Fin n → A) → Vec A n
+tabulate {zero} f = []
+tabulate {suc n} f = f zero ∷ tabulate (f ∘ suc)
+
-- Update.
infixl 6 _[_]≔_
@@ -195,8 +216,10 @@ _[_]≔_ : ∀ {a n} {A : Set a} → Vec A n → Fin n → A → Vec A n
(x ∷ xs) [ suc i ]≔ y = x ∷ xs [ i ]≔ y
-- Generates a vector containing all elements in Fin n. This function
--- is not placed in Data.Fin since Data.Vec depends on Data.Fin.
+-- is not placed in Data.Fin because Data.Vec depends on Data.Fin.
+--
+-- The implementation was suggested by Conor McBride ("Fwd: how to
+-- count 0..n-1", http://thread.gmane.org/gmane.comp.lang.agda/2554).
allFin : ∀ n → Vec (Fin n) n
-allFin zero = []
-allFin (suc n) = zero ∷ map suc (allFin n)
+allFin _ = tabulate id
diff --git a/src/Data/Vec/N-ary.agda b/src/Data/Vec/N-ary.agda
index bc5512d..0139ccf 100644
--- a/src/Data/Vec/N-ary.agda
+++ b/src/Data/Vec/N-ary.agda
@@ -6,9 +6,11 @@
module Data.Vec.N-ary where
-open import Function
-open import Data.Nat
+open import Data.Nat hiding (_⊔_)
+open import Data.Product as Prod
open import Data.Vec
+open import Function
+open import Function.Equivalence using (_⇔_; equivalent)
open import Level
open import Relation.Binary
open import Relation.Binary.PropositionalEquality
@@ -17,98 +19,153 @@ open import Relation.Nullary.Decidable
------------------------------------------------------------------------
-- N-ary functions
-N-ary : ∀ {ℓ} → ℕ → Set ℓ → Set ℓ → Set ℓ
+N-ary-level : Level → Level → ℕ → Level
+N-ary-level ℓ₁ ℓ₂ zero = ℓ₂
+N-ary-level ℓ₁ ℓ₂ (suc n) = ℓ₁ ⊔ N-ary-level ℓ₁ ℓ₂ n
+
+N-ary : ∀ {ℓ₁ ℓ₂} (n : ℕ) → Set ℓ₁ → Set ℓ₂ → Set (N-ary-level ℓ₁ ℓ₂ n)
N-ary zero A B = B
N-ary (suc n) A B = A → N-ary n A B
------------------------------------------------------------------------
-- Conversion
-curryⁿ : ∀ {n ℓ} {A B : Set ℓ} → (Vec A n → B) → N-ary n A B
+curryⁿ : ∀ {n a b} {A : Set a} {B : Set b} →
+ (Vec A n → B) → N-ary n A B
curryⁿ {zero} f = f []
curryⁿ {suc n} f = λ x → curryⁿ (f ∘ _∷_ x)
-_$ⁿ_ : ∀ {n ℓ} {A B : Set ℓ} → N-ary n A B → (Vec A n → B)
+_$ⁿ_ : ∀ {n a b} {A : Set a} {B : Set b} → N-ary n A B → (Vec A n → B)
f $ⁿ [] = f
f $ⁿ (x ∷ xs) = f x $ⁿ xs
------------------------------------------------------------------------
--- N-ary function equality
+-- Quantifiers
+
+-- Universal quantifier.
+
+∀ⁿ : ∀ n {a ℓ} {A : Set a} →
+ N-ary n A (Set ℓ) → Set (N-ary-level a ℓ n)
+∀ⁿ zero P = P
+∀ⁿ (suc n) P = ∀ x → ∀ⁿ n (P x)
+
+-- Universal quantifier with implicit (hidden) arguments.
-Eq-level : Level → Level → ℕ → Level
-Eq-level _ _ zero = _
-Eq-level _ _ (suc n) = _
+∀ⁿʰ : ∀ n {a ℓ} {A : Set a} →
+ N-ary n A (Set ℓ) → Set (N-ary-level a ℓ n)
+∀ⁿʰ zero P = P
+∀ⁿʰ (suc n) P = ∀ {x} → ∀ⁿʰ n (P x)
-Eq : ∀ {ℓ₁ ℓ₂} {A B C : Set ℓ₁} n →
- REL B C ℓ₂ → REL (N-ary n A B) (N-ary n A C) (Eq-level ℓ₁ ℓ₂ n)
-Eq zero _∼_ f g = f ∼ g
-Eq (suc n) _∼_ f g = ∀ x → Eq n _∼_ (f x) (g x)
+-- Existential quantifier.
+
+∃ⁿ : ∀ n {a ℓ} {A : Set a} →
+ N-ary n A (Set ℓ) → Set (N-ary-level a ℓ n)
+∃ⁿ zero P = P
+∃ⁿ (suc n) P = ∃ λ x → ∃ⁿ n (P x)
+
+------------------------------------------------------------------------
+-- N-ary function equality
+
+Eq : ∀ {a b c ℓ} {A : Set a} {B : Set b} {C : Set c} n →
+ REL B C ℓ → REL (N-ary n A B) (N-ary n A C) (N-ary-level a ℓ n)
+Eq n _∼_ f g = ∀ⁿ n (curryⁿ {n = n} λ xs → (f $ⁿ xs) ∼ (g $ⁿ xs))
-- A variant where all the arguments are implicit (hidden).
-Eqʰ : ∀ {ℓ₁ ℓ₂} {A B C : Set ℓ₁} n →
- REL B C ℓ₂ → REL (N-ary n A B) (N-ary n A C) (Eq-level ℓ₁ ℓ₂ n)
-Eqʰ zero _∼_ f g = f ∼ g
-Eqʰ (suc n) _∼_ f g = ∀ {x} → Eqʰ n _∼_ (f x) (g x)
+Eqʰ : ∀ {a b c ℓ} {A : Set a} {B : Set b} {C : Set c} n →
+ REL B C ℓ → REL (N-ary n A B) (N-ary n A C) (N-ary-level a ℓ n)
+Eqʰ n _∼_ f g = ∀ⁿʰ n (curryⁿ {n = n} λ xs → (f $ⁿ xs) ∼ (g $ⁿ xs))
------------------------------------------------------------------------
-- Some lemmas
-- The functions curryⁿ and _$ⁿ_ are inverses.
-left-inverse : ∀ {n ℓ} {A B : Set ℓ} (f : Vec A n → B) →
+left-inverse : ∀ {n a b} {A : Set a} {B : Set b} (f : Vec A n → B) →
∀ xs → curryⁿ f $ⁿ xs ≡ f xs
left-inverse f [] = refl
left-inverse f (x ∷ xs) = left-inverse (f ∘ _∷_ x) xs
-right-inverse : ∀ {ℓ} {A B : Set ℓ} n (f : N-ary n A B) →
+right-inverse : ∀ {a b} {A : Set a} {B : Set b} n (f : N-ary n A B) →
Eq n _≡_ (curryⁿ (_$ⁿ_ {n} f)) f
right-inverse zero f = refl
right-inverse (suc n) f = λ x → right-inverse n (f x)
+-- ∀ⁿ can be expressed in an "uncurried" way.
+
+uncurry-∀ⁿ : ∀ n {a ℓ} {A : Set a} {P : N-ary n A (Set ℓ)} →
+ ∀ⁿ n P ⇔ (∀ (xs : Vec A n) → P $ⁿ xs)
+uncurry-∀ⁿ n {a} {ℓ} {A} = equivalent (⇒ n) (⇐ n)
+ where
+ ⇒ : ∀ n {P : N-ary n A (Set ℓ)} →
+ ∀ⁿ n P → (∀ (xs : Vec A n) → P $ⁿ xs)
+ ⇒ zero p [] = p
+ ⇒ (suc n) p (x ∷ xs) = ⇒ n (p x) xs
+
+ ⇐ : ∀ n {P : N-ary n A (Set ℓ)} →
+ (∀ (xs : Vec A n) → P $ⁿ xs) → ∀ⁿ n P
+ ⇐ zero p = p []
+ ⇐ (suc n) p = λ x → ⇐ n (p ∘ _∷_ x)
+
+-- ∃ⁿ can be expressed in an "uncurried" way.
+
+uncurry-∃ⁿ : ∀ n {a ℓ} {A : Set a} {P : N-ary n A (Set ℓ)} →
+ ∃ⁿ n P ⇔ (∃ λ (xs : Vec A n) → P $ⁿ xs)
+uncurry-∃ⁿ n {a} {ℓ} {A} = equivalent (⇒ n) (⇐ n)
+ where
+ ⇒ : ∀ n {P : N-ary n A (Set ℓ)} →
+ ∃ⁿ n P → (∃ λ (xs : Vec A n) → P $ⁿ xs)
+ ⇒ zero p = ([] , p)
+ ⇒ (suc n) (x , p) = Prod.map (_∷_ x) id (⇒ n p)
+
+ ⇐ : ∀ n {P : N-ary n A (Set ℓ)} →
+ (∃ λ (xs : Vec A n) → P $ⁿ xs) → ∃ⁿ n P
+ ⇐ zero ([] , p) = p
+ ⇐ (suc n) (x ∷ xs , p) = (x , ⇐ n (xs , p))
+
-- Conversion preserves equality.
-curryⁿ-cong : ∀ {n ℓ₁ ℓ₂} {A B C : Set ℓ₁} {_∼_ : REL B C ℓ₂}
- (f : Vec A n → B) (g : Vec A n → C) →
+curryⁿ-cong : ∀ {n a b c ℓ} {A : Set a} {B : Set b} {C : Set c}
+ (_∼_ : REL B C ℓ) (f : Vec A n → B) (g : Vec A n → C) →
(∀ xs → f xs ∼ g xs) →
Eq n _∼_ (curryⁿ f) (curryⁿ g)
-curryⁿ-cong {zero} f g hyp = hyp []
-curryⁿ-cong {suc n} f g hyp = λ x →
- curryⁿ-cong (f ∘ _∷_ x) (g ∘ _∷_ x) (λ xs → hyp (x ∷ xs))
+curryⁿ-cong {zero} _∼_ f g hyp = hyp []
+curryⁿ-cong {suc n} _∼_ f g hyp = λ x →
+ curryⁿ-cong _∼_ (f ∘ _∷_ x) (g ∘ _∷_ x) (λ xs → hyp (x ∷ xs))
-curryⁿ-cong⁻¹ : ∀ {n ℓ₁ ℓ₂} {A B C : Set ℓ₁} {_∼_ : REL B C ℓ₂}
- (f : Vec A n → B) (g : Vec A n → C) →
+curryⁿ-cong⁻¹ : ∀ {n a b c ℓ} {A : Set a} {B : Set b} {C : Set c}
+ (_∼_ : REL B C ℓ) (f : Vec A n → B) (g : Vec A n → C) →
Eq n _∼_ (curryⁿ f) (curryⁿ g) →
∀ xs → f xs ∼ g xs
-curryⁿ-cong⁻¹ f g hyp [] = hyp
-curryⁿ-cong⁻¹ f g hyp (x ∷ xs) =
- curryⁿ-cong⁻¹ (f ∘ _∷_ x) (g ∘ _∷_ x) (hyp x) xs
+curryⁿ-cong⁻¹ _∼_ f g hyp [] = hyp
+curryⁿ-cong⁻¹ _∼_ f g hyp (x ∷ xs) =
+ curryⁿ-cong⁻¹ _∼_ (f ∘ _∷_ x) (g ∘ _∷_ x) (hyp x) xs
-appⁿ-cong : ∀ {n ℓ₁ ℓ₂} {A B C : Set ℓ₁} {_∼_ : REL B C ℓ₂}
- (f : N-ary n A B) (g : N-ary n A C) →
+appⁿ-cong : ∀ {n a b c ℓ} {A : Set a} {B : Set b} {C : Set c}
+ (_∼_ : REL B C ℓ) (f : N-ary n A B) (g : N-ary n A C) →
Eq n _∼_ f g →
(xs : Vec A n) → (f $ⁿ xs) ∼ (g $ⁿ xs)
-appⁿ-cong f g hyp [] = hyp
-appⁿ-cong f g hyp (x ∷ xs) = appⁿ-cong (f x) (g x) (hyp x) xs
+appⁿ-cong _∼_ f g hyp [] = hyp
+appⁿ-cong _∼_ f g hyp (x ∷ xs) = appⁿ-cong _∼_ (f x) (g x) (hyp x) xs
-appⁿ-cong⁻¹ : ∀ {n ℓ₁ ℓ₂} {A B C : Set ℓ₁} {_∼_ : REL B C ℓ₂}
- (f : N-ary n A B) (g : N-ary n A C) →
+appⁿ-cong⁻¹ : ∀ {n a b c ℓ} {A : Set a} {B : Set b} {C : Set c}
+ (_∼_ : REL B C ℓ) (f : N-ary n A B) (g : N-ary n A C) →
((xs : Vec A n) → (f $ⁿ xs) ∼ (g $ⁿ xs)) →
Eq n _∼_ f g
-appⁿ-cong⁻¹ {zero} f g hyp = hyp []
-appⁿ-cong⁻¹ {suc n} f g hyp = λ x →
- appⁿ-cong⁻¹ (f x) (g x) (λ xs → hyp (x ∷ xs))
+appⁿ-cong⁻¹ {zero} _∼_ f g hyp = hyp []
+appⁿ-cong⁻¹ {suc n} _∼_ f g hyp = λ x →
+ appⁿ-cong⁻¹ _∼_ (f x) (g x) (λ xs → hyp (x ∷ xs))
-- Eq and Eqʰ are equivalent.
-Eq-to-Eqʰ : ∀ {ℓ₁ ℓ₂} {A B C : Set ℓ₁} n {_∼_ : REL B C ℓ₂}
- {f : N-ary n A B} {g : N-ary n A C} →
+Eq-to-Eqʰ : ∀ {a b c ℓ} {A : Set a} {B : Set b} {C : Set c}
+ n (_∼_ : REL B C ℓ) {f : N-ary n A B} {g : N-ary n A C} →
Eq n _∼_ f g → Eqʰ n _∼_ f g
-Eq-to-Eqʰ zero eq = eq
-Eq-to-Eqʰ (suc n) eq = Eq-to-Eqʰ n (eq _)
+Eq-to-Eqʰ zero _∼_ eq = eq
+Eq-to-Eqʰ (suc n) _∼_ eq = Eq-to-Eqʰ n _∼_ (eq _)
-Eqʰ-to-Eq : ∀ {ℓ₁ ℓ₂} {A B C : Set ℓ₁} n {_∼_ : REL B C ℓ₂}
- {f : N-ary n A B} {g : N-ary n A C} →
+Eqʰ-to-Eq : ∀ {a b c ℓ} {A : Set a} {B : Set b} {C : Set c}
+ n (_∼_ : REL B C ℓ) {f : N-ary n A B} {g : N-ary n A C} →
Eqʰ n _∼_ f g → Eq n _∼_ f g
-Eqʰ-to-Eq zero eq = eq
-Eqʰ-to-Eq (suc n) eq = λ _ → Eqʰ-to-Eq n eq
+Eqʰ-to-Eq zero _∼_ eq = eq
+Eqʰ-to-Eq (suc n) _∼_ eq = λ _ → Eqʰ-to-Eq n _∼_ eq
diff --git a/src/Data/Vec/Properties.agda b/src/Data/Vec/Properties.agda
index 02fabc6..1fa6d08 100644
--- a/src/Data/Vec/Properties.agda
+++ b/src/Data/Vec/Properties.agda
@@ -7,11 +7,16 @@
module Data.Vec.Properties where
open import Algebra
+open import Category.Applicative.Indexed
+open import Category.Monad
+open import Category.Monad.Identity
open import Data.Vec
open import Data.Nat
import Data.Nat.Properties as Nat
-open import Data.Fin using (Fin; zero; suc)
+open import Data.Fin as Fin using (Fin; zero; suc; toℕ; fromℕ)
+open import Data.Fin.Props using (_+′_)
open import Function
+open import Function.Inverse using (_⇿_)
open import Level
open import Relation.Binary
@@ -41,13 +46,76 @@ module UsingVectorEquality {s₁ s₂} (S : Setoid s₁ s₂) where
open import Relation.Binary.PropositionalEquality as PropEq
open import Relation.Binary.HeterogeneousEquality as HetEq
--- lookup is natural.
+-- lookup is an applicative functor morphism.
-lookup-natural : ∀ {a b n} {A : Set a} {B : Set b}
- (f : A → B) (i : Fin n) →
- lookup i ∘ map f ≗ f ∘ lookup i
-lookup-natural f zero (x ∷ xs) = refl
-lookup-natural f (suc i) (x ∷ xs) = lookup-natural f i xs
+lookup-morphism :
+ ∀ {a n} (i : Fin n) →
+ Morphism (applicative {a = a})
+ (RawMonad.rawIApplicative IdentityMonad)
+lookup-morphism i = record
+ { op = lookup i
+ ; op-pure = lookup-replicate i
+ ; op-⊛ = lookup-⊛ i
+ }
+ where
+ lookup-replicate : ∀ {a n} {A : Set a} (i : Fin n) →
+ lookup i ∘ replicate {A = A} ≗ id {A = A}
+ lookup-replicate zero = λ _ → refl
+ lookup-replicate (suc i) = lookup-replicate i
+
+ lookup-⊛ : ∀ {a b n} {A : Set a} {B : Set b}
+ i (fs : Vec (A → B) n) (xs : Vec A n) →
+ lookup i (fs ⊛ xs) ≡ (lookup i fs $ lookup i xs)
+ lookup-⊛ zero (f ∷ fs) (x ∷ xs) = refl
+ lookup-⊛ (suc i) (f ∷ fs) (x ∷ xs) = lookup-⊛ i fs xs
+
+-- tabulate is an inverse of flip lookup.
+
+lookup∘tabulate : ∀ {a n} {A : Set a} (f : Fin n → A) (i : Fin n) →
+ lookup i (tabulate f) ≡ f i
+lookup∘tabulate f zero = refl
+lookup∘tabulate f (suc i) = lookup∘tabulate (f ∘ suc) i
+
+tabulate∘lookup : ∀ {a n} {A : Set a} (xs : Vec A n) →
+ tabulate (flip lookup xs) ≡ xs
+tabulate∘lookup [] = refl
+tabulate∘lookup (x ∷ xs) = PropEq.cong (_∷_ x) $ tabulate∘lookup xs
+
+-- If you look up an index in allFin n, then you get the index.
+
+lookup-allFin : ∀ {n} (i : Fin n) → lookup i (allFin n) ≡ i
+lookup-allFin = lookup∘tabulate id
+
+-- Various lemmas relating lookup and _++_.
+
+lookup-++-< : ∀ {a} {A : Set a} {m n}
+ (xs : Vec A m) (ys : Vec A n) i (i<m : toℕ i < m) →
+ lookup i (xs ++ ys) ≡ lookup (Fin.fromℕ≤ i<m) xs
+lookup-++-< [] ys i ()
+lookup-++-< (x ∷ xs) ys zero (s≤s z≤n) = refl
+lookup-++-< (x ∷ xs) ys (suc i) (s≤s (s≤s i<m)) =
+ lookup-++-< xs ys i (s≤s i<m)
+
+lookup-++-≥ : ∀ {a} {A : Set a} {m n}
+ (xs : Vec A m) (ys : Vec A n) i (i≥m : toℕ i ≥ m) →
+ lookup i (xs ++ ys) ≡ lookup (Fin.reduce≥ i i≥m) ys
+lookup-++-≥ [] ys i i≥m = refl
+lookup-++-≥ (x ∷ xs) ys zero ()
+lookup-++-≥ (x ∷ xs) ys (suc i) (s≤s i≥m) = lookup-++-≥ xs ys i i≥m
+
+lookup-++-inject+ : ∀ {a} {A : Set a} {m n}
+ (xs : Vec A m) (ys : Vec A n) i →
+ lookup (Fin.inject+ n i) (xs ++ ys) ≡ lookup i xs
+lookup-++-inject+ [] ys ()
+lookup-++-inject+ (x ∷ xs) ys zero = refl
+lookup-++-inject+ (x ∷ xs) ys (suc i) = lookup-++-inject+ xs ys i
+
+lookup-++-+′ : ∀ {a} {A : Set a} {m n}
+ (xs : Vec A m) (ys : Vec A n) i →
+ lookup (fromℕ m +′ i) (xs ++ ys) ≡ lookup i ys
+lookup-++-+′ [] ys zero = refl
+lookup-++-+′ [] (y ∷ xs) (suc i) = lookup-++-+′ [] xs i
+lookup-++-+′ (x ∷ xs) ys i = lookup-++-+′ xs ys i
-- map is a congruence.
@@ -68,6 +136,35 @@ map-∘ : ∀ {a b c n} {A : Set a} {B : Set b} {C : Set c}
map-∘ f g [] = refl
map-∘ f g (x ∷ xs) = PropEq.cong (_∷_ (f (g x))) (map-∘ f g xs)
+-- tabulate can be expressed using map and allFin.
+
+tabulate-∘ : ∀ {n a b} {A : Set a} {B : Set b}
+ (f : A → B) (g : Fin n → A) →
+ tabulate (f ∘ g) ≡ map f (tabulate g)
+tabulate-∘ {zero} f g = refl
+tabulate-∘ {suc n} f g =
+ PropEq.cong (_∷_ (f (g zero))) (tabulate-∘ f (g ∘ suc))
+
+tabulate-allFin : ∀ {n a} {A : Set a} (f : Fin n → A) →
+ tabulate f ≡ map f (allFin n)
+tabulate-allFin f = tabulate-∘ f id
+
+-- The positive case of allFin can be expressed recursively using map.
+
+allFin-map : ∀ n → allFin (suc n) ≡ zero ∷ map suc (allFin n)
+allFin-map n = PropEq.cong (_∷_ zero) $ tabulate-∘ suc id
+
+-- If you look up every possible index, in increasing order, then you
+-- get back the vector you started with.
+
+map-lookup-allFin : ∀ {a} {A : Set a} {n} (xs : Vec A n) →
+ map (λ x → lookup x xs) (allFin n) ≡ xs
+map-lookup-allFin {n = n} xs = begin
+ map (λ x → lookup x xs) (allFin n) ≡⟨ PropEq.sym $ tabulate-∘ (λ x → lookup x xs) id ⟩
+ tabulate (λ x → lookup x xs) ≡⟨ tabulate∘lookup xs ⟩
+ xs ∎
+ where open ≡-Reasoning
+
-- sum commutes with _++_.
sum-++-commute : ∀ {m n} (xs : Vec ℕ m) {ys : Vec ℕ n} →
@@ -149,3 +246,32 @@ foldr-fusion {B = B} {f} e {C} h fuse =
idIsFold : ∀ {a n} {A : Set a} → id ≗ foldr (Vec A) {n} _∷_ []
idIsFold = foldr-universal _ _ id refl (λ _ _ → refl)
+
+-- Proof irrelevance for _[_]=_.
+
+proof-irrelevance-[]= : ∀ {a} {A : Set a} {n} {xs : Vec A n} {i x} →
+ (p q : xs [ i ]= x) → p ≡ q
+proof-irrelevance-[]= here here = refl
+proof-irrelevance-[]= (there xs[i]=x) (there xs[i]=x') =
+ PropEq.cong there (proof-irrelevance-[]= xs[i]=x xs[i]=x')
+
+-- _[_]=_ can be expressed using lookup and _≡_.
+
+[]=⇿lookup : ∀ {a n i} {A : Set a} {x} {xs : Vec A n} →
+ xs [ i ]= x ⇿ lookup i xs ≡ x
+[]=⇿lookup {i = i} {x = x} {xs} = record
+ { to = PropEq.→-to-⟶ to
+ ; from = PropEq.→-to-⟶ (from i xs)
+ ; inverse-of = record
+ { left-inverse-of = λ _ → proof-irrelevance-[]= _ _
+ ; right-inverse-of = λ _ → PropEq.proof-irrelevance _ _
+ }
+ }
+ where
+ to : ∀ {n xs} {i : Fin n} → xs [ i ]= x → lookup i xs ≡ x
+ to here = refl
+ to (there xs[i]=x) = to xs[i]=x
+
+ from : ∀ {n} (i : Fin n) xs → lookup i xs ≡ x → xs [ i ]= x
+ from zero (.x ∷ _) refl = here
+ from (suc i) (_ ∷ xs) p = there (from i xs p)
diff --git a/src/Data/W.agda b/src/Data/W.agda
new file mode 100644
index 0000000..5e68a9d
--- /dev/null
+++ b/src/Data/W.agda
@@ -0,0 +1,31 @@
+------------------------------------------------------------------------
+-- W-types
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Data.W where
+
+open import Level
+open import Relation.Nullary
+
+-- The family of W-types.
+
+data W {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where
+ sup : (x : A) (f : B x → W A B) → W A B
+
+-- Projections.
+
+head : ∀ {a b} {A : Set a} {B : A → Set b} →
+ W A B → A
+head (sup x f) = x
+
+tail : ∀ {a b} {A : Set a} {B : A → Set b} →
+ (x : W A B) → B (head x) → W A B
+tail (sup x f) = f
+
+-- If B is always inhabited, then W A B is empty.
+
+inhabited⇒empty : ∀ {a b} {A : Set a} {B : A → Set b} →
+ (∀ x → B x) → ¬ W A B
+inhabited⇒empty b (sup x f) = inhabited⇒empty b (f (b x))
diff --git a/src/Function/Inverse/TypeIsomorphisms.agda b/src/Function/Inverse/TypeIsomorphisms.agda
index 4c58303..6a26889 100644
--- a/src/Function/Inverse/TypeIsomorphisms.agda
+++ b/src/Function/Inverse/TypeIsomorphisms.agda
@@ -15,7 +15,8 @@ open import Data.Unit
open import Level
open import Function
open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence using (_⇔_; module Equivalent)
+open import Function.Equivalence
+ using (_⇔_; equivalent; module Equivalent)
open import Function.Inverse as Inv
using (Kind; Isomorphism; _⇿_; module Inverse)
open import Relation.Binary
@@ -246,25 +247,75 @@ open import Relation.Nullary
}
------------------------------------------------------------------------
+-- _→_ preserves isomorphisms
+
+_→-cong-⇔_ :
+ ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
+ A ⇔ B → C ⇔ D → (A → C) ⇔ (B → D)
+A⇔B →-cong-⇔ C⇔D = record
+ { to = P.→-to-⟶ λ f x →
+ Equivalent.to C⇔D ⟨$⟩ f (Equivalent.from A⇔B ⟨$⟩ x)
+ ; from = P.→-to-⟶ λ f x →
+ Equivalent.from C⇔D ⟨$⟩ f (Equivalent.to A⇔B ⟨$⟩ x)
+ }
+
+→-cong :
+ ∀ {a b c d} →
+ P.Extensionality a c → P.Extensionality b d →
+ ∀ {k} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
+ Isomorphism k A B → Isomorphism k C D → Isomorphism k (A → C) (B → D)
+→-cong extAC extBD {Inv.equivalent} A⇔B C⇔D = A⇔B →-cong-⇔ C⇔D
+→-cong extAC extBD {Inv.inverse} A⇿B C⇿D = record
+ { to = Equivalent.to A→C⇔B→D
+ ; from = Equivalent.from A→C⇔B→D
+ ; inverse-of = record
+ { left-inverse-of = λ f → extAC λ x → begin
+ Inverse.from C⇿D ⟨$⟩ (Inverse.to C⇿D ⟨$⟩
+ f (Inverse.from A⇿B ⟨$⟩ (Inverse.to A⇿B ⟨$⟩ x))) ≡⟨ Inverse.left-inverse-of C⇿D _ ⟩
+ f (Inverse.from A⇿B ⟨$⟩ (Inverse.to A⇿B ⟨$⟩ x)) ≡⟨ P.cong f $ Inverse.left-inverse-of A⇿B x ⟩
+ f x ∎
+ ; right-inverse-of = λ f → extBD λ x → begin
+ Inverse.to C⇿D ⟨$⟩ (Inverse.from C⇿D ⟨$⟩
+ f (Inverse.to A⇿B ⟨$⟩ (Inverse.from A⇿B ⟨$⟩ x))) ≡⟨ Inverse.right-inverse-of C⇿D _ ⟩
+ f (Inverse.to A⇿B ⟨$⟩ (Inverse.from A⇿B ⟨$⟩ x)) ≡⟨ P.cong f $ Inverse.right-inverse-of A⇿B x ⟩
+ f x ∎
+ }
+ }
+ where
+ open P.≡-Reasoning
+ A→C⇔B→D = Inv.⇿⇒ A⇿B →-cong-⇔ Inv.⇿⇒ C⇿D
+
+------------------------------------------------------------------------
-- ¬_ preserves isomorphisms
¬-cong-⇔ : ∀ {a b} {A : Set a} {B : Set b} →
A ⇔ B → (¬ A) ⇔ (¬ B)
-¬-cong-⇔ A⇔B = record
- { to = P.→-to-⟶ (λ ¬a b → ¬a (Equivalent.from A⇔B ⟨$⟩ b))
- ; from = P.→-to-⟶ (λ ¬b a → ¬b (Equivalent.to A⇔B ⟨$⟩ a))
- }
+¬-cong-⇔ A⇔B = A⇔B →-cong-⇔ (⊥ ∎)
+ where open Inv.EquationalReasoning
¬-cong : ∀ {a b} →
P.Extensionality a zero → P.Extensionality b zero →
∀ {k} {A : Set a} {B : Set b} →
Isomorphism k A B → Isomorphism k (¬ A) (¬ B)
-¬-cong _ _ {k = Inv.equivalent} A⇔B = ¬-cong-⇔ A⇔B
-¬-cong extA extB {k = Inv.inverse} {A = A} {B} A⇿B = record
- { to = Equivalent.to ¬A⇔¬B
- ; from = Equivalent.from ¬A⇔¬B
- ; inverse-of = record
- { left-inverse-of = λ ¬a → extA (λ a → P.cong ¬a (Inverse.left-inverse-of A⇿B a))
- ; right-inverse-of = λ ¬b → extB (λ b → P.cong ¬b (Inverse.right-inverse-of A⇿B b))
- }
- } where ¬A⇔¬B = ¬-cong-⇔ (Inv.⇿⇒ A⇿B)
+¬-cong extA extB A≈B = →-cong extA extB A≈B (⊥ ∎)
+ where open Inv.EquationalReasoning
+
+------------------------------------------------------------------------
+-- _⇔_ preserves _⇔_
+
+-- The type of the following proof is a bit more general.
+
+Isomorphism-cong :
+ ∀ {k a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
+ Isomorphism k A B → Isomorphism k C D →
+ Isomorphism k A C ⇔ Isomorphism k B D
+Isomorphism-cong {A = A} {B} {C} {D} A≈B C≈D =
+ equivalent (λ A≈C → B ≈⟨ sym A≈B ⟩
+ A ≈⟨ A≈C ⟩
+ C ≈⟨ C≈D ⟩
+ D ∎)
+ (λ B≈D → A ≈⟨ A≈B ⟩
+ B ≈⟨ B≈D ⟩
+ D ≈⟨ sym C≈D ⟩
+ C ∎)
+ where open Inv.EquationalReasoning
diff --git a/src/Function/Surjection.agda b/src/Function/Surjection.agda
index a0269cc..fcf586d 100644
--- a/src/Function/Surjection.agda
+++ b/src/Function/Surjection.agda
@@ -14,6 +14,7 @@ open import Function.Injection hiding (id; _∘_)
open import Function.LeftInverse as Left hiding (id; _∘_)
open import Data.Product
open import Relation.Binary
+import Relation.Binary.PropositionalEquality as P
-- Surjective functions.
@@ -25,7 +26,7 @@ record Surjective {f₁ f₂ t₁ t₂}
from : To ⟶ From
right-inverse-of : from RightInverseOf to
--- The set of all surjections between two setoids.
+-- The set of all surjections from one setoid to another.
record Surjection {f₁ f₂ t₁ t₂}
(From : Setoid f₁ f₂) (To : Setoid t₁ t₂) :
@@ -55,6 +56,13 @@ record Surjection {f₁ f₂ t₁ t₂}
; from = from
}
+-- The set of all surjections from one set to another.
+
+infix 3 _↠_
+
+_↠_ : ∀ {f t} → Set f → Set t → Set _
+From ↠ To = Surjection (P.setoid From) (P.setoid To)
+
-- Identity and composition.
id : ∀ {s₁ s₂} {S : Setoid s₁ s₂} → Surjection S S
diff --git a/src/IO.agda b/src/IO.agda
index 94fbbaf..5d9ad1a 100644
--- a/src/IO.agda
+++ b/src/IO.agda
@@ -48,7 +48,7 @@ abstract
-- Because IO A lives in Set₁ I hesitate to define sequence, which
-- would require defining a Set₁ variant of Colist.
-mapM : ∀ {A B} → (A → IO B) → Colist A → IO (Colist B)
+mapM : {A B : Set} → (A → IO B) → Colist A → IO (Colist B)
mapM f [] = return []
mapM f (x ∷ xs) = ♯ f x >>= λ y →
♯ (♯ mapM f (♭ xs) >>= λ ys →
@@ -57,7 +57,7 @@ mapM f (x ∷ xs) = ♯ f x >>= λ y →
-- The reason for not defining mapM′ in terms of mapM is efficiency
-- (the unused results could cause unnecessary memory use).
-mapM′ : ∀ {A B} → (A → IO B) → Colist A → IO ⊤
+mapM′ : {A B : Set} → (A → IO B) → Colist A → IO ⊤
mapM′ f [] = return _
mapM′ f (x ∷ xs) = ♯ f x >> ♯ mapM′ f (♭ xs)
diff --git a/src/Reflection.agda b/src/Reflection.agda
index 974c9d8..4f18ede 100644
--- a/src/Reflection.agda
+++ b/src/Reflection.agda
@@ -4,9 +4,10 @@
module Reflection where
-open import Data.Bool using (Bool); open Data.Bool.Bool
-open import Data.List using (List)
-open import Data.Nat using (ℕ)
+open import Data.Bool as Bool using (Bool); open Bool.Bool
+open import Data.List using (List); open Data.List.List
+open import Data.Nat as Nat using (ℕ)
+open import Function
open import Relation.Binary
open import Relation.Binary.PropositionalEquality
open import Relation.Binary.PropositionalEquality.TrustMe
@@ -27,13 +28,15 @@ private
-- Equality of names is decidable.
-infix 4 _==_ _≟_
+infix 4 _==_ _≟-Name_
-_==_ : Name → Name → Bool
-_==_ = primQNameEquality
+private
+
+ _==_ : Name → Name → Bool
+ _==_ = primQNameEquality
-_≟_ : Decidable {A = Name} _≡_
-s₁ ≟ s₂ with s₁ == s₂
+_≟-Name_ : Decidable {A = Name} _≡_
+s₁ ≟-Name s₂ with s₁ == s₂
... | true = yes trustMe
... | false = no whatever
where postulate whatever : _
@@ -41,14 +44,15 @@ s₁ ≟ s₂ with s₁ == s₂
------------------------------------------------------------------------
-- Terms
--- Is the argument explicit?
+-- Is the argument implicit? (Here true stands for implicit and false
+-- for explicit.)
-Explicit? = Bool
+Implicit? = Bool
-- Arguments.
data Arg A : Set where
- arg : Explicit? → A → Arg A
+ arg : (im? : Implicit?) (x : A) → Arg A
{-# BUILTIN ARG Arg #-}
{-# BUILTIN ARGARG arg #-}
@@ -57,15 +61,15 @@ data Arg A : Set where
data Term : Set where
-- Variable applied to arguments.
- var : ℕ → List (Arg Term) → Term
+ var : (x : ℕ) (args : List (Arg Term)) → Term
-- Constructor applied to arguments.
- con : Name → List (Arg Term) → Term
+ con : (c : Name) (args : List (Arg Term)) → Term
-- Identifier applied to arguments.
- def : Name → List (Arg Term) → Term
+ def : (f : Name) (args : List (Arg Term)) → Term
-- Explicit or implicit λ abstraction.
- lam : Explicit? → Term → Term
+ lam : (im? : Implicit?) (t : Term) → Term
-- Pi-type.
- pi : Arg Term → Term → Term
+ pi : (t₁ : Arg Term) (t₂ : Term) → Term
-- An arbitrary sort (Set, for instance).
sort : Term
-- Anything.
@@ -79,3 +83,140 @@ data Term : Set where
{-# BUILTIN AGDATERMPI pi #-}
{-# BUILTIN AGDATERMSORT sort #-}
{-# BUILTIN AGDATERMUNSUPPORTED unknown #-}
+
+------------------------------------------------------------------------
+-- Term equality is decidable
+
+-- Boring helper functions.
+
+private
+
+ arg₁ : ∀ {A im? im?′} {x x′ : A} →
+ arg im? x ≡ arg im?′ x′ → im? ≡ im?′
+ arg₁ refl = refl
+
+ arg₂ : ∀ {A im? im?′} {x x′ : A} →
+ arg im? x ≡ arg im?′ x′ → x ≡ x′
+ arg₂ refl = refl
+
+ cons₁ : ∀ {A : Set} {x y} {xs ys : List A} → x ∷ xs ≡ y ∷ ys → x ≡ y
+ cons₁ refl = refl
+
+ cons₂ : ∀ {A : Set} {x y} {xs ys : List A} → x ∷ xs ≡ y ∷ ys → xs ≡ ys
+ cons₂ refl = refl
+
+ var₁ : ∀ {x x′ args args′} → var x args ≡ var x′ args′ → x ≡ x′
+ var₁ refl = refl
+
+ var₂ : ∀ {x x′ args args′} → var x args ≡ var x′ args′ → args ≡ args′
+ var₂ refl = refl
+
+ con₁ : ∀ {c c′ args args′} → con c args ≡ con c′ args′ → c ≡ c′
+ con₁ refl = refl
+
+ con₂ : ∀ {c c′ args args′} → con c args ≡ con c′ args′ → args ≡ args′
+ con₂ refl = refl
+
+ def₁ : ∀ {f f′ args args′} → def f args ≡ def f′ args′ → f ≡ f′
+ def₁ refl = refl
+
+ def₂ : ∀ {f f′ args args′} → def f args ≡ def f′ args′ → args ≡ args′
+ def₂ refl = refl
+
+ lam₁ : ∀ {im? im?′ t t′} → lam im? t ≡ lam im?′ t′ → im? ≡ im?′
+ lam₁ refl = refl
+
+ lam₂ : ∀ {im? im?′ t t′} → lam im? t ≡ lam im?′ t′ → t ≡ t′
+ lam₂ refl = refl
+
+ pi₁ : ∀ {t₁ t₁′ t₂ t₂′} → pi t₁ t₂ ≡ pi t₁′ t₂′ → t₁ ≡ t₁′
+ pi₁ refl = refl
+
+ pi₂ : ∀ {t₁ t₁′ t₂ t₂′} → pi t₁ t₂ ≡ pi t₁′ t₂′ → t₂ ≡ t₂′
+ pi₂ refl = refl
+
+mutual
+
+ infix 4 _≟-Arg_ _≟-Args_ _≟_
+
+ _≟-Arg_ : Decidable (_≡_ {A = Arg Term})
+ arg e a ≟-Arg arg e′ a′ with Bool._≟_ e e′ | a ≟ a′
+ arg e a ≟-Arg arg .e .a | yes refl | yes refl = yes refl
+ arg e a ≟-Arg arg e′ a′ | no e≢e′ | _ = no (e≢e′ ∘ arg₁)
+ arg e a ≟-Arg arg e′ a′ | _ | no a≢a′ = no (a≢a′ ∘ arg₂)
+
+ _≟-Args_ : Decidable (_≡_ {A = List (Arg Term)})
+ [] ≟-Args [] = yes refl
+ (a ∷ args) ≟-Args (a′ ∷ args′) with a ≟-Arg a′ | args ≟-Args args′
+ (a ∷ args) ≟-Args (.a ∷ .args) | yes refl | yes refl = yes refl
+ (a ∷ args) ≟-Args (a′ ∷ args′) | no a≢a′ | _ = no (a≢a′ ∘ cons₁)
+ (a ∷ args) ≟-Args (a′ ∷ args′) | _ | no args≢args′ = no (args≢args′ ∘ cons₂)
+ [] ≟-Args (_ ∷ _) = no λ()
+ (_ ∷ _) ≟-Args [] = no λ()
+
+ _≟_ : Decidable (_≡_ {A = Term})
+ var x args ≟ var x′ args′ with Nat._≟_ x x′ | args ≟-Args args′
+ var x args ≟ var .x .args | yes refl | yes refl = yes refl
+ var x args ≟ var x′ args′ | no x≢x′ | _ = no (x≢x′ ∘ var₁)
+ var x args ≟ var x′ args′ | _ | no args≢args′ = no (args≢args′ ∘ var₂)
+ con c args ≟ con c′ args′ with c ≟-Name c′ | args ≟-Args args′
+ con c args ≟ con .c .args | yes refl | yes refl = yes refl
+ con c args ≟ con c′ args′ | no c≢c′ | _ = no (c≢c′ ∘ con₁)
+ con c args ≟ con c′ args′ | _ | no args≢args′ = no (args≢args′ ∘ con₂)
+ def f args ≟ def f′ args′ with f ≟-Name f′ | args ≟-Args args′
+ def f args ≟ def .f .args | yes refl | yes refl = yes refl
+ def f args ≟ def f′ args′ | no f≢f′ | _ = no (f≢f′ ∘ def₁)
+ def f args ≟ def f′ args′ | _ | no args≢args′ = no (args≢args′ ∘ def₂)
+ lam im? t ≟ lam im?′ t′ with Bool._≟_ im? im?′ | t ≟ t′
+ lam im? t ≟ lam .im? .t | yes refl | yes refl = yes refl
+ lam im? t ≟ lam im?′ t′ | no im?≢im?′ | _ = no (im?≢im?′ ∘ lam₁)
+ lam im? t ≟ lam im?′ t′ | _ | no t≢t′ = no (t≢t′ ∘ lam₂)
+ pi t₁ t₂ ≟ pi t₁′ t₂′ with t₁ ≟-Arg t₁′ | t₂ ≟ t₂′
+ pi t₁ t₂ ≟ pi .t₁ .t₂ | yes refl | yes refl = yes refl
+ pi t₁ t₂ ≟ pi t₁′ t₂′ | no t₁≢t₁′ | _ = no (t₁≢t₁′ ∘ pi₁)
+ pi t₁ t₂ ≟ pi t₁′ t₂′ | _ | no t₂≢t₂′ = no (t₂≢t₂′ ∘ pi₂)
+ sort ≟ sort = yes refl
+ unknown ≟ unknown = yes refl
+
+ var x args ≟ con c args′ = no λ()
+ var x args ≟ def f args′ = no λ()
+ var x args ≟ lam im? t = no λ()
+ var x args ≟ pi t₁ t₂ = no λ()
+ var x args ≟ sort = no λ()
+ var x args ≟ unknown = no λ()
+ con c args ≟ var x args′ = no λ()
+ con c args ≟ def f args′ = no λ()
+ con c args ≟ lam im? t = no λ()
+ con c args ≟ pi t₁ t₂ = no λ()
+ con c args ≟ sort = no λ()
+ con c args ≟ unknown = no λ()
+ def f args ≟ var x args′ = no λ()
+ def f args ≟ con c args′ = no λ()
+ def f args ≟ lam im? t = no λ()
+ def f args ≟ pi t₁ t₂ = no λ()
+ def f args ≟ sort = no λ()
+ def f args ≟ unknown = no λ()
+ lam im? t ≟ var x args = no λ()
+ lam im? t ≟ con c args = no λ()
+ lam im? t ≟ def f args = no λ()
+ lam im? t ≟ pi t₁ t₂ = no λ()
+ lam im? t ≟ sort = no λ()
+ lam im? t ≟ unknown = no λ()
+ pi t₁ t₂ ≟ var x args = no λ()
+ pi t₁ t₂ ≟ con c args = no λ()
+ pi t₁ t₂ ≟ def f args = no λ()
+ pi t₁ t₂ ≟ lam im? t = no λ()
+ pi t₁ t₂ ≟ sort = no λ()
+ pi t₁ t₂ ≟ unknown = no λ()
+ sort ≟ var x args = no λ()
+ sort ≟ con c args = no λ()
+ sort ≟ def f args = no λ()
+ sort ≟ lam im? t = no λ()
+ sort ≟ pi t₁ t₂ = no λ()
+ sort ≟ unknown = no λ()
+ unknown ≟ var x args = no λ()
+ unknown ≟ con c args = no λ()
+ unknown ≟ def f args = no λ()
+ unknown ≟ lam im? t = no λ()
+ unknown ≟ pi t₁ t₂ = no λ()
+ unknown ≟ sort = no λ()
diff --git a/src/Relation/Binary/List/Pointwise.agda b/src/Relation/Binary/List/Pointwise.agda
index d0314f0..04a4fe0 100644
--- a/src/Relation/Binary/List/Pointwise.agda
+++ b/src/Relation/Binary/List/Pointwise.agda
@@ -2,6 +2,8 @@
-- Pointwise lifting of relations to lists
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Relation.Binary.List.Pointwise where
open import Function
@@ -13,47 +15,54 @@ open import Relation.Binary renaming (Rel to Rel₂)
open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
private
- module Dummy {A : Set} where
+ module Dummy {a} {A : Set a} where
infixr 5 _∷_
- data Rel (_∼_ : Rel₂ A zero) : List A → List A → Set where
+ data Rel {ℓ} (_∼_ : Rel₂ A ℓ) : List A → List A → Set (ℓ ⊔ a) where
[] : Rel _∼_ [] []
_∷_ : ∀ {x xs y ys} (x∼y : x ∼ y) (xs∼ys : Rel _∼_ xs ys) →
Rel _∼_ (x ∷ xs) (y ∷ ys)
- head : ∀ {_∼_ x y xs ys} → Rel _∼_ (x ∷ xs) (y ∷ ys) → x ∼ y
+ head : ∀ {ℓ} {_∼_ : Rel₂ A ℓ} {x y xs ys} →
+ Rel _∼_ (x ∷ xs) (y ∷ ys) → x ∼ y
head (x∼y ∷ xs∼ys) = x∼y
- tail : ∀ {_∼_ x y xs ys} → Rel _∼_ (x ∷ xs) (y ∷ ys) → Rel _∼_ xs ys
+ tail : ∀ {ℓ} {_∼_ : Rel₂ A ℓ} {x y xs ys} →
+ Rel _∼_ (x ∷ xs) (y ∷ ys) → Rel _∼_ xs ys
tail (x∼y ∷ xs∼ys) = xs∼ys
- reflexive : ∀ {_≈_ _∼_} → _≈_ ⇒ _∼_ → Rel _≈_ ⇒ Rel _∼_
+ reflexive : ∀ {ℓ₁ ℓ₂} {_≈_ : Rel₂ A ℓ₁} {_∼_ : Rel₂ A ℓ₂} →
+ _≈_ ⇒ _∼_ → Rel _≈_ ⇒ Rel _∼_
reflexive ≈⇒∼ [] = []
reflexive ≈⇒∼ (x≈y ∷ xs≈ys) = ≈⇒∼ x≈y ∷ reflexive ≈⇒∼ xs≈ys
- refl : ∀ {_∼_} → Reflexive _∼_ → Reflexive (Rel _∼_)
+ refl : ∀ {ℓ} {_∼_ : Rel₂ A ℓ} →
+ Reflexive _∼_ → Reflexive (Rel _∼_)
refl rfl {[]} = []
refl rfl {x ∷ xs} = rfl ∷ refl rfl
- symmetric : ∀ {_∼_} → Symmetric _∼_ → Symmetric (Rel _∼_)
+ symmetric : ∀ {ℓ} {_∼_ : Rel₂ A ℓ} →
+ Symmetric _∼_ → Symmetric (Rel _∼_)
symmetric sym [] = []
symmetric sym (x∼y ∷ xs∼ys) = sym x∼y ∷ symmetric sym xs∼ys
- transitive : ∀ {_∼_} → Transitive _∼_ → Transitive (Rel _∼_)
+ transitive : ∀ {ℓ} {_∼_ : Rel₂ A ℓ} →
+ Transitive _∼_ → Transitive (Rel _∼_)
transitive trans [] [] = []
transitive trans (x∼y ∷ xs∼ys) (y∼z ∷ ys∼zs) =
trans x∼y y∼z ∷ transitive trans xs∼ys ys∼zs
- antisymmetric : ∀ {_≈_ _≤_} → Antisymmetric _≈_ _≤_ →
+ antisymmetric : ∀ {ℓ₁ ℓ₂} {_≈_ : Rel₂ A ℓ₁} {_≤_ : Rel₂ A ℓ₂} →
+ Antisymmetric _≈_ _≤_ →
Antisymmetric (Rel _≈_) (Rel _≤_)
antisymmetric antisym [] [] = []
antisymmetric antisym (x∼y ∷ xs∼ys) (y∼x ∷ ys∼xs) =
antisym x∼y y∼x ∷ antisymmetric antisym xs∼ys ys∼xs
- respects₂ : ∀ {_≈_ _∼_} →
+ respects₂ : ∀ {ℓ₁ ℓ₂} {_≈_ : Rel₂ A ℓ₁} {_∼_ : Rel₂ A ℓ₂} →
_∼_ Respects₂ _≈_ → (Rel _∼_) Respects₂ (Rel _≈_)
- respects₂ {_≈_} {_∼_} resp =
+ respects₂ {_} {_} {_≈_} {_∼_} resp =
(λ {xs} {ys} {zs} → resp¹ {xs} {ys} {zs}) ,
(λ {xs} {ys} {zs} → resp² {xs} {ys} {zs})
where
@@ -67,7 +76,8 @@ private
resp² (x≈y ∷ xs≈ys) (x∼z ∷ xs∼zs) =
proj₂ resp x≈y x∼z ∷ resp² xs≈ys xs∼zs
- decidable : ∀ {_∼_} → Decidable _∼_ → Decidable (Rel _∼_)
+ decidable : ∀ {ℓ} {_∼_ : Rel₂ A ℓ} →
+ Decidable _∼_ → Decidable (Rel _∼_)
decidable dec [] [] = yes []
decidable dec [] (y ∷ ys) = no (λ ())
decidable dec (x ∷ xs) [] = no (λ ())
@@ -77,14 +87,15 @@ private
... | no ¬xs∼ys = no (¬xs∼ys ∘ tail)
... | yes xs∼ys = yes (x∼y ∷ xs∼ys)
- isEquivalence : ∀ {_≈_} → IsEquivalence _≈_ → IsEquivalence (Rel _≈_)
+ isEquivalence : ∀ {ℓ} {_≈_ : Rel₂ A ℓ} →
+ IsEquivalence _≈_ → IsEquivalence (Rel _≈_)
isEquivalence eq = record
{ refl = refl Eq.refl
; sym = symmetric Eq.sym
; trans = transitive Eq.trans
} where module Eq = IsEquivalence eq
- isPreorder : ∀ {_≈_ _∼_} →
+ isPreorder : ∀ {ℓ₁ ℓ₂} {_≈_ : Rel₂ A ℓ₁} {_∼_ : Rel₂ A ℓ₂} →
IsPreorder _≈_ _∼_ → IsPreorder (Rel _≈_) (Rel _∼_)
isPreorder pre = record
{ isEquivalence = isEquivalence Pre.isEquivalence
@@ -92,14 +103,15 @@ private
; trans = transitive Pre.trans
} where module Pre = IsPreorder pre
- isDecEquivalence : ∀ {_≈_} → IsDecEquivalence _≈_ →
+ isDecEquivalence : ∀ {ℓ} {_≈_ : Rel₂ A ℓ} → IsDecEquivalence _≈_ →
IsDecEquivalence (Rel _≈_)
isDecEquivalence eq = record
{ isEquivalence = isEquivalence Dec.isEquivalence
; _≟_ = decidable Dec._≟_
} where module Dec = IsDecEquivalence eq
- isPartialOrder : ∀ {_≈_ _≤_} → IsPartialOrder _≈_ _≤_ →
+ isPartialOrder : ∀ {ℓ₁ ℓ₂} {_≈_ : Rel₂ A ℓ₁} {_≤_ : Rel₂ A ℓ₂} →
+ IsPartialOrder _≈_ _≤_ →
IsPartialOrder (Rel _≈_) (Rel _≤_)
isPartialOrder po = record
{ isPreorder = isPreorder PO.isPreorder
@@ -118,22 +130,22 @@ private
open Dummy public
-preorder : Preorder _ _ _ → Preorder _ _ _
+preorder : ∀ {p₁ p₂ p₃} → Preorder p₁ p₂ p₃ → Preorder _ _ _
preorder p = record
{ isPreorder = isPreorder (Preorder.isPreorder p)
}
-setoid : Setoid _ _ → Setoid _ _
+setoid : ∀ {c ℓ} → Setoid c ℓ → Setoid _ _
setoid s = record
{ isEquivalence = isEquivalence (Setoid.isEquivalence s)
}
-decSetoid : DecSetoid _ _ → DecSetoid _ _
+decSetoid : ∀ {c ℓ} → DecSetoid c ℓ → DecSetoid _ _
decSetoid d = record
{ isDecEquivalence = isDecEquivalence (DecSetoid.isDecEquivalence d)
}
-poset : Poset _ _ _ → Poset _ _ _
+poset : ∀ {c ℓ₁ ℓ₂} → Poset c ℓ₁ ℓ₂ → Poset _ _ _
poset p = record
{ isPartialOrder = isPartialOrder (Poset.isPartialOrder p)
}
diff --git a/src/Relation/Binary/Product/StrictLex.agda b/src/Relation/Binary/Product/StrictLex.agda
index 3a3542b..42755b6 100644
--- a/src/Relation/Binary/Product/StrictLex.agda
+++ b/src/Relation/Binary/Product/StrictLex.agda
@@ -106,7 +106,7 @@ private
×-≈-respects₂ :
∀ {_≈₁_ _<₁_} → IsEquivalence _≈₁_ → _<₁_ Respects₂ _≈₁_ →
- ∀ {_≈₂_ _<₂_} → _<₂_ Respects₂ _≈₂_ →
+ {_≈₂_ _<₂_ : Rel A₂ ℓ₂} → _<₂_ Respects₂ _≈₂_ →
(×-Lex _≈₁_ _<₁_ _<₂_) Respects₂ (_≈₁_ ×-Rel _≈₂_)
×-≈-respects₂ {_≈₁_ = _≈₁_} {_<₁_ = _<₁_} eq₁ resp₁
{_≈₂_ = _≈₂_} {_<₂_ = _<₂_} resp₂ =
@@ -149,9 +149,10 @@ private
... | inj₁ x₁<y₁ = inj₁ (inj₁ x₁<y₁)
... | inj₂ x₁>y₁ = inj₂ (inj₁ x₁>y₁)
- ×-compare : ∀ {_≈₁_ _<₁_} → Symmetric _≈₁_ → Trichotomous _≈₁_ _<₁_ →
- ∀ {_≈₂_ _<₂_} → Trichotomous _≈₂_ _<₂_ →
- Trichotomous (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_)
+ ×-compare :
+ {_≈₁_ _<₁_ : Rel A₁ ℓ₁} → Symmetric _≈₁_ → Trichotomous _≈₁_ _<₁_ →
+ {_≈₂_ _<₂_ : Rel A₂ ℓ₂} → Trichotomous _≈₂_ _<₂_ →
+ Trichotomous (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_)
×-compare {_≈₁_} {_<₁_} sym₁ compare₁ {_≈₂_} {_<₂_} compare₂ = cmp
where
cmp″ : ∀ {x₁ y₁ x₂ y₂} →
diff --git a/src/Relation/Binary/Reflection.agda b/src/Relation/Binary/Reflection.agda
index 149f897..5f54c6a 100644
--- a/src/Relation/Binary/Reflection.agda
+++ b/src/Relation/Binary/Reflection.agda
@@ -6,11 +6,14 @@
{-# OPTIONS --universe-polymorphism #-}
open import Data.Fin
-open import Function
open import Data.Nat
open import Data.Vec as Vec
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence using (module Equivalent)
open import Level
open import Relation.Binary
+import Relation.Binary.PropositionalEquality as P
-- Think of the parameters as follows:
--
@@ -72,13 +75,26 @@ solve : ∀ n (f : N-ary n (Expr n) (Expr n × Expr n)) →
Eqʰ n _≈_ (curryⁿ ⟦ proj₁ (close n f) ⇓⟧) (curryⁿ ⟦ proj₂ (close n f) ⇓⟧) →
Eq n _≈_ (curryⁿ ⟦ proj₁ (close n f) ⟧) (curryⁿ ⟦ proj₂ (close n f) ⟧)
solve n f hyp =
- curryⁿ-cong ⟦ proj₁ (close n f) ⟧ ⟦ proj₂ (close n f) ⟧
+ curryⁿ-cong _≈_ ⟦ proj₁ (close n f) ⟧ ⟦ proj₂ (close n f) ⟧
(λ ρ → prove ρ (proj₁ (close n f)) (proj₂ (close n f))
- (curryⁿ-cong⁻¹ ⟦ proj₁ (close n f) ⇓⟧ ⟦ proj₂ (close n f) ⇓⟧
- (Eqʰ-to-Eq n hyp) ρ))
+ (curryⁿ-cong⁻¹ _≈_
+ ⟦ proj₁ (close n f) ⇓⟧ ⟦ proj₂ (close n f) ⇓⟧
+ (Eqʰ-to-Eq n _≈_ hyp) ρ))
+
+-- A variant of solve which does not require that the normal form
+-- equality is proved for an arbitrary environment.
+
+solve₁ : ∀ n (f : N-ary n (Expr n) (Expr n × Expr n)) →
+ ∀ⁿ n (curryⁿ λ ρ →
+ ⟦ proj₁ (close n f) ⇓⟧ ρ ≈ ⟦ proj₂ (close n f) ⇓⟧ ρ →
+ ⟦ proj₁ (close n f) ⟧ ρ ≈ ⟦ proj₂ (close n f) ⟧ ρ)
+solve₁ n f =
+ Equivalent.from (uncurry-∀ⁿ n) ⟨$⟩ λ ρ →
+ P.subst id (P.sym (left-inverse (λ _ → _ ≈ _ → _ ≈ _) ρ))
+ (prove ρ (proj₁ (close n f)) (proj₂ (close n f)))
--- A variant of _,_ which is intended to make uses of solve look a
--- bit nicer.
+-- A variant of _,_ which is intended to make uses of solve and solve₁
+-- look a bit nicer.
infix 4 _⊜_
diff --git a/src/Relation/Binary/Sigma/Pointwise.agda b/src/Relation/Binary/Sigma/Pointwise.agda
index dfac94e..dadec84 100644
--- a/src/Relation/Binary/Sigma/Pointwise.agda
+++ b/src/Relation/Binary/Sigma/Pointwise.agda
@@ -9,7 +9,8 @@ module Relation.Binary.Sigma.Pointwise where
open import Data.Product as Prod
open import Level
open import Function
-import Function.Equality as F
+open import Function.Equality as F using (_⟶_; _⟨$⟩_)
+ renaming (_∘_ to _⊚_)
open import Function.Equivalence as Eq
using (Equivalent; _⇔_; module Equivalent)
renaming (_∘_ to _⟨∘⟩_)
@@ -18,6 +19,7 @@ open import Function.Inverse as Inv
renaming (_∘_ to _⟪∘⟫_)
open import Function.LeftInverse
using (_LeftInverseOf_; _RightInverseOf_)
+open import Function.Surjection using (_↠_; module Surjection)
import Relation.Binary as B
open import Relation.Binary.Indexed as I using (_at_)
import Relation.Binary.HeterogeneousEquality as H
@@ -110,45 +112,101 @@ Rel⇿≡ {a} {b} {A} {B} = record
-- Equivalences and inverses are also preserved
equivalent :
- ∀ {i} {I : Set i}
- {f₁ f₂ t₁ t₂} {From : I.Setoid I f₁ f₂} {To : I.Setoid I t₁ t₂} →
- (∀ {i} → Equivalent (From at i) (To at i)) →
- Equivalent (setoid (P.setoid I) From) (setoid (P.setoid I) To)
-equivalent {I = I} {From = F} {T} F⇔T = record
+ ∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : I.Setoid A₁ b₁ b₁′} {B₂ : I.Setoid A₂ b₂ b₂′}
+ (A₁⇔A₂ : A₁ ⇔ A₂) →
+ (∀ {x} → B₁ at x ⟶ B₂ at (Equivalent.to A₁⇔A₂ ⟨$⟩ x)) →
+ (∀ {y} → B₂ at y ⟶ B₁ at (Equivalent.from A₁⇔A₂ ⟨$⟩ y)) →
+ Equivalent (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+equivalent {A₁ = A₁} {A₂} {B₁} {B₂} A₁⇔A₂ B-to B-from = record
{ to = record { _⟨$⟩_ = to; cong = to-cong }
; from = record { _⟨$⟩_ = from; cong = from-cong }
}
where
- open B.Setoid (setoid (P.setoid I) F) using () renaming (_≈_ to _≈F_)
- open B.Setoid (setoid (P.setoid I) T) using () renaming (_≈_ to _≈T_)
+ open B.Setoid (setoid (P.setoid A₁) B₁)
+ using () renaming (_≈_ to _≈₁_)
+ open B.Setoid (setoid (P.setoid A₂) B₂)
+ using () renaming (_≈_ to _≈₂_)
open B using (_=[_]⇒_)
- to = Prod.map id (F._⟨$⟩_ (Equivalent.to F⇔T))
+ to = Prod.map (_⟨$⟩_ (Equivalent.to A₁⇔A₂)) (_⟨$⟩_ B-to)
- to-cong : _≈F_ =[ to ]⇒ _≈T_
- to-cong (P.refl , ∼) = (P.refl , F.cong (Equivalent.to F⇔T) ∼)
+ to-cong : _≈₁_ =[ to ]⇒ _≈₂_
+ to-cong (P.refl , ∼) = (P.refl , F.cong B-to ∼)
- from = Prod.map id (F._⟨$⟩_ (Equivalent.from F⇔T))
+ from = Prod.map (_⟨$⟩_ (Equivalent.from A₁⇔A₂)) (_⟨$⟩_ B-from)
- from-cong : _≈T_ =[ from ]⇒ _≈F_
- from-cong (P.refl , ∼) = (P.refl , F.cong (Equivalent.from F⇔T) ∼)
+ from-cong : _≈₂_ =[ from ]⇒ _≈₁_
+ from-cong (P.refl , ∼) = (P.refl , F.cong B-from ∼)
-⇔ : ∀ {a b₁ b₂} {A : Set a} {B₁ : A → Set b₁} {B₂ : A → Set b₂} →
- (∀ {x} → B₁ x ⇔ B₂ x) → Σ A B₁ ⇔ Σ A B₂
-⇔ {B₁ = B₁} {B₂} B₁⇔B₂ =
+equivalent′ :
+ ∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : I.Setoid A₁ b₁ b₁′} (B₂ : I.Setoid A₂ b₂ b₂′)
+ (A₁↠A₂ : A₁ ↠ A₂) →
+ (∀ {x} → Equivalent (B₁ at x) (B₂ at (Surjection.to A₁↠A₂ ⟨$⟩ x))) →
+ Equivalent (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+equivalent′ {B₁ = B₁} B₂ A₁↠A₂ B₁⇔B₂ =
+ equivalent (Surjection.equivalent A₁↠A₂) B-to B-from
+ where
+ B-to : ∀ {x} → B₁ at x ⟶ B₂ at (Surjection.to A₁↠A₂ ⟨$⟩ x)
+ B-to = Equivalent.to B₁⇔B₂
+
+ subst-cong : ∀ {i a p} {I : Set i} {A : I → Set a}
+ (P : ∀ {i} → A i → A i → Set p) {i i′} {x y : A i}
+ (i≡i′ : P._≡_ i i′) →
+ P x y → P (P.subst A i≡i′ x) (P.subst A i≡i′ y)
+ subst-cong P P.refl p = p
+
+ B-from : ∀ {y} → B₂ at y ⟶ B₁ at (Surjection.from A₁↠A₂ ⟨$⟩ y)
+ B-from = record
+ { _⟨$⟩_ = λ x → Equivalent.from B₁⇔B₂ ⟨$⟩
+ P.subst (I.Setoid.Carrier B₂)
+ (P.sym $ Surjection.right-inverse-of A₁↠A₂ _)
+ x
+ ; cong = F.cong (Equivalent.from B₁⇔B₂) ∘
+ subst-cong (λ {x} → I.Setoid._≈_ B₂ {x} {x})
+ (P.sym (Surjection.right-inverse-of A₁↠A₂ _))
+ }
+
+⇔ : ∀ {a₁ a₂ b₁ b₂}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
+ (A₁⇔A₂ : A₁ ⇔ A₂) →
+ (∀ {x} → B₁ x → B₂ (Equivalent.to A₁⇔A₂ ⟨$⟩ x)) →
+ (∀ {y} → B₂ y → B₁ (Equivalent.from A₁⇔A₂ ⟨$⟩ y)) →
+ Σ A₁ B₁ ⇔ Σ A₂ B₂
+⇔ {B₁ = B₁} {B₂} A₁⇔A₂ B-to B-from =
Inverse.equivalent (Rel⇿≡ {B = B₂}) ⟨∘⟩
- equivalent (λ {x} →
- Inverse.equivalent (H.≡⇿≅ B₂) ⟨∘⟩
- B₁⇔B₂ {x} ⟨∘⟩
- Inverse.equivalent (Inv.sym (H.≡⇿≅ B₁))) ⟨∘⟩
+ equivalent A₁⇔A₂
+ (Inverse.to (H.≡⇿≅ B₂) ⊚ P.→-to-⟶ B-to ⊚ Inverse.from (H.≡⇿≅ B₁))
+ (Inverse.to (H.≡⇿≅ B₁) ⊚ P.→-to-⟶ B-from ⊚ Inverse.from (H.≡⇿≅ B₂))
+ ⟨∘⟩
+ Eq.sym (Inverse.equivalent (Rel⇿≡ {B = B₁}))
+
+⇔′ : ∀ {a₁ a₂ b₁ b₂}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
+ (A₁↠A₂ : A₁ ↠ A₂) →
+ (∀ {x} → _⇔_ (B₁ x) (B₂ (Surjection.to A₁↠A₂ ⟨$⟩ x))) →
+ _⇔_ (Σ A₁ B₁) (Σ A₂ B₂)
+⇔′ {B₁ = B₁} {B₂} A₁↠A₂ B₁⇔B₂ =
+ Inverse.equivalent (Rel⇿≡ {B = B₂}) ⟨∘⟩
+ equivalent′ (H.indexedSetoid B₂) A₁↠A₂
+ (λ {x} → Inverse.equivalent (H.≡⇿≅ B₂) ⟨∘⟩
+ B₁⇔B₂ {x} ⟨∘⟩
+ Inverse.equivalent (Inv.sym (H.≡⇿≅ B₁))) ⟨∘⟩
Eq.sym (Inverse.equivalent (Rel⇿≡ {B = B₁}))
inverse :
- ∀ {i} {I : Set i}
- {f₁ f₂ t₁ t₂} {From : I.Setoid I f₁ f₂} {To : I.Setoid I t₁ t₂} →
- (∀ {i} → Inverse (From at i) (To at i)) →
- Inverse (setoid (P.setoid I) From) (setoid (P.setoid I) To)
-inverse {I = I} {From = F} {T} F⇿T = record
+ ∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : I.Setoid A₁ b₁ b₁′} (B₂ : I.Setoid A₂ b₂ b₂′) →
+ (A₁⇿A₂ : A₁ ⇿ A₂) →
+ (∀ {x} → Inverse (B₁ at x) (B₂ at (Inverse.to A₁⇿A₂ ⟨$⟩ x))) →
+ Inverse (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+inverse {A₁ = A₁} {A₂} {B₁} B₂ A₁⇿A₂ B₁⇿B₂ = record
{ to = Equivalent.to eq
; from = Equivalent.from eq
; inverse-of = record
@@ -157,23 +215,59 @@ inverse {I = I} {From = F} {T} F⇿T = record
}
}
where
- eq = equivalent (Inverse.equivalent F⇿T)
+ eq = equivalent′ B₂ (Inverse.surjection A₁⇿A₂)
+ (Inverse.equivalent B₁⇿B₂)
left : Equivalent.from eq LeftInverseOf Equivalent.to eq
- left (x , y) = (P.refl , Inverse.left-inverse-of F⇿T y)
+ left (x , y) =
+ Inverse.left-inverse-of A₁⇿A₂ x ,
+ I.Setoid.trans B₁
+ (lemma (P.sym (Inverse.left-inverse-of A₁⇿A₂ x))
+ (P.sym (Inverse.right-inverse-of A₁⇿A₂
+ (Inverse.to A₁⇿A₂ ⟨$⟩ x))))
+ (Inverse.left-inverse-of B₁⇿B₂ y)
+ where
+ lemma :
+ ∀ {x x′ y} → P._≡_ x x′ →
+ (eq : P._≡_ (Inverse.to A₁⇿A₂ ⟨$⟩ x) (Inverse.to A₁⇿A₂ ⟨$⟩ x′)) →
+ I.Setoid._≈_ B₁
+ (Inverse.from B₁⇿B₂ ⟨$⟩ P.subst (I.Setoid.Carrier B₂) eq y)
+ (Inverse.from B₁⇿B₂ ⟨$⟩ y)
+ lemma P.refl P.refl = I.Setoid.refl B₁
right : Equivalent.from eq RightInverseOf Equivalent.to eq
- right (x , y) = (P.refl , Inverse.right-inverse-of F⇿T y)
+ right (x , y) =
+ Inverse.right-inverse-of A₁⇿A₂ x ,
+ I.Setoid.trans B₂
+ (Inverse.right-inverse-of B₁⇿B₂
+ (P.subst (I.Setoid.Carrier B₂) p y))
+ (lemma p)
+ where
+ p = P.sym $ Inverse.right-inverse-of A₁⇿A₂ x
+
+ lemma : ∀ {x x′ y} (eq : P._≡_ x x′) →
+ I.Setoid._≈_ B₂
+ (P.subst (I.Setoid.Carrier B₂) eq y)
+ y
+ lemma P.refl = I.Setoid.refl B₂
-⇿ : ∀ {a b₁ b₂} {A : Set a} {B₁ : A → Set b₁} {B₂ : A → Set b₂} →
- (∀ {x} → B₁ x ⇿ B₂ x) → Σ A B₁ ⇿ Σ A B₂
-⇿ {B₁ = B₁} {B₂} B₁⇿B₂ =
+⇿ : ∀ {a₁ a₂ b₁ b₂}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
+ (A₁⇿A₂ : A₁ ⇿ A₂) →
+ (∀ {x} → B₁ x ⇿ B₂ (Inverse.to A₁⇿A₂ ⟨$⟩ x)) →
+ Σ A₁ B₁ ⇿ Σ A₂ B₂
+⇿ {B₁ = B₁} {B₂} A₁⇿A₂ B₁⇿B₂ =
Rel⇿≡ {B = B₂} ⟪∘⟫
- inverse (λ {x} → H.≡⇿≅ B₂ ⟪∘⟫ B₁⇿B₂ {x} ⟪∘⟫ Inv.sym (H.≡⇿≅ B₁)) ⟪∘⟫
+ inverse (H.indexedSetoid B₂) A₁⇿A₂
+ (λ {x} → H.≡⇿≅ B₂ ⟪∘⟫ B₁⇿B₂ {x} ⟪∘⟫ Inv.sym (H.≡⇿≅ B₁)) ⟪∘⟫
Inv.sym (Rel⇿≡ {B = B₁})
-cong : ∀ {k a b₁ b₂} {A : Set a} {B₁ : A → Set b₁} {B₂ : A → Set b₂} →
- (∀ {x} → Isomorphism k (B₁ x) (B₂ x)) →
- Isomorphism k (Σ A B₁) (Σ A B₂)
-cong {k = Inv.equivalent} = ⇔
+cong : ∀ {k a₁ a₂ b₁ b₂}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
+ (A₁⇿A₂ : _⇿_ A₁ A₂) →
+ (∀ {x} → Isomorphism k (B₁ x) (B₂ (Inverse.to A₁⇿A₂ ⟨$⟩ x))) →
+ Isomorphism k (Σ A₁ B₁) (Σ A₂ B₂)
+cong {k = Inv.equivalent} = ⇔′ ∘ Inverse.surjection
cong {k = Inv.inverse} = ⇿
diff --git a/src/Relation/Binary/Vec/Pointwise.agda b/src/Relation/Binary/Vec/Pointwise.agda
new file mode 100644
index 0000000..9613aa4
--- /dev/null
+++ b/src/Relation/Binary/Vec/Pointwise.agda
@@ -0,0 +1,204 @@
+------------------------------------------------------------------------
+-- Pointwise lifting of relations to vectors
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Relation.Binary.Vec.Pointwise where
+
+open import Category.Applicative.Indexed
+open import Data.Fin
+open import Data.Plus as Plus hiding (equivalent; map)
+open import Data.Vec as Vec hiding ([_]; map)
+import Data.Vec.Properties as VecProp
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence as Equiv
+ using (_⇔_; module Equivalent)
+open import Level
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+open import Relation.Nullary
+
+private
+ module Dummy {a} {A : Set a} where
+
+ -- Functional definition.
+
+ record Pointwise (_∼_ : Rel A a) {n} (xs ys : Vec A n) : Set a where
+ constructor ext
+ field app : ∀ i → lookup i xs ∼ lookup i ys
+
+ -- Inductive definition.
+
+ infixr 5 _∷_
+
+ data Pointwise′ (_∼_ : Rel A a) :
+ ∀ {n} (xs ys : Vec A n) → Set a where
+ [] : Pointwise′ _∼_ [] []
+ _∷_ : ∀ {n x y} {xs ys : Vec A n}
+ (x∼y : x ∼ y) (xs∼ys : Pointwise′ _∼_ xs ys) →
+ Pointwise′ _∼_ (x ∷ xs) (y ∷ ys)
+
+ -- The two definitions are equivalent.
+
+ equivalent : ∀ {_∼_ : Rel A a} {n} {xs ys : Vec A n} →
+ Pointwise _∼_ xs ys ⇔ Pointwise′ _∼_ xs ys
+ equivalent {_∼_} = Equiv.equivalent (to _ _) from
+ where
+ to : ∀ {n} (xs ys : Vec A n) →
+ Pointwise _∼_ xs ys → Pointwise′ _∼_ xs ys
+ to [] [] xs∼ys = []
+ to (x ∷ xs) (y ∷ ys) xs∼ys =
+ Pointwise.app xs∼ys zero ∷
+ to xs ys (ext (Pointwise.app xs∼ys ∘ suc))
+
+ nil : Pointwise _∼_ [] []
+ nil = ext λ()
+
+ cons : ∀ {n x y} {xs ys : Vec A n} →
+ x ∼ y → Pointwise _∼_ xs ys → Pointwise _∼_ (x ∷ xs) (y ∷ ys)
+ cons {x = x} {y} {xs} {ys} x∼y xs∼ys = ext helper
+ where
+ helper : ∀ i → lookup i (x ∷ xs) ∼ lookup i (y ∷ ys)
+ helper zero = x∼y
+ helper (suc i) = Pointwise.app xs∼ys i
+
+ from : ∀ {n} {xs ys : Vec A n} →
+ Pointwise′ _∼_ xs ys → Pointwise _∼_ xs ys
+ from [] = nil
+ from (x∼y ∷ xs∼ys) = cons x∼y (from xs∼ys)
+
+ -- Pointwise preserves reflexivity.
+
+ refl : ∀ {_∼_ : Rel A a} {n} →
+ Reflexive _∼_ → Reflexive (Pointwise _∼_ {n = n})
+ refl rfl = ext (λ _ → rfl)
+
+ -- Pointwise preserves symmetry.
+
+ sym : ∀ {_∼_ : Rel A a} {n} →
+ Symmetric _∼_ → Symmetric (Pointwise _∼_ {n = n})
+ sym sm xs∼ys = ext λ i → sm (Pointwise.app xs∼ys i)
+
+ -- Pointwise preserves transitivity.
+
+ trans : ∀ {_∼_ : Rel A a} {n} →
+ Transitive _∼_ → Transitive (Pointwise _∼_ {n = n})
+ trans trns xs∼ys ys∼zs = ext λ i →
+ trns (Pointwise.app xs∼ys i) (Pointwise.app ys∼zs i)
+
+ -- Pointwise preserves equivalences.
+
+ isEquivalence :
+ ∀ {_∼_ : Rel A a} {n} →
+ IsEquivalence _∼_ → IsEquivalence (Pointwise _∼_ {n = n})
+ isEquivalence equiv = record
+ { refl = refl (IsEquivalence.refl equiv)
+ ; sym = sym (IsEquivalence.sym equiv)
+ ; trans = trans (IsEquivalence.trans equiv)
+ }
+
+ -- Pointwise _≡_ is equivalent to _≡_.
+
+ Pointwise-≡ : ∀ {n} {xs ys : Vec A n} →
+ Pointwise _≡_ xs ys ⇔ xs ≡ ys
+ Pointwise-≡ =
+ Equiv.equivalent
+ (to ∘ _⟨$⟩_ (Equivalent.to equivalent))
+ (λ xs≡ys → P.subst (Pointwise _≡_ _) xs≡ys (refl P.refl))
+ where
+ to : ∀ {n} {xs ys : Vec A n} → Pointwise′ _≡_ xs ys → xs ≡ ys
+ to [] = P.refl
+ to (P.refl ∷ xs∼ys) = P.cong (_∷_ _) $ to xs∼ys
+
+ -- Pointwise and Plus commute when the underlying relation is
+ -- reflexive.
+
+ ⁺∙⇒∙⁺ : ∀ {_∼_ : Rel A a} {n} {xs ys : Vec A n} →
+ Plus (Pointwise _∼_) xs ys → Pointwise (Plus _∼_) xs ys
+ ⁺∙⇒∙⁺ [ ρ≈ρ′ ] = ext (λ x → [ Pointwise.app ρ≈ρ′ x ])
+ ⁺∙⇒∙⁺ (ρ ∼⁺⟨ ρ≈ρ′ ⟩ ρ′≈ρ″) =
+ ext (λ x → _ ∼⁺⟨ Pointwise.app (⁺∙⇒∙⁺ ρ≈ρ′ ) x ⟩
+ Pointwise.app (⁺∙⇒∙⁺ ρ′≈ρ″) x)
+
+ ∙⁺⇒⁺∙ : ∀ {_∼_ : Rel A a} {n} {xs ys : Vec A n} →
+ Reflexive _∼_ →
+ Pointwise (Plus _∼_) xs ys → Plus (Pointwise _∼_) xs ys
+ ∙⁺⇒⁺∙ {_∼_} x∼x =
+ Plus.map (_⟨$⟩_ (Equivalent.from equivalent)) ∘
+ helper ∘
+ _⟨$⟩_ (Equivalent.to equivalent)
+ where
+ helper : ∀ {n} {xs ys : Vec A n} →
+ Pointwise′ (Plus _∼_) xs ys → Plus (Pointwise′ _∼_) xs ys
+ helper [] = [ [] ]
+ helper (_∷_ {x = x} {y = y} {xs = xs} {ys = ys} x∼y xs∼ys) =
+ x ∷ xs ∼⁺⟨ Plus.map (λ x∼y → x∼y ∷ xs∼xs) x∼y ⟩
+ y ∷ xs ∼⁺⟨ Plus.map (λ xs∼ys → x∼x ∷ xs∼ys) (helper xs∼ys) ⟩∎
+ y ∷ ys ∎
+ where
+ xs∼xs : Pointwise′ _∼_ xs xs
+ xs∼xs = Equivalent.to equivalent ⟨$⟩ refl x∼x
+
+open Dummy public
+
+-- Note that ∙⁺⇒⁺∙ cannot be defined if the requirement of reflexivity
+-- is dropped.
+
+private
+
+ module Counterexample where
+
+ data D : Set where
+ i j x y z : D
+
+ data _R_ : Rel D zero where
+ iRj : i R j
+ xRy : x R y
+ yRz : y R z
+
+ xR⁺z : x [ _R_ ]⁺ z
+ xR⁺z =
+ x ∼⁺⟨ [ xRy ] ⟩
+ y ∼⁺⟨ [ yRz ] ⟩∎
+ z ∎
+
+ ix = i ∷ x ∷ []
+ jz = j ∷ z ∷ []
+
+ ix∙⁺jz : Pointwise′ (Plus _R_) ix jz
+ ix∙⁺jz = [ iRj ] ∷ xR⁺z ∷ []
+
+ ¬ix⁺∙jz : ¬ Plus′ (Pointwise′ _R_) ix jz
+ ¬ix⁺∙jz [ iRj ∷ () ∷ [] ]
+ ¬ix⁺∙jz ((iRj ∷ xRy ∷ []) ∷ [ () ∷ yRz ∷ [] ])
+ ¬ix⁺∙jz ((iRj ∷ xRy ∷ []) ∷ (() ∷ yRz ∷ []) ∷ _)
+
+ counterexample :
+ ¬ (∀ {n} {xs ys : Vec D n} →
+ Pointwise (Plus _R_) xs ys →
+ Plus (Pointwise _R_) xs ys)
+ counterexample ∙⁺⇒⁺∙ =
+ ¬ix⁺∙jz (Equivalent.to Plus.equivalent ⟨$⟩
+ Plus.map (_⟨$⟩_ (Equivalent.to equivalent))
+ (∙⁺⇒⁺∙ (Equivalent.from equivalent ⟨$⟩ ix∙⁺jz)))
+
+-- Map.
+
+map : ∀ {a} {A : Set a} {_R_ _R′_ : Rel A a} {n} →
+ _R_ ⇒ _R′_ → Pointwise _R_ ⇒ Pointwise _R′_ {n}
+map R⇒R′ xsRys = ext λ i →
+ R⇒R′ (Pointwise.app xsRys i)
+
+-- A variant.
+
+gmap : ∀ {a} {A A′ : Set a}
+ {_R_ : Rel A a} {_R′_ : Rel A′ a} {f : A → A′} {n} →
+ _R_ =[ f ]⇒ _R′_ →
+ Pointwise _R_ =[ Vec.map {n = n} f ]⇒ Pointwise _R′_
+gmap {_R′_ = _R′_} {f} R⇒R′ {i = xs} {j = ys} xsRys = ext λ i →
+ let module M = Morphism (VecProp.lookup-morphism i) in
+ P.subst₂ _R′_ (P.sym $ M.op-<$> f xs)
+ (P.sym $ M.op-<$> f ys)
+ (R⇒R′ (Pointwise.app xsRys i))
diff --git a/src/Relation/Nullary/Negation.agda b/src/Relation/Nullary/Negation.agda
index 173ec65..a9ec05c 100644
--- a/src/Relation/Nullary/Negation.agda
+++ b/src/Relation/Nullary/Negation.agda
@@ -98,7 +98,7 @@ decidable-stable (no ¬p) ¬¬p = ⊥-elim (¬¬p ¬p)
-- Double-negation is a monad (if we assume that all elements of ¬ ¬ P
-- are equal).
-¬¬-Monad : RawMonad (λ P → ¬ ¬ P)
+¬¬-Monad : ∀ {p} → RawMonad (λ (P : Set p) → ¬ ¬ P)
¬¬-Monad = record
{ return = contradiction
; _>>=_ = λ x f → negated-stable (¬¬-map f x)
diff --git a/src/Universe.agda b/src/Universe.agda
new file mode 100644
index 0000000..4923a30
--- /dev/null
+++ b/src/Universe.agda
@@ -0,0 +1,19 @@
+------------------------------------------------------------------------
+-- Universes
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Universe where
+
+open import Level
+
+-- Universes.
+
+record Universe u e : Set (suc (u ⊔ e)) where
+ field
+ -- Codes.
+ U : Set u
+
+ -- Decoding function.
+ El : U → Set e