summaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorIain Lane <laney@ubuntu.com>2010-12-10 14:57:42 +0000
committerIain Lane <laney@ubuntu.com>2010-12-10 14:57:42 +0000
commitcc1f5c802508a4098a9d4ea99c356277cbe5ff0d (patch)
tree79907ab1f0044031e018bffe3ed0578fa5176446
parentf294a45d2691b750adc2ed9238db7232aa04ad7b (diff)
Imported Upstream version 0.4
-rw-r--r--AllNonAsciiChars.hs17
-rw-r--r--Everything.agda80
-rw-r--r--GNUmakefile3
-rw-r--r--GenerateEverything.hs7
-rw-r--r--LICENCE4
-rw-r--r--README.agda22
-rw-r--r--README.txt8
-rw-r--r--README/Nat.agda8
-rw-r--r--Setup.hs3
-rw-r--r--lib.cabal19
-rw-r--r--release-notes12
-rw-r--r--src/Algebra.agda160
-rw-r--r--src/Algebra/FunctionProperties.agda45
-rw-r--r--src/Algebra/FunctionProperties/Core.agda12
-rw-r--r--src/Algebra/Morphism.agda19
-rw-r--r--src/Algebra/Operations.agda8
-rw-r--r--src/Algebra/Props/AbelianGroup.agda9
-rw-r--r--src/Algebra/Props/BooleanAlgebra.agda102
-rw-r--r--src/Algebra/Props/DistributiveLattice.agda10
-rw-r--r--src/Algebra/Props/Group.agda8
-rw-r--r--src/Algebra/Props/Lattice.agda10
-rw-r--r--src/Algebra/Props/Ring.agda8
-rw-r--r--src/Algebra/RingSolver.agda74
-rw-r--r--src/Algebra/RingSolver/AlmostCommutativeRing.agda46
-rw-r--r--src/Algebra/RingSolver/Lemmas.agda9
-rw-r--r--src/Algebra/RingSolver/Simple.agda12
-rw-r--r--src/Algebra/Structures.agda170
-rw-r--r--src/Category/Applicative/Indexed.agda2
-rw-r--r--src/Category/Functor.agda2
-rw-r--r--src/Category/Monad.agda2
-rw-r--r--src/Category/Monad/Continuation.agda2
-rw-r--r--src/Category/Monad/Indexed.agda2
-rw-r--r--src/Category/Monad/Partiality.agda818
-rw-r--r--src/Category/Monad/State.agda2
-rw-r--r--src/Coinduction.agda63
-rw-r--r--src/Data/AVL/IndexedMap.agda6
-rw-r--r--src/Data/AVL/Sets.agda2
-rw-r--r--src/Data/Bin.agda2
-rw-r--r--src/Data/Bool.agda2
-rw-r--r--src/Data/Bool/Properties.agda137
-rw-r--r--src/Data/BoundedVec.agda2
-rw-r--r--src/Data/Colist.agda19
-rw-r--r--src/Data/Container.agda260
-rw-r--r--src/Data/Container/AlternativeBagAndSetEquality.agda145
-rw-r--r--src/Data/Container/Any.agda260
-rw-r--r--src/Data/Container/Combinator.agda222
-rw-r--r--src/Data/Covec.agda11
-rw-r--r--src/Data/DifferenceList.agda2
-rw-r--r--src/Data/DifferenceNat.agda2
-rw-r--r--src/Data/DifferenceVec.agda2
-rw-r--r--src/Data/Digit.agda2
-rw-r--r--src/Data/Fin.agda2
-rw-r--r--src/Data/Fin/Dec.agda4
-rw-r--r--src/Data/Fin/Props.agda8
-rw-r--r--src/Data/Fin/Subset/Props.agda2
-rw-r--r--src/Data/Fin/Substitution.agda2
-rw-r--r--src/Data/Fin/Substitution/Lemmas.agda2
-rw-r--r--src/Data/Fin/Substitution/List.agda2
-rw-r--r--src/Data/Function/Equality.agda108
-rw-r--r--src/Data/Function/LeftInverse.agda69
-rw-r--r--src/Data/Graph/Acyclic.agda2
-rw-r--r--src/Data/Integer.agda7
-rw-r--r--src/Data/Integer/Divisibility.agda2
-rw-r--r--src/Data/Integer/Properties.agda2
-rw-r--r--src/Data/List.agda4
-rw-r--r--src/Data/List/All.agda5
-rw-r--r--src/Data/List/All/Properties.agda8
-rw-r--r--src/Data/List/Any.agda78
-rw-r--r--src/Data/List/Any/BagAndSetEquality.agda167
-rw-r--r--src/Data/List/Any/Membership.agda238
-rw-r--r--src/Data/List/Any/Properties.agda1250
-rw-r--r--src/Data/List/Countdown.agda10
-rw-r--r--src/Data/List/NonEmpty.agda2
-rw-r--r--src/Data/List/NonEmpty/Properties.agda6
-rw-r--r--src/Data/List/Properties.agda273
-rw-r--r--src/Data/Map.agda85
-rw-r--r--src/Data/Maybe.agda73
-rw-r--r--src/Data/Nat.agda23
-rw-r--r--src/Data/Nat/Coprimality.agda2
-rw-r--r--src/Data/Nat/DivMod.agda6
-rw-r--r--src/Data/Nat/Divisibility.agda3
-rw-r--r--src/Data/Nat/GCD.agda36
-rw-r--r--src/Data/Nat/GCD/Lemmas.agda2
-rw-r--r--src/Data/Nat/InfinitelyOften.agda2
-rw-r--r--src/Data/Nat/LCM.agda2
-rw-r--r--src/Data/Nat/Properties.agda152
-rw-r--r--src/Data/Nat/Show.agda2
-rw-r--r--src/Data/Product.agda36
-rw-r--r--src/Data/Product/Record.agda65
-rw-r--r--src/Data/Rational.agda2
-rw-r--r--src/Data/Sets.agda95
-rw-r--r--src/Data/Sign/Properties.agda2
-rw-r--r--src/Data/Star.agda7
-rw-r--r--src/Data/Star/BoundedVec.agda2
-rw-r--r--src/Data/Star/Decoration.agda2
-rw-r--r--src/Data/Star/Nat.agda2
-rw-r--r--src/Data/Star/Pointer.agda2
-rw-r--r--src/Data/Star/Properties.agda3
-rw-r--r--src/Data/Star/Vec.agda2
-rw-r--r--src/Data/Stream.agda38
-rw-r--r--src/Data/String.agda8
-rw-r--r--src/Data/Sum.agda2
-rw-r--r--src/Data/Unit.agda1
-rw-r--r--src/Data/Vec/Equality.agda7
-rw-r--r--src/Data/Vec/N-ary.agda2
-rw-r--r--src/Data/Vec/Properties.agda2
-rw-r--r--src/Foreign/Haskell.agda10
-rw-r--r--src/Function.agda (renamed from src/Data/Function.agda)13
-rw-r--r--src/Function/Bijection.agda88
-rw-r--r--src/Function/Equality.agda101
-rw-r--r--src/Function/Equivalence.agda105
-rw-r--r--src/Function/Injection.agda (renamed from src/Data/Function/Injection.agda)12
-rw-r--r--src/Function/Inverse.agda273
-rw-r--r--src/Function/Inverse/TypeIsomorphisms.agda270
-rw-r--r--src/Function/LeftInverse.agda90
-rw-r--r--src/Function/Surjection.agda83
-rw-r--r--src/IO.agda8
-rw-r--r--src/IO/Primitive.agda2
-rw-r--r--src/Induction.agda3
-rw-r--r--src/Induction/Lexicographic.agda34
-rw-r--r--src/Induction/Nat.agda41
-rw-r--r--src/Induction/WellFounded.agda108
-rw-r--r--src/Level.agda9
-rw-r--r--src/Reflection.agda81
-rw-r--r--src/Relation/Binary.agda74
-rw-r--r--src/Relation/Binary/Consequences.agda64
-rw-r--r--src/Relation/Binary/Core.agda16
-rw-r--r--src/Relation/Binary/Flip.agda3
-rw-r--r--src/Relation/Binary/HeterogeneousEquality.agda28
-rw-r--r--src/Relation/Binary/Indexed.agda35
-rw-r--r--src/Relation/Binary/Indexed/Core.agda69
-rw-r--r--src/Relation/Binary/InducedPreorders.agda6
-rw-r--r--src/Relation/Binary/List/NonStrictLex.agda120
-rw-r--r--src/Relation/Binary/List/Pointwise.agda39
-rw-r--r--src/Relation/Binary/List/StrictLex.agda168
-rw-r--r--src/Relation/Binary/NonStrictToStrict.agda2
-rw-r--r--src/Relation/Binary/On.agda3
-rw-r--r--src/Relation/Binary/OrderMorphism.agda2
-rw-r--r--src/Relation/Binary/Product/NonStrictLex.agda138
-rw-r--r--src/Relation/Binary/Product/Pointwise.agda268
-rw-r--r--src/Relation/Binary/Product/StrictLex.agda236
-rw-r--r--src/Relation/Binary/PropositionalEquality.agda33
-rw-r--r--src/Relation/Binary/Props/Preorder.agda6
-rw-r--r--src/Relation/Binary/Props/StrictPartialOrder.agda1
-rw-r--r--src/Relation/Binary/Reflection.agda2
-rw-r--r--src/Relation/Binary/Sigma/Pointwise.agda179
-rw-r--r--src/Relation/Binary/Simple.agda13
-rw-r--r--src/Relation/Binary/StrictToNonStrict.agda4
-rw-r--r--src/Relation/Binary/Sum.agda245
-rw-r--r--src/Relation/Nullary.agda12
-rw-r--r--src/Relation/Nullary/Decidable.agda17
-rw-r--r--src/Relation/Nullary/Negation.agda24
-rw-r--r--src/Relation/Nullary/Product.agda7
-rw-r--r--src/Relation/Nullary/Universe.agda44
-rw-r--r--src/Relation/Unary.agda7
155 files changed, 6364 insertions, 2607 deletions
diff --git a/AllNonAsciiChars.hs b/AllNonAsciiChars.hs
index d46b952..50de1f9 100644
--- a/AllNonAsciiChars.hs
+++ b/AllNonAsciiChars.hs
@@ -1,17 +1,20 @@
-- | This module extracts all the non-ASCII characters used by the
-- library code (along with how many times they are used).
---
--- The implementation relies on the utf8-string and FileManip
--- libraries which are available from Hackage.
-module AllNonAsciiChars where
+module Main where
import qualified Data.List as L
import Data.Char
import Data.Function
import Control.Applicative
-import System.IO.UTF8 as U
import System.FilePath.Find
+import System.IO
+
+readUTF8File :: FilePath -> IO String
+readUTF8File f = do
+ h <- openFile f ReadMode
+ hSetEncoding h utf8
+ hGetContents h
main :: IO ()
main = do
@@ -19,9 +22,9 @@ main = do
(extension ~~? ".agda" ||? extension ~~? ".lagda")
"src"
nonAsciiChars <-
- filter (not . isAscii) . concat <$> mapM U.readFile agdaFiles
+ filter (not . isAscii) . concat <$> mapM readUTF8File agdaFiles
let table = reverse $
L.sortBy (compare `on` snd) $
map (\cs -> (head cs, length cs)) $
L.group $ L.sort $ nonAsciiChars
- mapM_ (\(c, count) -> U.putStrLn (c : ": " ++ show count)) table
+ mapM_ (\(c, count) -> putStrLn (c : ": " ++ show count)) table
diff --git a/Everything.agda b/Everything.agda
index 8aa747e..f60361b 100644
--- a/Everything.agda
+++ b/Everything.agda
@@ -82,7 +82,7 @@ import Category.Monad.Partiality
-- The state monad
import Category.Monad.State
--- Types used to make recursive arguments coinductive
+-- Basic types related to coinduction
import Coinduction
-- AVL trees
@@ -124,6 +124,18 @@ import Data.Colist
-- Coinductive "natural" numbers
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
+
-- Coinductive vectors
import Data.Covec
@@ -172,18 +184,6 @@ import Data.Fin.Substitution.Lemmas
-- Application of substitutions to lists, along with various lemmas
import Data.Fin.Substitution.List
--- Simple combinators working solely on and with functions
-import Data.Function
-
--- Function setoids and related constructions
-import Data.Function.Equality
-
--- Injections
-import Data.Function.Injection
-
--- Left inverses
-import Data.Function.LeftInverse
-
-- Directed acyclic multigraphs
import Data.Graph.Acyclic
@@ -208,7 +208,13 @@ import Data.List.All.Properties
-- Lists where at least one element satisfies a given property
import Data.List.Any
--- Properties relating Any to various list functions
+-- Properties related to bag and set equality
+import Data.List.Any.BagAndSetEquality
+
+-- Properties related to list membership
+import Data.List.Any.Membership
+
+-- Properties related to Any
import Data.List.Any.Properties
-- A data structure which keeps track of an upper bound on the number
@@ -227,10 +233,6 @@ import Data.List.Properties
-- Reverse view
import Data.List.Reverse
--- Finite maps, i.e. lookup tables (currently only some type
--- signatures)
-import Data.Map
-
-- The Maybe type
import Data.Maybe
@@ -267,15 +269,9 @@ import Data.Nat.Show
-- Products
import Data.Product
--- Products implemented using records
-import Data.Product.Record
-
-- Rational numbers
import Data.Rational
--- Finite sets (currently only some type signatures)
-import Data.Sets
-
-- Signs
import Data.Sign
@@ -339,6 +335,33 @@ import Data.Vec.Properties
-- Types used (only) when calling out to Haskell via the FFI
import Foreign.Haskell
+-- Simple combinators working solely on and with functions
+import Function
+
+-- Bijections
+import Function.Bijection
+
+-- Function setoids and related constructions
+import Function.Equality
+
+-- Equivalence (coinhabitance)
+import Function.Equivalence
+
+-- Injections
+import Function.Injection
+
+-- Inverses
+import Function.Inverse
+
+-- Various basic type isomorphisms
+import Function.Inverse.TypeIsomorphisms
+
+-- Left inverses
+import Function.LeftInverse
+
+-- Surjections
+import Function.Surjection
+
-- IO
import IO
@@ -360,6 +383,9 @@ import Induction.WellFounded
-- Universe levels
import Level
+-- Support for reflection
+import Reflection
+
-- Properties of homogeneous binary relations
import Relation.Binary
@@ -375,6 +401,9 @@ import Relation.Binary.Flip
-- Heterogeneous equality
import Relation.Binary.HeterogeneousEquality
+-- Indexed binary relations
+import Relation.Binary.Indexed
+
-- Induced preorders
import Relation.Binary.InducedPreorders
@@ -439,6 +468,9 @@ import Relation.Binary.Props.TotalOrder
-- proof by reflection
import Relation.Binary.Reflection
+-- Pointwise lifting of binary relations to sigma types
+import Relation.Binary.Sigma.Pointwise
+
-- Some simple binary relations
import Relation.Binary.Simple
diff --git a/GNUmakefile b/GNUmakefile
index 527e343..a180864 100644
--- a/GNUmakefile
+++ b/GNUmakefile
@@ -5,4 +5,5 @@ test: Everything.agda
.PHONY: Everything.agda
Everything.agda:
- runhaskell GenerateEverything.hs
+ cabal install
+ GenerateEverything
diff --git a/GenerateEverything.hs b/GenerateEverything.hs
index b90a443..eaed773 100644
--- a/GenerateEverything.hs
+++ b/GenerateEverything.hs
@@ -1,8 +1,5 @@
{-# LANGUAGE PatternGuards #-}
--- This program requires that the filepath and FileManip packages from
--- Hackage are installed.
-
import qualified Data.List as List
import Control.Applicative
import System.Environment
@@ -35,7 +32,9 @@ main = do
usage :: String
usage = unlines
- [ "Usage: runhaskell GenerateEverything.hs"
+ [ "GenerateEverything: A utility program for Agda's standard library."
+ , ""
+ , "Usage: GenerateEverything"
, ""
, "This program should be run in the base directory of a clean checkout of"
, "the library."
diff --git a/LICENCE b/LICENCE
index 6641cab..a83b34d 100644
--- a/LICENCE
+++ b/LICENCE
@@ -1,7 +1,7 @@
-Copyright (c) 2007-2009 Nils Anders Danielsson, Ulf Norell, Shin-Cheng
+Copyright (c) 2007-2010 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
+Darin Morrison, Peter Berry
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 65343a2..3749b79 100644
--- a/README.agda
+++ b/README.agda
@@ -1,15 +1,15 @@
module README where
------------------------------------------------------------------------
--- The Agda standard library, version 0.3
+-- The Agda standard library, version 0.4
--
-- Author: Nils Anders Danielsson, with contributions from
--- Jean-Philippe Bernardy, 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, Liang-Ting
+-- Chen, Dan Doel, 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.6.
+-- This version of the library has been tested using Agda 2.2.8.
-- Note that no guarantees are currently made about forwards or
-- backwards compatibility, the library is still at an experimental
@@ -47,9 +47,9 @@ module README where
-- • Coinduction
-- Support for coinduction.
-- • Data
--- Data types and properties about data types. (Also some
--- combinators working solely on and with functions; see
--- Data.Function.)
+-- Data types and properties about data types.
+-- • Function
+-- Combinators and properties related to functions.
-- • Foreign
-- Related to the foreign function interface.
-- • Induction
@@ -59,6 +59,8 @@ module README where
-- Input/output-related functions.
-- • Level
-- Universe levels.
+-- • Reflection
+-- Support for reflection.
-- • Relation
-- Properties of and proofs about relations (mostly homogeneous
-- binary relations).
@@ -180,12 +182,10 @@ import IO
-- -- Reflexivity is expressed in terms of an underlying equality:
-- reflexive : _≈_ ⇒ _∼_
-- trans : Transitive _∼_
--- ∼-resp-≈ : _∼_ Respects₂ _≈_
--
-- module Eq = IsEquivalence isEquivalence
--
--- refl : Reflexive _∼_
--- refl = reflexive Eq.refl
+-- ...
--
-- The Eq module in IsPreorder is not opened publicly, because it
-- contains some fields which clash with fields or other definitions
diff --git a/README.txt b/README.txt
index 96d7988..6361460 100644
--- a/README.txt
+++ b/README.txt
@@ -7,10 +7,8 @@ For information about the library, see README.agda.
The README module imports the Everything module. This module is
generated automatically; if you have downloaded the library from its
darcs repository and want to type check README you can construct
-Everything by running "runhaskell GenerateEverything.hs". Note that
-the GenerateEverything program depends on the filepath and FileManip
-packages from Hackage.
+Everything by running "cabal install && GenerateEverything".
Note that all library sources are located under src. The modules
-README and Everything are not really part of the library, so these
-modules are located in the top-level directory instead.
+README, README.* and Everything are not really part of the library, so
+these modules are located in the top-level directory instead.
diff --git a/README/Nat.agda b/README/Nat.agda
index 0ca3aa2..2717365 100644
--- a/README/Nat.agda
+++ b/README/Nat.agda
@@ -28,10 +28,10 @@ ex₂ = refl
open import Algebra
import Data.Nat.Properties as Nat
private
- module ℕ = CommutativeSemiring Nat.commutativeSemiring
+ module CS = CommutativeSemiring Nat.commutativeSemiring
ex₃ : ∀ m n → m * n ≡ n * m
-ex₃ m n = ℕ.*-comm m n
+ex₃ m n = CS.*-comm m n
-- The module ≡-Reasoning in Relation.Binary.PropositionalEquality
-- provides some combinators for equational reasoning.
@@ -41,8 +41,8 @@ open import Data.Product
ex₄ : ∀ m n → m * (n + 0) ≡ n * m
ex₄ m n = begin
- m * (n + 0) ≡⟨ cong (_*_ m) (proj₂ ℕ.+-identity n) ⟩
- m * n ≡⟨ ℕ.*-comm m n ⟩
+ m * (n + 0) ≡⟨ cong (_*_ m) (proj₂ CS.+-identity n) ⟩
+ m * n ≡⟨ CS.*-comm m n ⟩
n * m ∎
-- The module SemiringSolver in Data.Nat.Properties contains a solver
diff --git a/Setup.hs b/Setup.hs
new file mode 100644
index 0000000..e8ef27d
--- /dev/null
+++ b/Setup.hs
@@ -0,0 +1,3 @@
+import Distribution.Simple
+
+main = defaultMain
diff --git a/lib.cabal b/lib.cabal
new file mode 100644
index 0000000..215ac74
--- /dev/null
+++ b/lib.cabal
@@ -0,0 +1,19 @@
+name: lib
+version: 0.0
+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.*
+
+executable AllNonAsciiChars
+ hs-source-dirs: .
+ main-is: AllNonAsciiChars.hs
+ build-depends: base == 4.2.*,
+ FileManip == 0.3.*
diff --git a/release-notes b/release-notes
index 52e6587..467be99 100644
--- a/release-notes
+++ b/release-notes
@@ -1,4 +1,16 @@
------------------------------------------------------------------------
+Version 0.4
+------------------------------------------------------------------------
+
+Version 0.4 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.8.
+
+Note that no guarantees are made about backwards or forwards
+compatibility, the library is still at an experimental stage.
+
+------------------------------------------------------------------------
Version 0.3
------------------------------------------------------------------------
diff --git a/src/Algebra.agda b/src/Algebra.agda
index cd7fd08..bbc9e9b 100644
--- a/src/Algebra.agda
+++ b/src/Algebra.agda
@@ -3,23 +3,25 @@
-- (packed in records together with sets, operations, etc.)
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Algebra where
open import Relation.Binary
open import Algebra.FunctionProperties
open import Algebra.Structures
-open import Data.Function
+open import Function
open import Level
------------------------------------------------------------------------
-- Semigroups, (commutative) monoids and (abelian) groups
-record Semigroup : Set₁ where
+record Semigroup c ℓ : Set (suc (c ⊔ ℓ)) where
infixl 7 _∙_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_∙_ : Op₂ Carrier
isSemigroup : IsSemigroup _≈_ _∙_
@@ -30,63 +32,63 @@ record Semigroup : Set₁ where
-- A raw monoid is a monoid without any laws.
-record RawMonoid : Set₁ where
+record RawMonoid c ℓ : Set (suc (c ⊔ ℓ)) where
infixl 7 _∙_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_∙_ : Op₂ Carrier
ε : Carrier
-record Monoid : Set₁ where
+record Monoid c ℓ : Set (suc (c ⊔ ℓ)) where
infixl 7 _∙_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_∙_ : Op₂ Carrier
ε : Carrier
isMonoid : IsMonoid _≈_ _∙_ ε
open IsMonoid isMonoid public
- semigroup : Semigroup
+ semigroup : Semigroup _ _
semigroup = record { isSemigroup = isSemigroup }
open Semigroup semigroup public using (setoid)
- rawMonoid : RawMonoid
+ rawMonoid : RawMonoid _ _
rawMonoid = record
{ _≈_ = _≈_
; _∙_ = _∙_
; ε = ε
}
-record CommutativeMonoid : Set₁ where
+record CommutativeMonoid c ℓ : Set (suc (c ⊔ ℓ)) where
infixl 7 _∙_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_∙_ : Op₂ Carrier
ε : Carrier
isCommutativeMonoid : IsCommutativeMonoid _≈_ _∙_ ε
open IsCommutativeMonoid isCommutativeMonoid public
- monoid : Monoid
+ monoid : Monoid _ _
monoid = record { isMonoid = isMonoid }
open Monoid monoid public using (setoid; semigroup; rawMonoid)
-record Group : Set₁ where
+record Group c ℓ : Set (suc (c ⊔ ℓ)) where
infix 8 _⁻¹
infixl 7 _∙_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_∙_ : Op₂ Carrier
ε : Carrier
_⁻¹ : Op₁ Carrier
@@ -94,18 +96,18 @@ record Group : Set₁ where
open IsGroup isGroup public
- monoid : Monoid
+ monoid : Monoid _ _
monoid = record { isMonoid = isMonoid }
open Monoid monoid public using (setoid; semigroup; rawMonoid)
-record AbelianGroup : Set₁ where
+record AbelianGroup c ℓ : Set (suc (c ⊔ ℓ)) where
infix 8 _⁻¹
infixl 7 _∙_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_∙_ : Op₂ Carrier
ε : Carrier
_⁻¹ : Op₁ Carrier
@@ -113,25 +115,25 @@ record AbelianGroup : Set₁ where
open IsAbelianGroup isAbelianGroup public
- group : Group
+ group : Group _ _
group = record { isGroup = isGroup }
open Group group public using (setoid; semigroup; monoid; rawMonoid)
- commutativeMonoid : CommutativeMonoid
+ commutativeMonoid : CommutativeMonoid _ _
commutativeMonoid =
record { isCommutativeMonoid = isCommutativeMonoid }
------------------------------------------------------------------------
-- Various kinds of semirings
-record NearSemiring : Set₁ where
+record NearSemiring c ℓ : Set (suc (c ⊔ ℓ)) where
infixl 7 _*_
infixl 6 _+_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_+_ : Op₂ Carrier
_*_ : Op₂ Carrier
0# : Carrier
@@ -139,7 +141,7 @@ record NearSemiring : Set₁ where
open IsNearSemiring isNearSemiring public
- +-monoid : Monoid
+ +-monoid : Monoid _ _
+-monoid = record { isMonoid = +-isMonoid }
open Monoid +-monoid public
@@ -147,16 +149,16 @@ record NearSemiring : Set₁ where
renaming ( semigroup to +-semigroup
; rawMonoid to +-rawMonoid)
- *-semigroup : Semigroup
+ *-semigroup : Semigroup _ _
*-semigroup = record { isSemigroup = *-isSemigroup }
-record SemiringWithoutOne : Set₁ where
+record SemiringWithoutOne c ℓ : Set (suc (c ⊔ ℓ)) where
infixl 7 _*_
infixl 6 _+_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_+_ : Op₂ Carrier
_*_ : Op₂ Carrier
0# : Carrier
@@ -164,7 +166,7 @@ record SemiringWithoutOne : Set₁ where
open IsSemiringWithoutOne isSemiringWithoutOne public
- nearSemiring : NearSemiring
+ nearSemiring : NearSemiring _ _
nearSemiring = record { isNearSemiring = isNearSemiring }
open NearSemiring nearSemiring public
@@ -173,17 +175,17 @@ record SemiringWithoutOne : Set₁ where
; *-semigroup
)
- +-commutativeMonoid : CommutativeMonoid
+ +-commutativeMonoid : CommutativeMonoid _ _
+-commutativeMonoid =
record { isCommutativeMonoid = +-isCommutativeMonoid }
-record SemiringWithoutAnnihilatingZero : Set₁ where
+record SemiringWithoutAnnihilatingZero c ℓ : Set (suc (c ⊔ ℓ)) where
infixl 7 _*_
infixl 6 _+_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_+_ : Op₂ Carrier
_*_ : Op₂ Carrier
0# : Carrier
@@ -194,7 +196,7 @@ record SemiringWithoutAnnihilatingZero : Set₁ where
open IsSemiringWithoutAnnihilatingZero
isSemiringWithoutAnnihilatingZero public
- +-commutativeMonoid : CommutativeMonoid
+ +-commutativeMonoid : CommutativeMonoid _ _
+-commutativeMonoid =
record { isCommutativeMonoid = +-isCommutativeMonoid }
@@ -205,7 +207,7 @@ record SemiringWithoutAnnihilatingZero : Set₁ where
; monoid to +-monoid
)
- *-monoid : Monoid
+ *-monoid : Monoid _ _
*-monoid = record { isMonoid = *-isMonoid }
open Monoid *-monoid public
@@ -214,13 +216,13 @@ record SemiringWithoutAnnihilatingZero : Set₁ where
; rawMonoid to *-rawMonoid
)
-record Semiring : Set₁ where
+record Semiring c ℓ : Set (suc (c ⊔ ℓ)) where
infixl 7 _*_
infixl 6 _+_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_+_ : Op₂ Carrier
_*_ : Op₂ Carrier
0# : Carrier
@@ -229,7 +231,7 @@ record Semiring : Set₁ where
open IsSemiring isSemiring public
- semiringWithoutAnnihilatingZero : SemiringWithoutAnnihilatingZero
+ semiringWithoutAnnihilatingZero : SemiringWithoutAnnihilatingZero _ _
semiringWithoutAnnihilatingZero = record
{ isSemiringWithoutAnnihilatingZero =
isSemiringWithoutAnnihilatingZero
@@ -243,20 +245,20 @@ record Semiring : Set₁ where
; *-semigroup; *-rawMonoid; *-monoid
)
- semiringWithoutOne : SemiringWithoutOne
+ semiringWithoutOne : SemiringWithoutOne _ _
semiringWithoutOne =
record { isSemiringWithoutOne = isSemiringWithoutOne }
open SemiringWithoutOne semiringWithoutOne public
using (nearSemiring)
-record CommutativeSemiringWithoutOne : Set₁ where
+record CommutativeSemiringWithoutOne c ℓ : Set (suc (c ⊔ ℓ)) where
infixl 7 _*_
infixl 6 _+_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_+_ : Op₂ Carrier
_*_ : Op₂ Carrier
0# : Carrier
@@ -266,7 +268,7 @@ record CommutativeSemiringWithoutOne : Set₁ where
open IsCommutativeSemiringWithoutOne
isCommutativeSemiringWithoutOne public
- semiringWithoutOne : SemiringWithoutOne
+ semiringWithoutOne : SemiringWithoutOne _ _
semiringWithoutOne =
record { isSemiringWithoutOne = isSemiringWithoutOne }
@@ -278,13 +280,13 @@ record CommutativeSemiringWithoutOne : Set₁ where
; nearSemiring
)
-record CommutativeSemiring : Set₁ where
+record CommutativeSemiring c ℓ : Set (suc (c ⊔ ℓ)) where
infixl 7 _*_
infixl 6 _+_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_+_ : Op₂ Carrier
_*_ : Op₂ Carrier
0# : Carrier
@@ -293,7 +295,7 @@ record CommutativeSemiring : Set₁ where
open IsCommutativeSemiring isCommutativeSemiring public
- semiring : Semiring
+ semiring : Semiring _ _
semiring = record { isSemiring = isSemiring }
open Semiring semiring public
@@ -305,11 +307,11 @@ record CommutativeSemiring : Set₁ where
; semiringWithoutAnnihilatingZero
)
- *-commutativeMonoid : CommutativeMonoid
+ *-commutativeMonoid : CommutativeMonoid _ _
*-commutativeMonoid =
record { isCommutativeMonoid = *-isCommutativeMonoid }
- commutativeSemiringWithoutOne : CommutativeSemiringWithoutOne
+ commutativeSemiringWithoutOne : CommutativeSemiringWithoutOne _ _
commutativeSemiringWithoutOne = record
{ isCommutativeSemiringWithoutOne = isCommutativeSemiringWithoutOne
}
@@ -319,26 +321,26 @@ record CommutativeSemiring : Set₁ where
-- A raw ring is a ring without any laws.
-record RawRing : Set₁ where
+record RawRing c : Set (suc c) where
infix 8 -_
infixl 7 _*_
infixl 6 _+_
field
- Carrier : Set
+ Carrier : Set c
_+_ : Op₂ Carrier
_*_ : Op₂ Carrier
-_ : Op₁ Carrier
0# : Carrier
1# : Carrier
-record Ring : Set₁ where
+record Ring c ℓ : Set (suc (c ⊔ ℓ)) where
infix 8 -_
infixl 7 _*_
infixl 6 _+_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_+_ : Op₂ Carrier
_*_ : Op₂ Carrier
-_ : Op₁ Carrier
@@ -348,10 +350,10 @@ record Ring : Set₁ where
open IsRing isRing public
- +-abelianGroup : AbelianGroup
+ +-abelianGroup : AbelianGroup _ _
+-abelianGroup = record { isAbelianGroup = +-isAbelianGroup }
- semiring : Semiring
+ semiring : Semiring _ _
semiring = record { isSemiring = isSemiring }
open Semiring semiring public
@@ -366,7 +368,7 @@ record Ring : Set₁ where
open AbelianGroup +-abelianGroup public
using () renaming (group to +-group)
- rawRing : RawRing
+ rawRing : RawRing _
rawRing = record
{ _+_ = _+_
; _*_ = _*_
@@ -375,14 +377,14 @@ record Ring : Set₁ where
; 1# = 1#
}
-record CommutativeRing : Set₁ where
+record CommutativeRing c ℓ : Set (suc (c ⊔ ℓ)) where
infix 8 -_
infixl 7 _*_
infixl 6 _+_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_+_ : Op₂ Carrier
_*_ : Op₂ Carrier
-_ : Op₁ Carrier
@@ -392,10 +394,10 @@ record CommutativeRing : Set₁ where
open IsCommutativeRing isCommutativeRing public
- ring : Ring
+ ring : Ring _ _
ring = record { isRing = isRing }
- commutativeSemiring : CommutativeSemiring
+ commutativeSemiring : CommutativeSemiring _ _
commutativeSemiring =
record { isCommutativeSemiring = isCommutativeSemiring }
@@ -412,13 +414,13 @@ record CommutativeRing : Set₁ where
------------------------------------------------------------------------
-- (Distributive) lattices and boolean algebras
-record Lattice : Set₁ where
+record Lattice c ℓ : Set (suc (c ⊔ ℓ)) where
infixr 7 _∧_
infixr 6 _∨_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_∨_ : Op₂ Carrier
_∧_ : Op₂ Carrier
isLattice : IsLattice _≈_ _∨_ _∧_
@@ -428,32 +430,32 @@ record Lattice : Set₁ where
setoid : Setoid _ _
setoid = record { isEquivalence = isEquivalence }
-record DistributiveLattice : Set₁ where
+record DistributiveLattice c ℓ : Set (suc (c ⊔ ℓ)) where
infixr 7 _∧_
infixr 6 _∨_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_∨_ : Op₂ Carrier
_∧_ : Op₂ Carrier
isDistributiveLattice : IsDistributiveLattice _≈_ _∨_ _∧_
open IsDistributiveLattice isDistributiveLattice public
- lattice : Lattice
+ lattice : Lattice _ _
lattice = record { isLattice = isLattice }
open Lattice lattice public using (setoid)
-record BooleanAlgebra : Set₁ where
+record BooleanAlgebra c ℓ : Set (suc (c ⊔ ℓ)) where
infix 8 ¬_
infixr 7 _∧_
infixr 6 _∨_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_∨_ : Op₂ Carrier
_∧_ : Op₂ Carrier
¬_ : Op₁ Carrier
@@ -463,7 +465,7 @@ record BooleanAlgebra : Set₁ where
open IsBooleanAlgebra isBooleanAlgebra public
- distributiveLattice : DistributiveLattice
+ distributiveLattice : DistributiveLattice _ _
distributiveLattice =
record { isDistributiveLattice = isDistributiveLattice }
diff --git a/src/Algebra/FunctionProperties.agda b/src/Algebra/FunctionProperties.agda
index dfc0d64..df1d94c 100644
--- a/src/Algebra/FunctionProperties.agda
+++ b/src/Algebra/FunctionProperties.agda
@@ -2,6 +2,8 @@
-- Properties of functions, such as associativity and commutativity
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- These properties can (for instance) be used to define algebraic
-- structures.
@@ -11,7 +13,8 @@ open import Relation.Binary
-- The properties are specified using the following relation as
-- "equality".
-module Algebra.FunctionProperties {A} (_≈_ : Rel A zero) where
+module Algebra.FunctionProperties
+ {a ℓ} {A : Set a} (_≈_ : Rel A ℓ) where
open import Data.Product
@@ -23,64 +26,64 @@ open import Algebra.FunctionProperties.Core public
------------------------------------------------------------------------
-- Properties of operations
-Associative : Op₂ A → Set
+Associative : Op₂ A → Set _
Associative _∙_ = ∀ x y z → ((x ∙ y) ∙ z) ≈ (x ∙ (y ∙ z))
-Commutative : Op₂ A → Set
+Commutative : Op₂ A → Set _
Commutative _∙_ = ∀ x y → (x ∙ y) ≈ (y ∙ x)
-LeftIdentity : A → Op₂ A → Set
+LeftIdentity : A → Op₂ A → Set _
LeftIdentity e _∙_ = ∀ x → (e ∙ x) ≈ x
-RightIdentity : A → Op₂ A → Set
+RightIdentity : A → Op₂ A → Set _
RightIdentity e _∙_ = ∀ x → (x ∙ e) ≈ x
-Identity : A → Op₂ A → Set
+Identity : A → Op₂ A → Set _
Identity e ∙ = LeftIdentity e ∙ × RightIdentity e ∙
-LeftZero : A → Op₂ A → Set
+LeftZero : A → Op₂ A → Set _
LeftZero z _∙_ = ∀ x → (z ∙ x) ≈ z
-RightZero : A → Op₂ A → Set
+RightZero : A → Op₂ A → Set _
RightZero z _∙_ = ∀ x → (x ∙ z) ≈ z
-Zero : A → Op₂ A → Set
+Zero : A → Op₂ A → Set _
Zero z ∙ = LeftZero z ∙ × RightZero z ∙
-LeftInverse : A → Op₁ A → Op₂ A → Set
+LeftInverse : A → Op₁ A → Op₂ A → Set _
LeftInverse e _⁻¹ _∙_ = ∀ x → (x ⁻¹ ∙ x) ≈ e
-RightInverse : A → Op₁ A → Op₂ A → Set
+RightInverse : A → Op₁ A → Op₂ A → Set _
RightInverse e _⁻¹ _∙_ = ∀ x → (x ∙ (x ⁻¹)) ≈ e
-Inverse : A → Op₁ A → Op₂ A → Set
+Inverse : A → Op₁ A → Op₂ A → Set _
Inverse e ⁻¹ ∙ = LeftInverse e ⁻¹ ∙ × RightInverse e ⁻¹ ∙
-_DistributesOverˡ_ : Op₂ A → Op₂ A → Set
+_DistributesOverˡ_ : Op₂ A → Op₂ A → Set _
_*_ DistributesOverˡ _+_ =
∀ x y z → (x * (y + z)) ≈ ((x * y) + (x * z))
-_DistributesOverʳ_ : Op₂ A → Op₂ A → Set
+_DistributesOverʳ_ : Op₂ A → Op₂ A → Set _
_*_ DistributesOverʳ _+_ =
∀ x y z → ((y + z) * x) ≈ ((y * x) + (z * x))
-_DistributesOver_ : Op₂ A → Op₂ A → Set
+_DistributesOver_ : Op₂ A → Op₂ A → Set _
* DistributesOver + = (* DistributesOverˡ +) × (* DistributesOverʳ +)
-_IdempotentOn_ : Op₂ A → A → Set
+_IdempotentOn_ : Op₂ A → A → Set _
_∙_ IdempotentOn x = (x ∙ x) ≈ x
-Idempotent : Op₂ A → Set
+Idempotent : Op₂ A → Set _
Idempotent ∙ = ∀ x → ∙ IdempotentOn x
-IdempotentFun : Op₁ A → Set
+IdempotentFun : Op₁ A → Set _
IdempotentFun f = ∀ x → f (f x) ≈ f x
-_Absorbs_ : Op₂ A → Op₂ A → Set
+_Absorbs_ : Op₂ A → Op₂ A → Set _
_∙_ Absorbs _∘_ = ∀ x y → (x ∙ (x ∘ y)) ≈ x
-Absorptive : Op₂ A → Op₂ A → Set
+Absorptive : Op₂ A → Op₂ A → Set _
Absorptive ∙ ∘ = (∙ Absorbs ∘) × (∘ Absorbs ∙)
-Involutive : Op₁ A → Set
+Involutive : Op₁ A → Set _
Involutive f = ∀ x → f (f x) ≈ x
diff --git a/src/Algebra/FunctionProperties/Core.agda b/src/Algebra/FunctionProperties/Core.agda
index a9facc4..1b604f0 100644
--- a/src/Algebra/FunctionProperties/Core.agda
+++ b/src/Algebra/FunctionProperties/Core.agda
@@ -2,18 +2,22 @@
-- Properties of functions, such as associativity and commutativity
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- This file contains some core definitions which are reexported by
-- Algebra.FunctionProperties. They are placed here because
--- Algebra.FunctionProperties is a parameterised module, and the
--- parameters are irrelevant for these definitions.
+-- Algebra.FunctionProperties is a parameterised module, and some of
+-- the parameters are irrelevant for these definitions.
module Algebra.FunctionProperties.Core where
+open import Level
+
------------------------------------------------------------------------
-- Unary and binary operations
-Op₁ : Set → Set
+Op₁ : ∀ {ℓ} → Set ℓ → Set ℓ
Op₁ A = A → A
-Op₂ : Set → Set
+Op₂ : ∀ {ℓ} → Set ℓ → Set ℓ
Op₂ A = A → A → A
diff --git a/src/Algebra/Morphism.agda b/src/Algebra/Morphism.agda
index d694e34..3bf8c49 100644
--- a/src/Algebra/Morphism.agda
+++ b/src/Algebra/Morphism.agda
@@ -2,13 +2,15 @@
-- Morphisms between algebraic structures
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Algebra.Morphism where
open import Relation.Binary
open import Algebra
open import Algebra.FunctionProperties
import Algebra.Props.Group as GroupP
-open import Data.Function
+open import Function
open import Data.Product
open import Level
import Relation.Binary.EqReasoning as EqR
@@ -16,17 +18,18 @@ import Relation.Binary.EqReasoning as EqR
------------------------------------------------------------------------
-- Basic definitions
-module Definitions (From To : Set) (_≈_ : Rel To zero) where
- Morphism : Set
+module Definitions {f t ℓ}
+ (From : Set f) (To : Set t) (_≈_ : Rel To ℓ) where
+ Morphism : Set _
Morphism = From → To
- Homomorphic₀ : Morphism → From → To → Set
+ Homomorphic₀ : Morphism → From → To → Set _
Homomorphic₀ ⟦_⟧ ∙ ∘ = ⟦ ∙ ⟧ ≈ ∘
- Homomorphic₁ : Morphism → Fun₁ From → Op₁ To → Set
+ Homomorphic₁ : Morphism → Fun₁ From → Op₁ To → Set _
Homomorphic₁ ⟦_⟧ ∙_ ∘_ = ∀ x → ⟦ ∙ x ⟧ ≈ ∘ ⟦ x ⟧
- Homomorphic₂ : Morphism → Fun₂ From → Op₂ To → Set
+ Homomorphic₂ : Morphism → Fun₂ From → Op₂ To → Set _
Homomorphic₂ ⟦_⟧ _∙_ _∘_ =
∀ x y → ⟦ x ∙ y ⟧ ≈ (⟦ x ⟧ ∘ ⟦ y ⟧)
@@ -35,7 +38,9 @@ module Definitions (From To : Set) (_≈_ : Rel To zero) where
-- Ring homomorphisms.
-record _-Ring⟶_ (From To : Ring) : Set where
+record _-Ring⟶_ {r₁ r₂ r₃ r₄}
+ (From : Ring r₁ r₂) (To : Ring r₃ r₄) :
+ Set (r₁ ⊔ r₂ ⊔ r₃ ⊔ r₄) where
private
module F = Ring From
module T = Ring To
diff --git a/src/Algebra/Operations.agda b/src/Algebra/Operations.agda
index 57a62d4..d8b6e13 100644
--- a/src/Algebra/Operations.agda
+++ b/src/Algebra/Operations.agda
@@ -3,13 +3,15 @@
-- exponentiation)
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
open import Algebra
-module Algebra.Operations (s : Semiring) where
+module Algebra.Operations {s₁ s₂} (S : Semiring s₁ s₂) where
-open Semiring s hiding (zero)
+open Semiring S hiding (zero)
open import Data.Nat using (zero; suc; ℕ)
-open import Data.Function
+open import Function
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
import Relation.Binary.EqReasoning as EqR
diff --git a/src/Algebra/Props/AbelianGroup.agda b/src/Algebra/Props/AbelianGroup.agda
index 4314b61..774c061 100644
--- a/src/Algebra/Props/AbelianGroup.agda
+++ b/src/Algebra/Props/AbelianGroup.agda
@@ -2,13 +2,16 @@
-- Some derivable properties
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
open import Algebra
-module Algebra.Props.AbelianGroup (g : AbelianGroup) where
+module Algebra.Props.AbelianGroup
+ {g₁ g₂} (G : AbelianGroup g₁ g₂) where
-open AbelianGroup g
+open AbelianGroup G
import Relation.Binary.EqReasoning as EqR; open EqR setoid
-open import Data.Function
+open import Function
open import Data.Product
private
diff --git a/src/Algebra/Props/BooleanAlgebra.agda b/src/Algebra/Props/BooleanAlgebra.agda
index e8ef0dd..8efac7b 100644
--- a/src/Algebra/Props/BooleanAlgebra.agda
+++ b/src/Algebra/Props/BooleanAlgebra.agda
@@ -2,18 +2,22 @@
-- Some derivable properties
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
open import Algebra
-module Algebra.Props.BooleanAlgebra (b : BooleanAlgebra) where
+module Algebra.Props.BooleanAlgebra
+ {b₁ b₂} (B : BooleanAlgebra b₁ b₂)
+ where
-open BooleanAlgebra b
+open BooleanAlgebra B
import Algebra.Props.DistributiveLattice as DL
open DL distributiveLattice public
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 Data.Function
+open import Function
open import Data.Product
------------------------------------------------------------------------
@@ -48,7 +52,7 @@ open import Data.Product
; ¬-cong = ¬-cong
}
-∧-∨-booleanAlgebra : BooleanAlgebra
+∧-∨-booleanAlgebra : BooleanAlgebra _ _
∧-∨-booleanAlgebra = record
{ _∧_ = _∨_
; _∨_ = _∧_
@@ -93,35 +97,29 @@ private
∨-∧-isCommutativeSemiring : IsCommutativeSemiring _≈_ _∨_ _∧_ ⊥ ⊤
∨-∧-isCommutativeSemiring = record
- { isSemiring = record
- { isSemiringWithoutAnnihilatingZero = record
- { +-isCommutativeMonoid = record
- { isMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∨-assoc
- ; ∙-cong = ∨-cong
- }
- ; identity = ∨-identity
- }
- ; comm = ∨-comm
- }
- ; *-isMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∧-assoc
- ; ∙-cong = ∧-cong
- }
- ; identity = ∧-identity
- }
- ; distrib = ∧-∨-distrib
+ { +-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∨-assoc
+ ; ∙-cong = ∨-cong
}
- ; zero = ∧-zero
+ ; identityˡ = proj₁ ∨-identity
+ ; comm = ∨-comm
}
- ; *-comm = ∧-comm
+ ; *-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∧-assoc
+ ; ∙-cong = ∧-cong
+ }
+ ; identityˡ = proj₁ ∧-identity
+ ; comm = ∧-comm
+ }
+ ; distribʳ = proj₂ ∧-∨-distrib
+ ; zeroˡ = proj₁ ∧-zero
}
-∨-∧-commutativeSemiring : CommutativeSemiring
+∨-∧-commutativeSemiring : CommutativeSemiring _ _
∨-∧-commutativeSemiring = record
{ _+_ = _∨_
; _*_ = _∧_
@@ -148,35 +146,29 @@ private
∧-∨-isCommutativeSemiring : IsCommutativeSemiring _≈_ _∧_ _∨_ ⊤ ⊥
∧-∨-isCommutativeSemiring = record
- { isSemiring = record
- { isSemiringWithoutAnnihilatingZero = record
- { +-isCommutativeMonoid = record
- { isMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∧-assoc
- ; ∙-cong = ∧-cong
- }
- ; identity = ∧-identity
- }
- ; comm = ∧-comm
- }
- ; *-isMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∨-assoc
- ; ∙-cong = ∨-cong
- }
- ; identity = ∨-identity
- }
- ; distrib = ∨-∧-distrib
+ { +-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∧-assoc
+ ; ∙-cong = ∧-cong
+ }
+ ; identityˡ = proj₁ ∧-identity
+ ; comm = ∧-comm
+ }
+ ; *-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∨-assoc
+ ; ∙-cong = ∨-cong
}
- ; zero = ∨-zero
+ ; identityˡ = proj₁ ∨-identity
+ ; comm = ∨-comm
}
- ; *-comm = ∨-comm
+ ; distribʳ = proj₂ ∨-∧-distrib
+ ; zeroˡ = proj₁ ∨-zero
}
-∧-∨-commutativeSemiring : CommutativeSemiring
+∧-∨-commutativeSemiring : CommutativeSemiring _ _
∧-∨-commutativeSemiring =
record { isCommutativeSemiring = ∧-∨-isCommutativeSemiring }
@@ -534,7 +526,7 @@ module XorRing
; *-comm = ∧-comm
}
- commutativeRing : CommutativeRing
+ commutativeRing : CommutativeRing _ _
commutativeRing = record
{ _+_ = _⊕_
; _*_ = _∧_
diff --git a/src/Algebra/Props/DistributiveLattice.agda b/src/Algebra/Props/DistributiveLattice.agda
index 79a1782..4dde88d 100644
--- a/src/Algebra/Props/DistributiveLattice.agda
+++ b/src/Algebra/Props/DistributiveLattice.agda
@@ -2,18 +2,20 @@
-- Some derivable properties
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
open import Algebra
module Algebra.Props.DistributiveLattice
- (dl : DistributiveLattice)
+ {dl₁ dl₂} (DL : DistributiveLattice dl₁ dl₂)
where
-open DistributiveLattice dl
+open DistributiveLattice DL
import Algebra.Props.Lattice as L; open L lattice public
open import Algebra.Structures
import Algebra.FunctionProperties as P; open P _≈_
import Relation.Binary.EqReasoning as EqR; open EqR setoid
-open import Data.Function
+open import Function
open import Data.Product
∨-∧-distrib : _∨_ DistributesOver _∧_
@@ -54,7 +56,7 @@ open import Data.Product
; ∨-∧-distribʳ = proj₂ ∧-∨-distrib
}
-∧-∨-distributiveLattice : DistributiveLattice
+∧-∨-distributiveLattice : DistributiveLattice _ _
∧-∨-distributiveLattice = record
{ _∧_ = _∨_
; _∨_ = _∧_
diff --git a/src/Algebra/Props/Group.agda b/src/Algebra/Props/Group.agda
index 0bbc5d1..e7dfe84 100644
--- a/src/Algebra/Props/Group.agda
+++ b/src/Algebra/Props/Group.agda
@@ -2,14 +2,16 @@
-- Some derivable properties
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
open import Algebra
-module Algebra.Props.Group (g : Group) where
+module Algebra.Props.Group {g₁ g₂} (G : Group g₁ g₂) where
-open Group g
+open Group G
import Algebra.FunctionProperties as P; open P _≈_
import Relation.Binary.EqReasoning as EqR; open EqR setoid
-open import Data.Function
+open import Function
open import Data.Product
⁻¹-involutive : ∀ x → x ⁻¹ ⁻¹ ≈ x
diff --git a/src/Algebra/Props/Lattice.agda b/src/Algebra/Props/Lattice.agda
index d1113f7..e8f46e0 100644
--- a/src/Algebra/Props/Lattice.agda
+++ b/src/Algebra/Props/Lattice.agda
@@ -2,15 +2,17 @@
-- Some derivable properties
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
open import Algebra
-module Algebra.Props.Lattice (l : Lattice) where
+module Algebra.Props.Lattice {l₁ l₂} (L : Lattice l₁ l₂) where
-open Lattice l
+open Lattice L
open import Algebra.Structures
import Algebra.FunctionProperties as P; open P _≈_
import Relation.Binary.EqReasoning as EqR; open EqR setoid
-open import Data.Function
+open import Function
open import Data.Product
∧-idempotent : Idempotent _∧_
@@ -39,7 +41,7 @@ open import Data.Product
; absorptive = swap absorptive
}
-∧-∨-lattice : Lattice
+∧-∨-lattice : Lattice _ _
∧-∨-lattice = record
{ _∧_ = _∨_
; _∨_ = _∧_
diff --git a/src/Algebra/Props/Ring.agda b/src/Algebra/Props/Ring.agda
index e41e590..a2c05d8 100644
--- a/src/Algebra/Props/Ring.agda
+++ b/src/Algebra/Props/Ring.agda
@@ -2,13 +2,15 @@
-- Some derivable properties
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
open import Algebra
-module Algebra.Props.Ring (r : Ring) where
+module Algebra.Props.Ring {r₁ r₂} (R : Ring r₁ r₂) where
-open Ring r
+open Ring R
import Relation.Binary.EqReasoning as EqR; open EqR setoid
-open import Data.Function
+open import Function
open import Data.Product
-‿*-distribˡ : ∀ x y → - x * y ≈ - (x * y)
diff --git a/src/Algebra/RingSolver.agda b/src/Algebra/RingSolver.agda
index d31c9c2..c678408 100644
--- a/src/Algebra/RingSolver.agda
+++ b/src/Algebra/RingSolver.agda
@@ -2,6 +2,8 @@
-- Solver for commutative ring or semiring equalities
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- Uses ideas from the Coq ring tactic. See "Proving Equalities in a
-- Commutative Ring Done Right in Coq" by Grégoire and Mahboubi. The
-- code below is not optimised like theirs, though.
@@ -10,14 +12,15 @@ open import Algebra
open import Algebra.RingSolver.AlmostCommutativeRing
module Algebra.RingSolver
- (coeff : RawRing) -- Coefficient "ring".
- (r : AlmostCommutativeRing) -- Main "ring".
- (morphism : coeff -Raw-AlmostCommutative⟶ r)
+ {r₁ r₂ r₃}
+ (Coeff : RawRing r₁) -- Coefficient "ring".
+ (R : AlmostCommutativeRing r₂ r₃) -- Main "ring".
+ (morphism : Coeff -Raw-AlmostCommutative⟶ R)
where
-import Algebra.RingSolver.Lemmas as L; open L coeff r morphism
-private module C = RawRing coeff
-open AlmostCommutativeRing r hiding (zero)
+import Algebra.RingSolver.Lemmas as L; open L Coeff R morphism
+private module C = RawRing Coeff
+open AlmostCommutativeRing R hiding (zero)
import Algebra.FunctionProperties as P; open P _≈_
open import Algebra.Morphism
open _-Raw-AlmostCommutative⟶_ morphism renaming (⟦_⟧ to ⟦_⟧')
@@ -30,7 +33,8 @@ import Relation.Binary.Reflection as Reflection
open import Data.Nat using (ℕ; suc; zero) renaming (_+_ to _ℕ-+_)
open import Data.Fin as Fin using (Fin; zero; suc)
open import Data.Vec
-open import Data.Function hiding (_∶_)
+open import Function hiding (type-signature)
+open import Level
infix 9 _↑ :-_ -‿NF_
infixr 9 _:^_ _^-NF_ _:↑_
@@ -48,7 +52,7 @@ data Op : Set where
-- The polynomials are indexed over the number of variables.
-data Polynomial (m : ℕ) : Set where
+data Polynomial (m : ℕ) : Set r₁ where
op : (o : Op) (p₁ : Polynomial m) (p₂ : Polynomial m) → Polynomial m
con : (c : C.Carrier) → Polynomial m
var : (x : Fin m) → Polynomial m
@@ -79,12 +83,12 @@ sem [*] = _*_
⟦ p :^ n ⟧ ρ = ⟦ p ⟧ ρ ^ n
⟦ :- p ⟧ ρ = - ⟦ p ⟧ ρ
-private
+-- Equality.
- -- Equality.
+_≛_ : ∀ {n} → Polynomial n → Polynomial n → Set _
+p₁ ≛ p₂ = ∀ {ρ} → ⟦ p₁ ⟧ ρ ≈ ⟦ p₂ ⟧ ρ
- _≛_ : ∀ {n} → Polynomial n → Polynomial n → Set
- p₁ ≛ p₂ = ∀ {ρ} → ⟦ p₁ ⟧ ρ ≈ ⟦ p₂ ⟧ ρ
+private
-- Reindexing.
@@ -98,24 +102,22 @@ private
------------------------------------------------------------------------
-- Normal forms of polynomials
-private
-
- -- The normal forms (Horner forms) are indexed over
- -- * the number of variables in the polynomial, and
- -- * an equivalent polynomial.
+-- The normal forms (Horner forms) are indexed over
+-- * the number of variables in the polynomial, and
+-- * an equivalent polynomial.
- data Normal : (n : ℕ) → Polynomial n → Set where
- con : (c : C.Carrier) → Normal 0 (con c)
- _↑ : ∀ {n p'} (p : Normal n p') → Normal (suc n) (p' :↑ 1)
- _*x+_ : ∀ {n p' c'} (p : Normal (suc n) p') (c : Normal n c') →
- Normal (suc n) (p' :* var zero :+ c' :↑ 1)
- _∶_ : ∀ {n p₁ p₂} (p : Normal n p₁) (eq : p₁ ≛ p₂) → Normal n p₂
+data Normal : (n : ℕ) → Polynomial n → Set (r₁ ⊔ r₂ ⊔ r₃) where
+ con : (c : C.Carrier) → Normal 0 (con c)
+ _↑ : ∀ {n p'} (p : Normal n p') → Normal (suc n) (p' :↑ 1)
+ _*x+_ : ∀ {n p' c'} (p : Normal (suc n) p') (c : Normal n c') →
+ Normal (suc n) (p' :* var zero :+ c' :↑ 1)
+ _∶_ : ∀ {n p₁ p₂} (p : Normal n p₁) (eq : p₁ ≛ p₂) → Normal n p₂
- ⟦_⟧-NF : ∀ {n p} → Normal n p → Vec Carrier n → Carrier
- ⟦ p ∶ _ ⟧-NF ρ = ⟦ p ⟧-NF ρ
- ⟦ con c ⟧-NF ρ = ⟦ c ⟧'
- ⟦ p ↑ ⟧-NF (x ∷ ρ) = ⟦ p ⟧-NF ρ
- ⟦ p *x+ c ⟧-NF (x ∷ ρ) = (⟦ p ⟧-NF (x ∷ ρ) * x) + ⟦ c ⟧-NF ρ
+⟦_⟧-Normal : ∀ {n p} → Normal n p → Vec Carrier n → Carrier
+⟦ p ∶ _ ⟧-Normal ρ = ⟦ p ⟧-Normal ρ
+⟦ con c ⟧-Normal ρ = ⟦ c ⟧'
+⟦ p ↑ ⟧-Normal (x ∷ ρ) = ⟦ p ⟧-Normal ρ
+⟦ p *x+ c ⟧-Normal (x ∷ ρ) = (⟦ p ⟧-Normal (x ∷ ρ) * x) + ⟦ c ⟧-Normal ρ
------------------------------------------------------------------------
-- Normalisation
@@ -184,15 +186,15 @@ private
normaliseOp [+] = _+-NF_
normaliseOp [*] = _*-NF_
- normalise : ∀ {n} (p : Polynomial n) → Normal n p
- normalise (op o p₁ p₂) = normalise p₁ ⟨ normaliseOp o ⟩ normalise p₂
- normalise (con c) = con-NF c
- normalise (var i) = var-NF i
- normalise (p :^ n) = normalise p ^-NF n
- normalise (:- p) = -‿NF normalise p
+normalise : ∀ {n} (p : Polynomial n) → Normal n p
+normalise (op o p₁ p₂) = normalise p₁ ⟨ normaliseOp o ⟩ normalise p₂
+normalise (con c) = con-NF c
+normalise (var i) = var-NF i
+normalise (p :^ n) = normalise p ^-NF n
+normalise (:- p) = -‿NF normalise p
⟦_⟧↓ : ∀ {n} → Polynomial n → Vec Carrier n → Carrier
-⟦ p ⟧↓ ρ = ⟦ normalise p ⟧-NF ρ
+⟦ p ⟧↓ ρ = ⟦ normalise p ⟧-Normal ρ
------------------------------------------------------------------------
-- Correctness
@@ -213,7 +215,7 @@ private
raise-sem (:- p) ρ = -‿cong (raise-sem p ρ)
nf-sound : ∀ {n p} (nf : Normal n p) ρ →
- ⟦ nf ⟧-NF ρ ≈ ⟦ p ⟧ ρ
+ ⟦ nf ⟧-Normal ρ ≈ ⟦ p ⟧ ρ
nf-sound (nf ∶ eq) ρ = nf-sound nf ρ ⟨ trans ⟩ eq
nf-sound (con c) ρ = refl
nf-sound (_↑ {p' = p'} nf) (x ∷ ρ) =
diff --git a/src/Algebra/RingSolver/AlmostCommutativeRing.agda b/src/Algebra/RingSolver/AlmostCommutativeRing.agda
index 961b4d3..946ff9c 100644
--- a/src/Algebra/RingSolver/AlmostCommutativeRing.agda
+++ b/src/Algebra/RingSolver/AlmostCommutativeRing.agda
@@ -3,6 +3,8 @@
-- commutative rings), used by the ring solver
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Algebra.RingSolver.AlmostCommutativeRing where
open import Relation.Binary
@@ -10,14 +12,15 @@ open import Algebra
open import Algebra.Structures
open import Algebra.FunctionProperties
import Algebra.Morphism as Morphism
-open import Data.Function
+open import Function
open import Level
------------------------------------------------------------------------
-- Definitions
-record IsAlmostCommutativeRing {A} (_≈_ : Rel A zero)
- (_+_ _*_ : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set where
+record IsAlmostCommutativeRing
+ {a ℓ} {A : Set a} (_≈_ : Rel A ℓ)
+ (_+_ _*_ : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
field
isCommutativeSemiring : IsCommutativeSemiring _≈_ _+_ _*_ 0# 1#
-‿cong : -_ Preserves _≈_ ⟶ _≈_
@@ -26,14 +29,14 @@ record IsAlmostCommutativeRing {A} (_≈_ : Rel A zero)
open IsCommutativeSemiring isCommutativeSemiring public
-record AlmostCommutativeRing : Set₁ where
+record AlmostCommutativeRing c ℓ : Set (suc (c ⊔ ℓ)) where
infix 8 -_
infixl 7 _*_
infixl 6 _+_
infix 4 _≈_
field
- Carrier : Set
- _≈_ : Rel Carrier zero
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ
_+_ : Op₂ Carrier
_*_ : Op₂ Carrier
-_ : Op₁ Carrier
@@ -44,7 +47,7 @@ record AlmostCommutativeRing : Set₁ where
open IsAlmostCommutativeRing isAlmostCommutativeRing public
- commutativeSemiring : CommutativeSemiring
+ commutativeSemiring : CommutativeSemiring _ _
commutativeSemiring =
record { isCommutativeSemiring = isCommutativeSemiring }
@@ -55,7 +58,7 @@ record AlmostCommutativeRing : Set₁ where
; semiring
)
- rawRing : RawRing
+ rawRing : RawRing _
rawRing = record
{ _+_ = _+_
; _*_ = _*_
@@ -68,8 +71,9 @@ record AlmostCommutativeRing : Set₁ where
-- Homomorphisms
record _-Raw-AlmostCommutative⟶_
- (From : RawRing)
- (To : AlmostCommutativeRing) : Set where
+ {r₁ r₂ r₃}
+ (From : RawRing r₁)
+ (To : AlmostCommutativeRing r₂ r₃) : Set (r₁ ⊔ r₂ ⊔ r₃) where
private
module F = RawRing From
module T = AlmostCommutativeRing To
@@ -83,9 +87,9 @@ record _-Raw-AlmostCommutative⟶_
1-homo : Homomorphic₀ ⟦_⟧ F.1# T.1#
-raw-almostCommutative⟶
- : ∀ r →
- AlmostCommutativeRing.rawRing r -Raw-AlmostCommutative⟶ r
--raw-almostCommutative⟶ r = record
+ : ∀ {r₁ r₂} (R : AlmostCommutativeRing r₁ r₂) →
+ AlmostCommutativeRing.rawRing R -Raw-AlmostCommutative⟶ R
+-raw-almostCommutative⟶ R = record
{ ⟦_⟧ = id
; +-homo = λ _ _ → refl
; *-homo = λ _ _ → refl
@@ -93,15 +97,16 @@ record _-Raw-AlmostCommutative⟶_
; 0-homo = refl
; 1-homo = refl
}
- where open AlmostCommutativeRing r
+ where open AlmostCommutativeRing R
------------------------------------------------------------------------
-- Conversions
-- Commutative rings are almost commutative rings.
-fromCommutativeRing : CommutativeRing → AlmostCommutativeRing
-fromCommutativeRing cr = record
+fromCommutativeRing :
+ ∀ {r₁ r₂} → CommutativeRing r₁ r₂ → AlmostCommutativeRing _ _
+fromCommutativeRing CR = record
{ isAlmostCommutativeRing = record
{ isCommutativeSemiring = isCommutativeSemiring
; -‿cong = -‿cong
@@ -110,15 +115,16 @@ fromCommutativeRing cr = record
}
}
where
- open CommutativeRing cr
+ open CommutativeRing CR
import Algebra.Props.Ring as R; open R ring
import Algebra.Props.AbelianGroup as AG; open AG +-abelianGroup
-- Commutative semirings can be viewed as almost commutative rings by
-- using identity as the "almost negation".
-fromCommutativeSemiring : CommutativeSemiring → AlmostCommutativeRing
-fromCommutativeSemiring cs = record
+fromCommutativeSemiring :
+ ∀ {r₁ r₂} → CommutativeSemiring r₁ r₂ → AlmostCommutativeRing _ _
+fromCommutativeSemiring CS = record
{ -_ = id
; isAlmostCommutativeRing = record
{ isCommutativeSemiring = isCommutativeSemiring
@@ -127,4 +133,4 @@ fromCommutativeSemiring cs = record
; -‿+-comm = λ _ _ → refl
}
}
- where open CommutativeSemiring cs
+ where open CommutativeSemiring CS
diff --git a/src/Algebra/RingSolver/Lemmas.agda b/src/Algebra/RingSolver/Lemmas.agda
index 42e5bf8..2003794 100644
--- a/src/Algebra/RingSolver/Lemmas.agda
+++ b/src/Algebra/RingSolver/Lemmas.agda
@@ -2,6 +2,8 @@
-- Some boring lemmas used by the ring solver
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- Note that these proofs use all "almost commutative ring" properties
-- except for zero and -‿cong.
@@ -9,8 +11,9 @@ open import Algebra
open import Algebra.RingSolver.AlmostCommutativeRing
module Algebra.RingSolver.Lemmas
- (coeff : RawRing)
- (r : AlmostCommutativeRing)
+ {r₁ r₂ r₃}
+ (coeff : RawRing r₁)
+ (r : AlmostCommutativeRing r₂ r₃)
(morphism : coeff -Raw-AlmostCommutative⟶ r)
where
@@ -20,7 +23,7 @@ open AlmostCommutativeRing r
open import Algebra.Morphism
open _-Raw-AlmostCommutative⟶_ morphism
import Relation.Binary.EqReasoning as EqR; open EqR setoid
-open import Data.Function
+open import Function
open import Data.Product
lemma₀ : ∀ x → x + ⟦ C.0# ⟧ ≈ x
diff --git a/src/Algebra/RingSolver/Simple.agda b/src/Algebra/RingSolver/Simple.agda
index faa5ed3..077d910 100644
--- a/src/Algebra/RingSolver/Simple.agda
+++ b/src/Algebra/RingSolver/Simple.agda
@@ -2,10 +2,14 @@
-- Instantiates the ring solver with two copies of the same ring
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
open import Algebra.RingSolver.AlmostCommutativeRing
-module Algebra.RingSolver.Simple (r : AlmostCommutativeRing) where
+module Algebra.RingSolver.Simple
+ {r₁ r₂} (R : AlmostCommutativeRing r₁ r₂)
+ where
-open AlmostCommutativeRing r
-import Algebra.RingSolver as R
-open R rawRing r (-raw-almostCommutative⟶ r) public
+open AlmostCommutativeRing R
+import Algebra.RingSolver as RS
+open RS rawRing R (-raw-almostCommutative⟶ R) public
diff --git a/src/Algebra/Structures.agda b/src/Algebra/Structures.agda
index 6a1f1de..314683a 100644
--- a/src/Algebra/Structures.agda
+++ b/src/Algebra/Structures.agda
@@ -3,21 +3,25 @@
-- etc.)
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
open import Relation.Binary
module Algebra.Structures where
import Algebra.FunctionProperties as FunctionProperties
-open FunctionProperties using (Op₁; Op₂)
-import Relation.Binary.EqReasoning as EqR
-open import Data.Function
open import Data.Product
-import Level as L
+open import Function
+open import Level hiding (zero)
+import Relation.Binary.EqReasoning as EqR
+
+open FunctionProperties using (Op₁; Op₂)
------------------------------------------------------------------------
-- One binary operation
-record IsSemigroup {A} (≈ : Rel A L.zero) (∙ : Op₂ A) : Set where
+record IsSemigroup {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (∙ : Op₂ A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
isEquivalence : IsEquivalence ≈
@@ -26,7 +30,8 @@ record IsSemigroup {A} (≈ : Rel A L.zero) (∙ : Op₂ A) : Set where
open IsEquivalence isEquivalence public
-record IsMonoid {A} (≈ : Rel A L.zero) (∙ : Op₂ A) (ε : A) : Set where
+record IsMonoid {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (∙ : Op₂ A) (ε : A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
isSemigroup : IsSemigroup ≈ ∙
@@ -34,17 +39,35 @@ record IsMonoid {A} (≈ : Rel A L.zero) (∙ : Op₂ A) (ε : A) : Set where
open IsSemigroup isSemigroup public
-record IsCommutativeMonoid {A} (≈ : Rel A L.zero)
- (∙ : Op₂ A) (ε : A) : Set where
+record IsCommutativeMonoid {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (_∙_ : Op₂ A) (ε : A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
- isMonoid : IsMonoid ≈ ∙ ε
- comm : Commutative ∙
+ isSemigroup : IsSemigroup ≈ _∙_
+ identityˡ : LeftIdentity ε _∙_
+ comm : Commutative _∙_
- open IsMonoid isMonoid public
+ open IsSemigroup isSemigroup public
+
+ identity : Identity ε _∙_
+ identity = (identityˡ , identityʳ)
+ where
+ open EqR (record { isEquivalence = isEquivalence })
+
+ identityʳ : RightIdentity ε _∙_
+ identityʳ = λ x → begin
+ (x ∙ ε) ≈⟨ comm x ε ⟩
+ (ε ∙ x) ≈⟨ identityˡ x ⟩
+ x ∎
+
+ isMonoid : IsMonoid ≈ _∙_ ε
+ isMonoid = record
+ { isSemigroup = isSemigroup
+ ; identity = identity
+ }
-record IsGroup {A} (≈ : Rel A L.zero)
- (_∙_ : Op₂ A) (ε : A) (_⁻¹ : Op₁ A) : Set where
+record IsGroup {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (_∙_ : Op₂ A) (ε : A) (_⁻¹ : Op₁ A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
infixl 7 _-_
field
@@ -57,8 +80,9 @@ record IsGroup {A} (≈ : Rel A L.zero)
_-_ : Op₂ A
x - y = x ∙ (y ⁻¹)
-record IsAbelianGroup {A} (≈ : Rel A L.zero)
- (∙ : Op₂ A) (ε : A) (⁻¹ : Op₁ A) : Set where
+record IsAbelianGroup
+ {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (∙ : Op₂ A) (ε : A) (⁻¹ : Op₁ A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
isGroup : IsGroup ≈ ∙ ε ⁻¹
@@ -68,15 +92,16 @@ record IsAbelianGroup {A} (≈ : Rel A L.zero)
isCommutativeMonoid : IsCommutativeMonoid ≈ ∙ ε
isCommutativeMonoid = record
- { isMonoid = isMonoid
- ; comm = comm
+ { isSemigroup = isSemigroup
+ ; identityˡ = proj₁ identity
+ ; comm = comm
}
------------------------------------------------------------------------
-- Two binary operations
-record IsNearSemiring {A} (≈ : Rel A L.zero)
- (+ * : Op₂ A) (0# : A) : Set where
+record IsNearSemiring {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (+ * : Op₂ A) (0# : A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
+-isMonoid : IsMonoid ≈ + 0#
@@ -97,8 +122,8 @@ record IsNearSemiring {A} (≈ : Rel A L.zero)
; ∙-cong to *-cong
)
-record IsSemiringWithoutOne {A} (≈ : Rel A L.zero)
- (+ * : Op₂ A) (0# : A) : Set where
+record IsSemiringWithoutOne {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (+ * : Op₂ A) (0# : A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
+-isCommutativeMonoid : IsCommutativeMonoid ≈ + 0#
@@ -107,6 +132,7 @@ record IsSemiringWithoutOne {A} (≈ : Rel A L.zero)
zero : Zero 0# *
open IsCommutativeMonoid +-isCommutativeMonoid public
+ hiding (identityˡ)
renaming ( assoc to +-assoc
; ∙-cong to +-cong
; isSemigroup to +-isSemigroup
@@ -130,7 +156,8 @@ record IsSemiringWithoutOne {A} (≈ : Rel A L.zero)
}
record IsSemiringWithoutAnnihilatingZero
- {A} (≈ : Rel A L.zero) (+ * : Op₂ A) (0# 1# : A) : Set where
+ {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
-- Note that these structures do have an additive unit, but this
@@ -140,6 +167,7 @@ record IsSemiringWithoutAnnihilatingZero
distrib : * DistributesOver +
open IsCommutativeMonoid +-isCommutativeMonoid public
+ hiding (identityˡ)
renaming ( assoc to +-assoc
; ∙-cong to +-cong
; isSemigroup to +-isSemigroup
@@ -156,8 +184,8 @@ record IsSemiringWithoutAnnihilatingZero
; identity to *-identity
)
-record IsSemiring {A} (≈ : Rel A L.zero)
- (+ * : Op₂ A) (0# 1# : A) : Set where
+record IsSemiring {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
isSemiringWithoutAnnihilatingZero :
@@ -179,7 +207,8 @@ record IsSemiring {A} (≈ : Rel A L.zero)
using (isNearSemiring)
record IsCommutativeSemiringWithoutOne
- {A} (≈ : Rel A L.zero) (+ * : Op₂ A) (0# : A) : Set where
+ {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (+ * : Op₂ A) (0# : A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
isSemiringWithoutOne : IsSemiringWithoutOne ≈ + * 0#
@@ -188,29 +217,63 @@ record IsCommutativeSemiringWithoutOne
open IsSemiringWithoutOne isSemiringWithoutOne public
record IsCommutativeSemiring
- {A} (≈ : Rel A L.zero) (+ * : Op₂ A) (0# 1# : A) : Set where
+ {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (_+_ _*_ : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
- isSemiring : IsSemiring ≈ + * 0# 1#
- *-comm : Commutative *
+ +-isCommutativeMonoid : IsCommutativeMonoid ≈ _+_ 0#
+ *-isCommutativeMonoid : IsCommutativeMonoid ≈ _*_ 1#
+ distribʳ : _*_ DistributesOverʳ _+_
+ zeroˡ : LeftZero 0# _*_
+
+ private
+ module +-CM = IsCommutativeMonoid +-isCommutativeMonoid
+ open module *-CM = IsCommutativeMonoid *-isCommutativeMonoid public
+ using () renaming (comm to *-comm)
+ open EqR (record { isEquivalence = +-CM.isEquivalence })
+
+ distrib : _*_ DistributesOver _+_
+ distrib = (distribˡ , distribʳ)
+ where
+ distribˡ : _*_ DistributesOverˡ _+_
+ distribˡ x y z = begin
+ (x * (y + z)) ≈⟨ *-comm x (y + z) ⟩
+ ((y + z) * x) ≈⟨ distribʳ x y z ⟩
+ ((y * x) + (z * x)) ≈⟨ *-comm y x ⟨ +-CM.∙-cong ⟩ *-comm z x ⟩
+ ((x * y) + (x * z)) ∎
- open IsSemiring isSemiring public
+ zero : Zero 0# _*_
+ zero = (zeroˡ , zeroʳ)
+ where
+ zeroʳ : RightZero 0# _*_
+ zeroʳ x = begin
+ (x * 0#) ≈⟨ *-comm x 0# ⟩
+ (0# * x) ≈⟨ zeroˡ x ⟩
+ 0# ∎
- *-isCommutativeMonoid : IsCommutativeMonoid ≈ * 1#
- *-isCommutativeMonoid = record
- { isMonoid = *-isMonoid
- ; comm = *-comm
+ isSemiring : IsSemiring ≈ _+_ _*_ 0# 1#
+ isSemiring = record
+ { isSemiringWithoutAnnihilatingZero = record
+ { +-isCommutativeMonoid = +-isCommutativeMonoid
+ ; *-isMonoid = *-CM.isMonoid
+ ; distrib = distrib
}
+ ; zero = zero
+ }
+
+ open IsSemiring isSemiring public
+ hiding (distrib; zero; +-isCommutativeMonoid)
isCommutativeSemiringWithoutOne :
- IsCommutativeSemiringWithoutOne ≈ + * 0#
+ IsCommutativeSemiringWithoutOne ≈ _+_ _*_ 0#
isCommutativeSemiringWithoutOne = record
{ isSemiringWithoutOne = isSemiringWithoutOne
- ; *-comm = *-comm
+ ; *-comm = *-CM.comm
}
-record IsRing {A} (≈ : Rel A L.zero)
- (_+_ _*_ : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set where
+record IsRing
+ {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (_+_ _*_ : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
+-isAbelianGroup : IsAbelianGroup ≈ _+_ 0# -_
@@ -285,9 +348,9 @@ record IsRing {A} (≈ : Rel A L.zero)
open IsSemiring isSemiring public
using (isNearSemiring; isSemiringWithoutOne)
-record IsCommutativeRing {A}
- (≈ : Rel A L.zero)
- (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) : Set where
+record IsCommutativeRing
+ {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
isRing : IsRing ≈ + * - 0# 1#
@@ -297,14 +360,23 @@ record IsCommutativeRing {A}
isCommutativeSemiring : IsCommutativeSemiring ≈ + * 0# 1#
isCommutativeSemiring = record
- { isSemiring = isSemiring
- ; *-comm = *-comm
+ { +-isCommutativeMonoid = +-isCommutativeMonoid
+ ; *-isCommutativeMonoid = record
+ { isSemigroup = *-isSemigroup
+ ; identityˡ = proj₁ *-identity
+ ; comm = *-comm
+ }
+ ; distribʳ = proj₂ distrib
+ ; zeroˡ = proj₁ zero
}
open IsCommutativeSemiring isCommutativeSemiring public
- using (isCommutativeSemiringWithoutOne)
+ using ( *-isCommutativeMonoid
+ ; isCommutativeSemiringWithoutOne
+ )
-record IsLattice {A} (≈ : Rel A L.zero) (∨ ∧ : Op₂ A) : Set where
+record IsLattice {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (∨ ∧ : Op₂ A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
isEquivalence : IsEquivalence ≈
@@ -318,8 +390,8 @@ record IsLattice {A} (≈ : Rel A L.zero) (∨ ∧ : Op₂ A) : Set where
open IsEquivalence isEquivalence public
-record IsDistributiveLattice {A} (≈ : Rel A L.zero)
- (∨ ∧ : Op₂ A) : Set where
+record IsDistributiveLattice {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (∨ ∧ : Op₂ A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
isLattice : IsLattice ≈ ∨ ∧
@@ -327,9 +399,9 @@ record IsDistributiveLattice {A} (≈ : Rel A L.zero)
open IsLattice isLattice public
-record IsBooleanAlgebra {A}
- (≈ : Rel A L.zero)
- (∨ ∧ : Op₂ A) (¬ : Op₁ A) (⊤ ⊥ : A) : Set where
+record IsBooleanAlgebra
+ {a ℓ} {A : Set a} (≈ : Rel A ℓ)
+ (∨ ∧ : Op₂ A) (¬ : Op₁ A) (⊤ ⊥ : A) : Set (a ⊔ ℓ) where
open FunctionProperties ≈
field
isDistributiveLattice : IsDistributiveLattice ≈ ∨ ∧
diff --git a/src/Category/Applicative/Indexed.agda b/src/Category/Applicative/Indexed.agda
index ba21606..680c7e5 100644
--- a/src/Category/Applicative/Indexed.agda
+++ b/src/Category/Applicative/Indexed.agda
@@ -7,7 +7,7 @@
module Category.Applicative.Indexed where
-open import Data.Function
+open import Function
open import Data.Product
open import Category.Functor
diff --git a/src/Category/Functor.agda b/src/Category/Functor.agda
index 5963aaa..39b9311 100644
--- a/src/Category/Functor.agda
+++ b/src/Category/Functor.agda
@@ -6,7 +6,7 @@
module Category.Functor where
-open import Data.Function
+open import Function
record RawFunctor (f : Set → Set) : Set₁ where
infixl 4 _<$>_ _<$_
diff --git a/src/Category/Monad.agda b/src/Category/Monad.agda
index 72ddf28..54cc482 100644
--- a/src/Category/Monad.agda
+++ b/src/Category/Monad.agda
@@ -6,7 +6,7 @@
module Category.Monad where
-open import Data.Function
+open import Function
open import Category.Monad.Indexed
open import Data.Unit
diff --git a/src/Category/Monad/Continuation.agda b/src/Category/Monad/Continuation.agda
index 880a565..2dceb53 100644
--- a/src/Category/Monad/Continuation.agda
+++ b/src/Category/Monad/Continuation.agda
@@ -10,7 +10,7 @@ open import Category.Monad
open import Category.Monad.Indexed
open import Category.Monad.Identity
-open import Data.Function
+open import Function
------------------------------------------------------------------------
-- Delimited continuation monads
diff --git a/src/Category/Monad/Indexed.agda b/src/Category/Monad/Indexed.agda
index 86d2622..a7d4cd3 100644
--- a/src/Category/Monad/Indexed.agda
+++ b/src/Category/Monad/Indexed.agda
@@ -6,7 +6,7 @@
module Category.Monad.Indexed where
-open import Data.Function
+open import Function
open import Category.Applicative.Indexed
record RawIMonad {I : Set} (M : IFun I) : Set₁ where
diff --git a/src/Category/Monad/Partiality.agda b/src/Category/Monad/Partiality.agda
index e07923d..1e16d89 100644
--- a/src/Category/Monad/Partiality.agda
+++ b/src/Category/Monad/Partiality.agda
@@ -6,10 +6,21 @@ module Category.Monad.Partiality where
open import Coinduction
open import Category.Monad
-open import Data.Nat
open import Data.Bool
+open import Data.Nat using (ℕ; zero; suc; _+_)
+open import Data.Product as Prod
+open import Data.Sum
+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.Negation
--- The partiality monad.
+------------------------------------------------------------------------
+-- The partiality monad
data _⊥ (A : Set) : Set where
now : (x : A) → A ⊥
@@ -23,7 +34,14 @@ monad = record
where
_>>=_ : ∀ {A B} → A ⊥ → (A → B ⊥) → B ⊥
now x >>= f = f x
- later x >>= f = later (♯ ♭ x >>= f)
+ later x >>= f = later (♯ (♭ x >>= f))
+
+private module M = RawMonad monad
+
+-- Non-termination.
+
+never : {A : Set} → A ⊥
+never = later (♯ never)
-- run x for n steps peels off at most n "later" constructors from x.
@@ -39,51 +57,795 @@ isNow (now x) = true
isNow (later x) = false
------------------------------------------------------------------------
+-- Kinds
+
+-- The partiality monad comes with two forms of equality (weak and
+-- strong) and one ordering. Strong equality is stronger than the
+-- ordering, which is stronger than weak equality.
+
+-- The three relations are defined using a single data type, indexed
+-- by a "kind".
+
+data OtherKind : Set where
+ geq weak : OtherKind
+
+data Kind : Set where
+ strong : Kind
+ other : (k : OtherKind) → Kind
+
+-- Kind equality is decidable.
+
+_≟-Kind_ : Decidable (_≡_ {A = Kind})
+_≟-Kind_ strong strong = yes P.refl
+_≟-Kind_ strong (other k) = no λ()
+_≟-Kind_ (other k) strong = no λ()
+_≟-Kind_ (other geq) (other geq) = yes P.refl
+_≟-Kind_ (other geq) (other weak) = no λ()
+_≟-Kind_ (other weak) (other geq) = no λ()
+_≟-Kind_ (other weak) (other weak) = yes P.refl
+
+-- A predicate which is satisfied only for equalities. Note that, for
+-- concrete inputs, this predicate evaluates to ⊤ or ⊥.
+
+Equality : Kind → Set
+Equality k = False (k ≟-Kind other geq)
+
+------------------------------------------------------------------------
+-- Equality/ordering
+
+module Equality {A : Set} -- The "return type".
+ (_R_ : 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)
+ 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
+ _≅_ = Rel strong
+
+ _≳_ : A ⊥ → A ⊥ → Set
+ _≳_ = Rel (other geq)
+
+ _≲_ : A ⊥ → A ⊥ → Set
+ _≲_ = flip _≳_
+
+ _≈_ : A ⊥ → A ⊥ → Set
+ _≈_ = Rel (other weak)
+
+------------------------------------------------------------------------
+-- Lemmas relating the three relations
+
+ -- All relations include strong equality.
+
+ ≅⇒ : ∀ {k} {x y : A ⊥} → x ≅ y → Rel k x y
+ ≅⇒ (now xRy) = now xRy
+ ≅⇒ (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
+ ≳⇒ (later x≳y) = later (♯ ≳⇒ (♭ x≳y))
+ ≳⇒ (laterˡ x≳y) = laterˡ (≳⇒ x≳y )
+
+ -- Weak equality includes the other relations.
+
+ ⇒≈ : ∀ {k} {x y : A ⊥} → Rel k x y → x ≈ y
+ ⇒≈ {strong} = ≅⇒
+ ⇒≈ {other geq} = ≳⇒
+ ⇒≈ {other weak} = id
+
+ -- The relations agree for non-terminating computations.
+
+ never⇒never : ∀ {k₁ k₂} {x : A ⊥} →
+ Rel k₁ x never → Rel k₂ x never
+ never⇒never (later x∼never) = later (♯ never⇒never (♭ x∼never))
+ never⇒never (laterʳ x≈never) = never⇒never x≈never
+ never⇒never (laterˡ x∼never) = later (♯ never⇒never x∼never)
+
+ -- The "other" relations agree when the right-hand side is a value.
+
+ 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 (laterˡ x∼now) = laterˡ (now⇒now x∼now)
+
+------------------------------------------------------------------------
+-- Later can be dropped
+
+ laterʳ⁻¹ : ∀ {k} {x : A ⊥} {y} →
+ Rel (other k) x (later y) → Rel (other k) x (♭ y)
+ laterʳ⁻¹ (later x∼y) = laterˡ (♭ x∼y)
+ laterʳ⁻¹ (laterʳ x≈y) = x≈y
+ laterʳ⁻¹ (laterˡ x∼ly) = laterˡ (laterʳ⁻¹ x∼ly)
+
+ laterˡ⁻¹ : ∀ {x} {y : A ⊥} → later x ≈ y → ♭ x ≈ y
+ laterˡ⁻¹ (later x≈y) = laterʳ (♭ x≈y)
+ laterˡ⁻¹ (laterʳ lx≈y) = laterʳ (laterˡ⁻¹ lx≈y)
+ laterˡ⁻¹ (laterˡ x≈y) = x≈y
+
+ later⁻¹ : ∀ {k} {x y : ∞ (A ⊥)} →
+ Rel k (later x) (later y) → Rel k (♭ x) (♭ y)
+ later⁻¹ (later x∼y) = ♭ x∼y
+ later⁻¹ (laterʳ lx≈y) = laterˡ⁻¹ lx≈y
+ later⁻¹ (laterˡ x∼ly) = laterʳ⁻¹ x∼ly
+
+------------------------------------------------------------------------
+-- Terminating computations
+
+ -- 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
+
+ -- 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⇓)
+
+------------------------------------------------------------------------
+-- Non-terminating computations
+
+ -- x ⇑ means that x does not terminate.
+
+ infix 4 _⇑[_] _⇑
+
+ _⇑[_] : A ⊥ → Kind → Set
+ x ⇑[ k ] = Rel k x never
+
+ _⇑ : A ⊥ → Set
+ x ⇑ = x ⇑[ other weak ]
+
+------------------------------------------------------------------------
+-- The relations are equivalences or partial orders
+
+-- The precondition that the underlying relation is an equivalence can
+-- be weakened for some of the properties.
+
+module Properties (S : Setoid zero zero) where
+
+ private
+ open module R = Setoid S
+ using () renaming (Carrier to A; _≈_ to _R_)
+ open Equality _R_
+
+ -- All the relations are preorders.
+
+ preorder : Kind → Preorder _ _ _
+ preorder k = record
+ { Carrier = A ⊥
+ ; _≈_ = _≡_
+ ; _∼_ = Rel k
+ ; isPreorder = record
+ { isEquivalence = P.isEquivalence
+ ; reflexive = refl _
+ ; trans = trans
+ }
+ }
+ 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)
+
+ 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)
+
+ -- The two equalities are equivalence relations.
+
+ setoid : (k : Kind) {eq : Equality k} → Setoid _ _
+ setoid k {eq} = record
+ { Carrier = A ⊥
+ ; _≈_ = Rel k
+ ; isEquivalence = record
+ { refl = Pre.refl
+ ; sym = sym 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})
+
+ -- The order is a partial order, with strong equality as the
+ -- underlying equality.
+
+ ≳-poset : Poset _ _ _
+ ≳-poset = record
+ { Carrier = A ⊥
+ ; _≈_ = _≅_
+ ; _≤_ = _≳_
+ ; isPartialOrder = record
+ { antisym = antisym
+ ; isPreorder = record
+ { isEquivalence = S.isEquivalence
+ ; reflexive = ≅⇒
+ ; trans = Pre.trans
+ }
+ }
+ }
+ where
+ antisym : {x y : A ⊥} → x ≳ y → x ≲ y → x ≅ y
+ antisym (now xRy) (now _) = now xRy
+ 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))
+ antisym (laterˡ x≳ly) (laterˡ x≲ly) = later (♯ antisym (laterʳ⁻¹ x≳ly) (laterʳ⁻¹ x≲ly))
+
+ -- Equational reasoning.
+
+ module Reasoning where
+
+ infix 2 _∎
+ infixr 2 _≅⟨_⟩_ _≳⟨_⟩_ _≈⟨_⟩_
+
+ _≅⟨_⟩_ : ∀ {k} x {y z : A ⊥} → x ≅ y → Rel k y z → Rel k x z
+ _ ≅⟨ x≅y ⟩ y∼z = Pre.trans (≅⇒ x≅y) y∼z
+
+ _≳⟨_⟩_ : ∀ {k} x {y z : A ⊥} →
+ x ≳ y → Rel (other k) y z → Rel (other k) x z
+ _ ≳⟨ x≳y ⟩ y∼z = Pre.trans (≳⇒ x≳y) y∼z
+
+ _≈⟨_⟩_ : ∀ x {y z : A ⊥} → x ≈ y → y ≈ z → x ≈ z
+ _ ≈⟨ x≈y ⟩ y≈z = Pre.trans x≈y y≈z
+
+ sym : ∀ {k} {eq : Equality k} {x y : A ⊥} →
+ Rel k x y → Rel k y x
+ sym {eq = eq} = S.sym {eq = eq}
+
+ _∎ : ∀ {k} (x : A ⊥) → Rel k x x
+ x ∎ = Pre.refl
+
+------------------------------------------------------------------------
+-- Lemmas related to now and never
+
+ -- Now is not never.
+
+ 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).
+
+ now-or-never : ∀ {k} (x : A ⊥) →
+ ¬ ¬ ((∃ λ y → x ⇓[ other k ] y) ⊎ x ⇑[ other k ])
+ now-or-never 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 (later x) hyp =
+ later (♯ not-now-is-never (♭ x) (hyp ∘ Prod.map id laterˡ))
+
+ helper : Dec (∃ λ y → x ≳ now y) → _
+ helper (yes ≳now) = inj₁ $ Prod.map id ≳⇒ ≳now
+ helper (no ≵now) = inj₂ $ ≳⇒ $ not-now-is-never x ≵now
+
+------------------------------------------------------------------------
+-- Lemmas related to steps
+
+ module Steps where
+
+ open P.≡-Reasoning
+ open Reasoning using (_≅⟨_⟩_)
+
+ private
+
+ 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
+ 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)
+
+------------------------------------------------------------------------
+-- Laws related to bind
+
+ -- Never is a left and right "zero" of bind.
+
+ left-zero : {B : Set} (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
+ (x >>= λ _ → never) ≅ never
+ right-zero (now x) = S.refl
+ right-zero (later x) = later (♯ right-zero (♭ x))
+
+ -- Now is a left and right identity of bind.
+
+ left-identity : {B : Set} (x : B) (f : B → A ⊥) → let open M in
+ (now x >>= f) ≅ f x
+ left-identity x f = S.refl
+
+ right-identity : (x : A ⊥) → let open M in
+ (x >>= now) ≅ x
+ right-identity (now x) = now R.refl
+ right-identity (later x) = later (♯ right-identity (♭ x))
+
+ -- Bind is associative.
+
+ associative : {B C : Set} (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)
+
+module Properties₂ (S₁ S₂ : Setoid zero zero) 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
+
+ -- 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₂
+ 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₂)
+
+ -- Inversion lemmas for bind.
+
+ >>=-inversion-⇓ :
+ ∀ {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⇓) =
+ Prod.map id (Prod.map laterˡ (Prod.map id (P.cong suc))) $
+ >>=-inversion-⇓ (♭ x) x>>=f⇓
+
+ >>=-inversion-⇑ : ∀ {k} x {f : A → B ⊥} → let open M in
+ EqB.Rel (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
+ where
+ open RawMonad ¬¬-Monad using (_<$>_)
+ open M using (_>>=_)
+
+ 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⁻¹
+
+ 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⇑)
+
+------------------------------------------------------------------------
+-- Instantiation of the modules above using propositional equality
+
+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
+
+ -- 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.
+
+ ≡⇒ : ∀ {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)
+
+------------------------------------------------------------------------
-- Productivity checker workaround
-- The monad can be awkward to use, due to the limitations of guarded
-- coinduction. The following code provides a (limited) workaround.
-infixl 1 _>>=_
+module Workaround where
+
+ infixl 1 _>>=_
+
+ data _⊥P : Set → Set₁ 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
+ now : ∀ {A} (x : A) → A ⊥W
+ later : ∀ {A} (x : A ⊥P) → A ⊥W
+
+ mutual
+
+ _>>=W_ : ∀ {A B} → A ⊥W → (A → B ⊥P) → B ⊥W
+ now x >>=W f = whnf (f x)
+ later x >>=W f = later (x >>= f)
+
+ whnf : ∀ {A} → A ⊥P → A ⊥W
+ whnf (now x) = now x
+ whnf (later x) = later (♭ x)
+ whnf (x >>= f) = whnf x >>=W f
+
+ mutual
+
+ private
+
+ ⟦_⟧W : ∀ {A} → A ⊥W → A ⊥
+ ⟦ now x ⟧W = now x
+ ⟦ later x ⟧W = later (♯ ⟦ x ⟧P)
+
+ ⟦_⟧P : ∀ {A} → A ⊥P → A ⊥
+ ⟦ x ⟧P = ⟦ whnf x ⟧W
+
+ -- The definitions above make sense. ⟦_⟧P is homomorphic with
+ -- respect to now, later and _>>=_.
+
+ module Correct where
+
+ open Propositional
+ open Reasoning
+
+ now-hom : ∀ {A} (x : A) → ⟦ now x ⟧P ≅ now x
+ now-hom x = now x ∎
+
+ later-hom : ∀ {A} (x : ∞ (A ⊥P)) →
+ ⟦ later x ⟧P ≅ later (♯ ⟦ ♭ x ⟧P)
+ later-hom x = later (♯ (⟦ ♭ x ⟧P ∎))
+
+ mutual
+
+ private
+
+ >>=-homW : ∀ {A B} (x : B ⊥W) (f : B → A ⊥P) →
+ ⟦ x >>=W f ⟧W ≅ M._>>=_ ⟦ x ⟧W (λ y → ⟦ f y ⟧P)
+ >>=-homW (now x) f = ⟦ f x ⟧P ∎
+ >>=-homW (later x) f = later (♯ >>=-hom x f)
+
+ >>=-hom : ∀ {A B} (x : B ⊥P) (f : B → A ⊥P) →
+ ⟦ x >>= f ⟧P ≅ M._>>=_ ⟦ x ⟧P (λ y → ⟦ f y ⟧P)
+ >>=-hom x f = >>=-homW (whnf x) f
+
+------------------------------------------------------------------------
+-- An alternative, but equivalent, formulation of equality/ordering
+
+module AlternativeEquality where
+
+ private
+
+ El : Setoid zero zero → Set
+ El = Setoid.Carrier
+
+ Eq : ∀ S → B.Rel (El S) _
+ Eq = Setoid._≈_
+
+ open Equality using (Rel)
+ open Equality.Rel
+
+ infix 4 _∣_≅P_ _∣_≳P_ _∣_≈P_
+ infix 2 _∎
+ infixr 2 _≅⟨_⟩_ _≳⟨_⟩_ _≳⟨_⟩≅_ _≳⟨_⟩≈_ _≈⟨_⟩≅_ _≈⟨_⟩≲_
+ infixl 1 _>>=_
+
+ mutual
+
+ -- Proof "programs".
+
+ _∣_≅P_ : ∀ S → B.Rel (El S ⊥) _
+ _∣_≅P_ = flip RelP strong
+
+ _∣_≳P_ : ∀ S → B.Rel (El S ⊥) _
+ _∣_≳P_ = flip RelP (other geq)
+
+ _∣_≈P_ : ∀ S → B.Rel (El S ⊥) _
+ _∣_≈P_ = flip RelP (other weak)
+
+ data RelP S : Kind → B.Rel (El S ⊥) (suc zero) 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₂}
+ {f₁ f₂ : El S′ → El S ⊥} →
+ let open M in
+ (x₁∼x₂ : RelP S′ k x₁ x₂)
+ (f₁∼f₂ : ∀ {x y} → x ⟨ Eq S′ ⟩ y →
+ RelP S k (f₁ x) (f₂ y)) →
+ RelP S k (x₁ >>= f₁) (x₂ >>= f₂)
+
+ -- Ordering/weak equality.
+
+ laterʳ : ∀ {x y} (x≈y : RelP S (other weak) x (♭ y)) → RelP S (other weak) x (later y)
+ laterˡ : ∀ {k x y} (x∼y : RelP S (other k) (♭ x) y) → RelP S (other k) (later x) y
+
+ -- Equational reasoning. Note that including full transitivity
+ -- for weak equality would make _∣_≈P_ trivial; a similar
+ -- problem applies to _∣_≳P_ (A ∣ never ≳P now x would be
+ -- provable). Instead the definition of RelP includes limited
+ -- notions of transitivity, similar to weak bisimulation up-to
+ -- various things.
+
+ _∎ : ∀ {k} x → RelP S k x x
+ sym : ∀ {k x y} {eq : Equality k} (x∼y : RelP S k x y) → RelP S k y x
+ _≅⟨_⟩_ : ∀ {k} x {y z} (x≅y : S ∣ x ≅P y) (y∼z : RelP S k y z) → RelP S k x z
+ _≳⟨_⟩_ : let open Equality (Eq S) in
+ ∀ x {y z} (x≳y : x ≳ y) (y≳z : S ∣ y ≳P z) → S ∣ x ≳P z
+ _≳⟨_⟩≅_ : ∀ x {y z} (x≳y : S ∣ x ≳P y) (y≅z : S ∣ y ≅P z) → S ∣ x ≳P z
+ _≳⟨_⟩≈_ : ∀ x {y z} (x≳y : S ∣ x ≳P y) (y≈z : S ∣ y ≈P z) → S ∣ x ≈P z
+ _≈⟨_⟩≅_ : ∀ x {y z} (x≈y : S ∣ x ≈P y) (y≅z : S ∣ y ≅P z) → S ∣ x ≈P z
+ _≈⟨_⟩≲_ : ∀ x {y z} (x≈y : S ∣ x ≈P y) (y≲z : S ∣ z ≳P y) → S ∣ x ≈P z
+
+ -- If any of the following transitivity-like rules were added to
+ -- RelP, then RelP and Rel would no longer be equivalent:
+ --
+ -- x ≳P y → y ≳P z → x ≳P z
+ -- x ≳P y → y ≳ z → x ≳P z
+ -- x ≲P y → y ≈P z → x ≈P z
+ -- x ≈P y → y ≳P z → x ≈P z
+ -- x ≲ y → y ≈P z → x ≈P z
+ -- x ≈P y → y ≳ z → x ≈P z
+ -- x ≈P y → y ≈P z → x ≈P z
+ -- x ≈P y → y ≈ z → x ≈P z
+ -- x ≈ y → y ≈P z → x ≈P z
+ --
+ -- The reason is that any of these rules would make it possible
+ -- to derive that never and now x are related.
+
+ -- RelP is complete with respect to Rel.
+
+ complete : ∀ {S k} {x y : El S ⊥} →
+ Equality.Rel (Eq S) k x y → RelP S k x y
+ complete (now xRy) = now xRy
+ complete (later x∼y) = later (♯ complete (♭ x∼y))
+ complete (laterʳ x≈y) = laterʳ (complete x≈y)
+ complete (laterˡ x∼y) = laterˡ (complete x∼y)
+
+ -- RelP is sound with respect to Rel.
+
+ private
+
+ -- Proof WHNFs.
+
+ data RelW S : Kind → B.Rel (El S ⊥) (suc zero) 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)
+ laterˡ : ∀ {k x y} (x∼y : RelW S (other k) (♭ x) y) → RelW S (other k) (later x) y
+
+ -- WHNFs can be turned into programs.
+
+ program : ∀ {S k x y} → RelW S k x y → RelP S k x y
+ program (now xRy) = now xRy
+ program (later x∼y) = later (♯ x∼y)
+ program (laterˡ x∼y) = laterˡ (program x∼y)
+ program (laterʳ x≈y) = laterʳ (program x≈y)
+
+ -- Lemmas for WHNFs.
+
+ _>>=W_ : ∀ {A B k x₁ x₂} {f₁ f₂ : El A → El B ⊥} →
+ RelW A k x₁ x₂ →
+ (∀ {x y} → x ⟨ Eq A ⟩ y → RelW B k (f₁ x) (f₂ y)) →
+ RelW B k (M._>>=_ x₁ f₁) (M._>>=_ x₂ f₂)
+ now xRy >>=W f₁∼f₂ = f₁∼f₂ xRy
+ later x∼y >>=W f₁∼f₂ = later (x∼y >>= program ∘ f₁∼f₂)
+ laterʳ x≈y >>=W f₁≈f₂ = laterʳ (x≈y >>=W f₁≈f₂)
+ laterˡ x∼y >>=W f₁∼f₂ = laterˡ (x∼y >>=W f₁∼f₂)
+
+ reflW : ∀ {S k} x → RelW S k x x
+ reflW {S} (now x) = now (Setoid.refl S)
+ reflW (later x) = later (♭ x ∎)
+
+ symW : ∀ {S k x y} → Equality k → RelW S k x y → RelW S k y x
+ symW {S} eq (now xRy) = now (Setoid.sym S xRy)
+ symW eq (later x≈y) = later (sym {eq = eq} x≈y)
+ symW eq (laterʳ x≈y) = laterˡ (symW eq x≈y)
+ symW eq (laterˡ {weak} x≈y) = laterʳ (symW eq x≈y)
+ symW () (laterˡ {geq} x≈y)
+
+ trans≅W : ∀ {S x y z} →
+ RelW S strong x y → RelW S strong y z → RelW S strong x z
+ trans≅W {S} (now xRy) (now yRz) = now (Setoid.trans S xRy yRz)
+ trans≅W (later x≅y) (later y≅z) = later (_ ≅⟨ x≅y ⟩ y≅z)
+
+ trans≳-W : ∀ {S x y z} → let open Equality (Eq S) in
+ x ≳ y → RelW S (other geq) y z → RelW S (other geq) x z
+ trans≳-W {S} (now xRy) (now yRz) = now (Setoid.trans S xRy yRz)
+ trans≳-W (later x≳y) (later y≳z) = later (_ ≳⟨ ♭ x≳y ⟩ y≳z)
+ trans≳-W (later x≳y) (laterˡ y≳z) = laterˡ (trans≳-W (♭ x≳y) y≳z)
+ trans≳-W (laterˡ x≳y) y≳z = laterˡ (trans≳-W x≳y y≳z)
+
+ -- Strong equality programs can be turned into WHNFs.
+
+ whnf≅ : ∀ {S x y} → S ∣ x ≅P y → RelW S strong x y
+ whnf≅ (now xRy) = now xRy
+ whnf≅ (later x≅y) = later (♭ x≅y)
+ whnf≅ (x₁≅x₂ >>= f₁≅f₂) = whnf≅ x₁≅x₂ >>=W λ xRy → whnf≅ (f₁≅f₂ xRy)
+ whnf≅ (x ∎) = reflW x
+ whnf≅ (sym x≅y) = symW _ (whnf≅ x≅y)
+ whnf≅ (x ≅⟨ x≅y ⟩ y≅z) = trans≅W (whnf≅ x≅y) (whnf≅ y≅z)
+
+ -- More transitivity lemmas.
+
+ _⟨_⟩≅_ : ∀ {S k} x {y z} →
+ RelP S k x y → S ∣ y ≅P z → RelP S k x z
+ _⟨_⟩≅_ {k = strong} x x≅y y≅z = x ≅⟨ x≅y ⟩ y≅z
+ _⟨_⟩≅_ {k = other geq} x x≳y y≅z = x ≳⟨ x≳y ⟩≅ y≅z
+ _⟨_⟩≅_ {k = other weak} x x≈y y≅z = x ≈⟨ x≈y ⟩≅ y≅z
+
+ trans∼≅W : ∀ {S k x y z} →
+ RelW S k x y → RelW S strong y z → RelW S k x z
+ trans∼≅W {S} (now xRy) (now yRz) = now (Setoid.trans S xRy yRz)
+ trans∼≅W (later x∼y) (later y≅z) = later (_ ⟨ x∼y ⟩≅ y≅z)
+ trans∼≅W (laterʳ x≈y) (later y≅z) = laterʳ (trans∼≅W x≈y (whnf≅ y≅z))
+ trans∼≅W (laterˡ x∼y) y≅z = laterˡ (trans∼≅W x∼y y≅z)
+
+ trans≅∼W : ∀ {S k x y z} →
+ RelW S strong x y → RelW S k y z → RelW S k x z
+ trans≅∼W {S} (now xRy) (now yRz) = now (Setoid.trans S xRy yRz)
+ trans≅∼W (later x≅y) (later y∼z) = later (_ ≅⟨ x≅y ⟩ y∼z)
+ trans≅∼W (later x≅y) (laterˡ y∼z) = laterˡ (trans≅∼W (whnf≅ x≅y) y∼z)
+ trans≅∼W x≅y (laterʳ ly≈z) = laterʳ (trans≅∼W x≅y ly≈z)
+
+ -- Order programs can be turned into WHNFs.
+
+ whnf≳ : ∀ {S x y} → S ∣ x ≳P y → RelW S (other geq) x y
+ whnf≳ (now xRy) = now xRy
+ whnf≳ (later x∼y) = later (♭ x∼y)
+ whnf≳ (laterˡ x≲y) = laterˡ (whnf≳ x≲y)
+ whnf≳ (x₁∼x₂ >>= f₁∼f₂) = whnf≳ x₁∼x₂ >>=W λ xRy → whnf≳ (f₁∼f₂ xRy)
+ whnf≳ (x ∎) = reflW x
+ whnf≳ (sym {eq = ()} x≅y)
+ whnf≳ (x ≅⟨ x≅y ⟩ y≳z) = trans≅∼W (whnf≅ x≅y) (whnf≳ y≳z)
+ whnf≳ (x ≳⟨ x≳y ⟩ y≳z) = trans≳-W x≳y (whnf≳ y≳z)
+ whnf≳ (x ≳⟨ x≳y ⟩≅ y≅z) = trans∼≅W (whnf≳ x≳y) (whnf≅ y≅z)
+
+ -- Another transitivity lemma.
+
+ trans≳≈W : ∀ {S x y z} →
+ RelW S (other geq) x y → RelW S (other weak) y z →
+ RelW S (other weak) x z
+ trans≳≈W {S} (now xRy) (now yRz) = now (Setoid.trans S xRy yRz)
+ trans≳≈W (later x≳y) (later y≈z) = later (_ ≳⟨ x≳y ⟩≈ y≈z)
+ trans≳≈W (laterˡ x≳y) y≈z = laterˡ (trans≳≈W x≳y y≈z)
+ trans≳≈W x≳y (laterʳ y≈z) = laterʳ (trans≳≈W x≳y y≈z)
+ trans≳≈W (later x≳y) (laterˡ y≈z) = laterˡ (trans≳≈W (whnf≳ x≳y) y≈z)
+
+ -- All programs can be turned into WHNFs.
+
+ whnf : ∀ {S k x y} → RelP S k x y → RelW S k x y
+ whnf (now xRy) = now xRy
+ whnf (later x∼y) = later (♭ x∼y)
+ whnf (laterʳ x≈y) = laterʳ (whnf x≈y)
+ whnf (laterˡ x∼y) = laterˡ (whnf x∼y)
+ whnf (x₁∼x₂ >>= f₁∼f₂) = whnf x₁∼x₂ >>=W λ xRy → whnf (f₁∼f₂ xRy)
+ whnf (x ∎) = reflW x
+ whnf (sym {eq = eq} x≈y) = symW eq (whnf x≈y)
+ whnf (x ≅⟨ x≅y ⟩ y∼z) = trans≅∼W (whnf x≅y) (whnf y∼z)
+ whnf (x ≳⟨ x≳y ⟩ y≳z) = trans≳-W x≳y (whnf y≳z)
+ whnf (x ≳⟨ x≳y ⟩≅ y≅z) = trans∼≅W (whnf x≳y) (whnf y≅z)
+ whnf (x ≳⟨ x≳y ⟩≈ y≈z) = trans≳≈W (whnf x≳y) (whnf y≈z)
+ whnf (x ≈⟨ x≈y ⟩≅ y≅z) = trans∼≅W (whnf x≈y) (whnf y≅z)
+ whnf (x ≈⟨ x≈y ⟩≲ y≲z) = symW _ (trans≳≈W (whnf y≲z) (symW _ (whnf x≈y)))
+
+ mutual
-data _⊥-prog : Set → Set₁ where
- now : ∀ {A} (x : A) → A ⊥-prog
- later : ∀ {A} (x : ∞ (A ⊥-prog)) → A ⊥-prog
- _>>=_ : ∀ {A B} (x : A ⊥-prog) (f : A → B ⊥-prog) → B ⊥-prog
+ -- Soundness.
-data _⊥-whnf : Set → Set₁ where
- now : ∀ {A} (x : A) → A ⊥-whnf
- later : ∀ {A} (x : A ⊥-prog) → A ⊥-whnf
+ private
-whnf : ∀ {A} → A ⊥-prog → A ⊥-whnf
-whnf (now x) = now x
-whnf (later x) = later (♭ x)
-whnf (x >>= f) with whnf x
-whnf (x >>= f) | now x′ = whnf (f x′)
-whnf (x >>= f) | later x′ = later (x′ >>= f)
+ soundW : ∀ {S k x y} → RelW S k x y → Rel (Eq S) k x y
+ soundW (now xRy) = now xRy
+ soundW (later x∼y) = later (♯ sound x∼y)
+ soundW (laterʳ x≈y) = laterʳ (soundW x≈y)
+ soundW (laterˡ x∼y) = laterˡ (soundW x∼y)
-mutual
+ sound : ∀ {S k x y} → RelP S k x y → Rel (Eq S) k x y
+ sound x∼y = soundW (whnf x∼y)
- value : ∀ {A} → A ⊥-whnf → A ⊥
- value (now x) = now x
- value (later x) = later (♯ run x)
+ -- RelP and Rel are equivalent (when the underlying relation is an
+ -- equivalence).
- run : ∀ {A} → A ⊥-prog → A ⊥
- run x = value (whnf x)
+ correct : ∀ {S k x y} → RelP S k x y ⇔ Rel (Eq S) k x y
+ correct = equivalent sound complete
------------------------------------------------------------------------
--- Examples
+-- Example
-module Examples where
+private
+ module Example where
- open import Relation.Nullary
+ open Data.Nat
+ open Workaround
-- McCarthy's f91:
- f91′ : ℕ → ℕ ⊥-prog
+ f91′ : ℕ → ℕ ⊥P
f91′ n with n ≤? 100
... | yes _ = later (♯ (f91′ (11 + n) >>= f91′))
... | no _ = now (n ∸ 10)
f91 : ℕ → ℕ ⊥
- f91 n = run (f91′ n)
+ f91 n = ⟦ f91′ n ⟧P
diff --git a/src/Category/Monad/State.agda b/src/Category/Monad/State.agda
index e75fe2a..7626540 100644
--- a/src/Category/Monad/State.agda
+++ b/src/Category/Monad/State.agda
@@ -8,7 +8,7 @@ open import Category.Monad
open import Category.Monad.Indexed
open import Category.Monad.Identity
open import Data.Product
-open import Data.Function
+open import Function
open import Data.Unit
open import Category.Applicative.Indexed
diff --git a/src/Coinduction.agda b/src/Coinduction.agda
index 4334e44..3fe1592 100644
--- a/src/Coinduction.agda
+++ b/src/Coinduction.agda
@@ -1,21 +1,62 @@
------------------------------------------------------------------------
--- Types used to make recursive arguments coinductive
+-- Basic types related to coinduction
------------------------------------------------------------------------
{-# OPTIONS --universe-polymorphism #-}
--- See Data.Colist for examples of how this type is used, or
--- http://article.gmane.org/gmane.comp.lang.agda/633 for a longer
--- explanation.
-
module Coinduction where
-open import Level
+import Level
+
+------------------------------------------------------------------------
+-- A type used to make recursive arguments coinductive
+
+infix 1000 ♯_
+
+postulate
+ ∞ : ∀ {a} (A : Set a) → Set a
+ ♯_ : ∀ {a} {A : Set a} → A → ∞ A
+ ♭ : ∀ {a} {A : Set a} → ∞ A → A
+
+{-# BUILTIN INFINITY ∞ #-}
+{-# BUILTIN SHARP ♯_ #-}
+{-# BUILTIN FLAT ♭ #-}
+
+------------------------------------------------------------------------
+-- Rec, a type which is analogous to the Rec type constructor used in
+-- (the current version of) ΠΣ
+
+data Rec {a} (A : ∞ (Set a)) : Set a where
+ fold : (x : ♭ A) → Rec A
+
+unfold : ∀ {a} {A : ∞ (Set a)} → Rec A → ♭ A
+unfold (fold x) = x
+
+{-
+
+ -- If --guardedness-preserving-type-constructors is enabled one can
+ -- define types like ℕ by recursion:
+
+ open import Data.Sum
+ open import Data.Unit
+
+ ℕ : Set
+ ℕ = ⊤ ⊎ Rec (♯ ℕ)
+
+ zero : ℕ
+ zero = inj₁ _
+
+ suc : ℕ → ℕ
+ suc n = inj₂ (fold n)
-infix 10 ♯_
+ ℕ-rec : (P : ℕ → Set) →
+ P zero →
+ (∀ n → P n → P (suc n)) →
+ ∀ n → P n
+ ℕ-rec P z s (inj₁ _) = z
+ ℕ-rec P z s (inj₂ (fold n)) = s n (ℕ-rec P z s n)
-codata ∞ {a} (A : Set a) : Set a where
- ♯_ : (x : A) → ∞ A
+ -- This feature is very experimental, though: it may lead to
+ -- inconsistencies.
-♭ : ∀ {a} {A : Set a} → ∞ A → A
-♭ (♯ x) = x
+-}
diff --git a/src/Data/AVL/IndexedMap.agda b/src/Data/AVL/IndexedMap.agda
index 6110c9e..f5f2d59 100644
--- a/src/Data/AVL/IndexedMap.agda
+++ b/src/Data/AVL/IndexedMap.agda
@@ -16,7 +16,7 @@ module Data.AVL.IndexedMap
where
import Data.AVL
-open import Data.Function
+open import Function
open import Data.Maybe as Maybe
open import Data.Bool
open import Data.List as List using (List)
@@ -32,10 +32,10 @@ KV = ∃ λ i → Key i × Value i
private
- fromKV : KV → Σ (∃ Key) λ ik → Value (proj₁ ik)
+ fromKV : KV → Σ[ ik ∶ ∃ Key ] Value (proj₁ ik)
fromKV (i , k , v) = ((i , k) , v)
- toKV : Σ (∃ Key) (λ ik → Value (proj₁ ik)) → KV
+ toKV : Σ[ ik ∶ ∃ Key ] Value (proj₁ ik) → KV
toKV ((i , k) , v) = (i , k , v)
private
diff --git a/src/Data/AVL/Sets.agda b/src/Data/AVL/Sets.agda
index edca723..93fd8fa 100644
--- a/src/Data/AVL/Sets.agda
+++ b/src/Data/AVL/Sets.agda
@@ -11,7 +11,7 @@ module Data.AVL.Sets
import Data.AVL as AVL
open StrictTotalOrder OrderedKeySet renaming (Carrier to Key)
open import Data.Unit
-open import Data.Function
+open import Function
open import Data.Product as Prod using (_×_; _,_; proj₁)
open import Data.Maybe as Maybe
open import Data.Bool
diff --git a/src/Data/Bin.agda b/src/Data/Bin.agda
index 91ff853..9ef0611 100644
--- a/src/Data/Bin.agda
+++ b/src/Data/Bin.agda
@@ -12,7 +12,7 @@ open import Data.Digit
open import Data.Fin as Fin using (Fin; zero) renaming (suc to 1+_)
open import Data.Fin.Props as FP using (_+′_)
open import Data.List as List hiding (downFrom)
-open import Data.Function
+open import Function
open import Data.Product
open import Algebra
open import Relation.Binary
diff --git a/src/Data/Bool.agda b/src/Data/Bool.agda
index da468ce..d1502a4 100644
--- a/src/Data/Bool.agda
+++ b/src/Data/Bool.agda
@@ -6,7 +6,7 @@
module Data.Bool where
-open import Data.Function
+open import Function
open import Data.Unit using (⊤)
open import Data.Empty
open import Level
diff --git a/src/Data/Bool/Properties.agda b/src/Data/Bool/Properties.agda
index 288ed5d..4debaa0 100644
--- a/src/Data/Bool/Properties.agda
+++ b/src/Data/Bool/Properties.agda
@@ -6,12 +6,14 @@ module Data.Bool.Properties where
open import Data.Bool as Bool
open import Data.Fin
-open import Data.Function
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence
+ using (_⇔_; equivalent; module Equivalent)
open import Algebra
open import Algebra.Structures
import Algebra.RingSolver.Simple as Solver
import Algebra.RingSolver.AlmostCommutativeRing as ACR
-open import Relation.Nullary using (_⇔_)
open import Relation.Binary.PropositionalEquality
hiding (proof-irrelevance)
open ≡-Reasoning
@@ -50,15 +52,6 @@ private
∧-comm false true = refl
∧-comm false false = refl
- ∨-identity : Identity false _∨_
- ∨-identity = (λ _ → refl) , (λ x → ∨-comm x false)
-
- ∧-identity : Identity true _∧_
- ∧-identity = (λ _ → refl) , (λ x → ∧-comm x true)
-
- zero-∧ : Zero false _∧_
- zero-∧ = (λ _ → refl) , (λ x → ∧-comm x false)
-
distrib-∧-∨ : _∧_ DistributesOver _∨_
distrib-∧-∨ = distˡ , distʳ
where
@@ -81,35 +74,29 @@ private
isCommutativeSemiring-∨-∧
: IsCommutativeSemiring _≡_ _∨_ _∧_ false true
isCommutativeSemiring-∨-∧ = record
- { isSemiring = record
- { isSemiringWithoutAnnihilatingZero = record
- { +-isCommutativeMonoid = record
- { isMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∨-assoc
- ; ∙-cong = cong₂ _∨_
- }
- ; identity = ∨-identity
- }
- ; comm = ∨-comm
- }
- ; *-isMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∧-assoc
- ; ∙-cong = cong₂ _∧_
- }
- ; identity = ∧-identity
- }
- ; distrib = distrib-∧-∨
+ { +-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∨-assoc
+ ; ∙-cong = cong₂ _∨_
+ }
+ ; identityˡ = λ _ → refl
+ ; comm = ∨-comm
+ }
+ ; *-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∧-assoc
+ ; ∙-cong = cong₂ _∧_
}
- ; zero = zero-∧
+ ; identityˡ = λ _ → refl
+ ; comm = ∧-comm
}
- ; *-comm = ∧-comm
+ ; distribʳ = proj₂ distrib-∧-∨
+ ; zeroˡ = λ _ → refl
}
-commutativeSemiring-∨-∧ : CommutativeSemiring
+commutativeSemiring-∨-∧ : CommutativeSemiring _ _
commutativeSemiring-∨-∧ = record
{ _+_ = _∨_
; _*_ = _∧_
@@ -126,9 +113,6 @@ module RingSolver =
private
- zero-∨ : Zero true _∨_
- zero-∨ = (λ _ → refl) , (λ x → ∨-comm x true)
-
distrib-∨-∧ : _∨_ DistributesOver _∧_
distrib-∨-∧ = distˡ , distʳ
where
@@ -151,35 +135,29 @@ private
isCommutativeSemiring-∧-∨
: IsCommutativeSemiring _≡_ _∧_ _∨_ true false
isCommutativeSemiring-∧-∨ = record
- { isSemiring = record
- { isSemiringWithoutAnnihilatingZero = record
- { +-isCommutativeMonoid = record
- { isMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∧-assoc
- ; ∙-cong = cong₂ _∧_
- }
- ; identity = ∧-identity
- }
- ; comm = ∧-comm
- }
- ; *-isMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∨-assoc
- ; ∙-cong = cong₂ _∨_
- }
- ; identity = ∨-identity
- }
- ; distrib = distrib-∨-∧
+ { +-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∧-assoc
+ ; ∙-cong = cong₂ _∧_
+ }
+ ; identityˡ = λ _ → refl
+ ; comm = ∧-comm
+ }
+ ; *-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∨-assoc
+ ; ∙-cong = cong₂ _∨_
}
- ; zero = zero-∨
+ ; identityˡ = λ _ → refl
+ ; comm = ∨-comm
}
- ; *-comm = ∨-comm
+ ; distribʳ = proj₂ distrib-∨-∧
+ ; zeroˡ = λ _ → refl
}
-commutativeSemiring-∧-∨ : CommutativeSemiring
+commutativeSemiring-∧-∨ : CommutativeSemiring _ _
commutativeSemiring-∧-∨ = record
{ _+_ = _∧_
; _*_ = _∨_
@@ -240,7 +218,7 @@ isBooleanAlgebra = record
; ¬-cong = cong not
}
-booleanAlgebra : BooleanAlgebra
+booleanAlgebra : BooleanAlgebra _ _
booleanAlgebra = record
{ _∨_ = _∨_
; _∧_ = _∧_
@@ -257,9 +235,10 @@ private
xor-is-ok : ∀ x y → x xor y ≡ (x ∨ y) ∧ not (x ∧ y)
xor-is-ok true y = refl
- xor-is-ok false y = sym $ proj₂ ∧-identity _
+ xor-is-ok false y = sym $ proj₂ CS.*-identity _
+ where module CS = CommutativeSemiring commutativeSemiring-∨-∧
-commutativeRing-xor-∧ : CommutativeRing
+commutativeRing-xor-∧ : CommutativeRing _ _
commutativeRing-xor-∧ = commutativeRing
where
import Algebra.Props.BooleanAlgebra as BA
@@ -288,25 +267,25 @@ not-¬ {false} refl ()
⇔→≡ : {b₁ b₂ b : Bool} → b₁ ≡ b ⇔ b₂ ≡ b → b₁ ≡ b₂
⇔→≡ {true } {true } hyp = refl
-⇔→≡ {true } {false} {true } hyp = sym (proj₁ hyp refl)
-⇔→≡ {true } {false} {false} hyp = proj₂ hyp refl
-⇔→≡ {false} {true } {true } hyp = proj₂ hyp refl
-⇔→≡ {false} {true } {false} hyp = sym (proj₁ hyp refl)
+⇔→≡ {true } {false} {true } hyp = sym (Equivalent.to hyp ⟨$⟩ refl)
+⇔→≡ {true } {false} {false} hyp = Equivalent.from hyp ⟨$⟩ refl
+⇔→≡ {false} {true } {true } hyp = Equivalent.from hyp ⟨$⟩ refl
+⇔→≡ {false} {true } {false} hyp = sym (Equivalent.to hyp ⟨$⟩ refl)
⇔→≡ {false} {false} hyp = refl
T-≡ : ∀ {b} → T b ⇔ b ≡ true
-T-≡ {false} = ((λ ()) , λ ())
-T-≡ {true} = (const refl , const _)
+T-≡ {false} = equivalent (λ ()) (λ ())
+T-≡ {true} = equivalent (const refl) (const _)
T-∧ : ∀ {b₁ b₂} → T (b₁ ∧ b₂) ⇔ (T b₁ × T b₂)
-T-∧ {true} {true} = (const (_ , _) , const _)
-T-∧ {true} {false} = ((λ ()) , proj₂)
-T-∧ {false} {_} = ((λ ()) , proj₁)
+T-∧ {true} {true} = equivalent (const (_ , _)) (const _)
+T-∧ {true} {false} = equivalent (λ ()) proj₂
+T-∧ {false} {_} = equivalent (λ ()) proj₁
T-∨ : ∀ {b₁ b₂} → T (b₁ ∨ b₂) ⇔ (T b₁ ⊎ T b₂)
-T-∨ {true} {b₂} = (inj₁ , const _)
-T-∨ {false} {true} = (inj₂ , const _)
-T-∨ {false} {false} = (inj₁ , [ id , id ])
+T-∨ {true} {b₂} = equivalent inj₁ (const _)
+T-∨ {false} {true} = equivalent inj₂ (const _)
+T-∨ {false} {false} = equivalent inj₁ [ id , id ]
proof-irrelevance : ∀ {b} (p q : T b) → p ≡ q
proof-irrelevance {true} _ _ = refl
diff --git a/src/Data/BoundedVec.agda b/src/Data/BoundedVec.agda
index 2ee4892..b839607 100644
--- a/src/Data/BoundedVec.agda
+++ b/src/Data/BoundedVec.agda
@@ -11,7 +11,7 @@ open import Data.List as List using (List)
open import Data.Vec as Vec using (Vec)
open import Relation.Binary.PropositionalEquality
open import Data.Nat.Properties
-open SemiringSolver
+open SemiringSolver hiding (_↑)
------------------------------------------------------------------------
-- The type
diff --git a/src/Data/Colist.agda b/src/Data/Colist.agda
index 358104c..2e8be02 100644
--- a/src/Data/Colist.agda
+++ b/src/Data/Colist.agda
@@ -18,9 +18,10 @@ open import Data.BoundedVec.Inefficient as BVec
using (BoundedVec; []; _∷_)
open import Data.Product using (_,_)
open import Data.Sum using (_⊎_; inj₁; inj₂)
-open import Data.Function
+open import Function
open import Relation.Binary
import Relation.Binary.InducedPreorders as Ind
+open import Relation.Binary.PropositionalEquality using (_≡_)
open import Relation.Nullary
open import Relation.Nullary.Negation
open RawMonad ¬¬-Monad
@@ -143,9 +144,8 @@ map-cong f (x ∷ xs≈) = f x ∷ ♯ map-cong f (♭ xs≈)
infix 4 _∈_
-data _∈_ {A : Set} : A → Colist A → Set where
- here : ∀ {x xs} → x ∈ x ∷ xs
- there : ∀ {x y xs} (x∈xs : x ∈ ♭ xs) → x ∈ y ∷ xs
+_∈_ : {A : Set} → A → Colist A → Set
+x ∈ xs = Any (_≡_ x) xs
-- xs ⊆ ys means that xs is a subset of ys.
@@ -166,7 +166,7 @@ data _⊑_ {A : Set} : Colist A → Colist A → Set where
⊑⇒⊆ : {A : Set} → _⊑_ {A = A} ⇒ _⊆_
⊑⇒⊆ [] ()
-⊑⇒⊆ (x ∷ xs⊑ys) here = here
+⊑⇒⊆ (x ∷ xs⊑ys) (here ≡x) = here ≡x
⊑⇒⊆ (_ ∷ xs⊑ys) (there x∈xs) = there (⊑⇒⊆ (♭ xs⊑ys) x∈xs)
-- The prefix relation forms a poset.
@@ -181,7 +181,6 @@ data _⊑_ {A : Set} : Colist A → Colist A → Set where
{ isEquivalence = Setoid.isEquivalence (setoid A)
; reflexive = reflexive
; trans = trans
- ; ∼-resp-≈ = ((λ {_} → ⊑-resp-≈ˡ) , λ {_} → ⊑-resp-≈ʳ)
}
; antisym = antisym
}
@@ -195,14 +194,6 @@ data _⊑_ {A : Set} : Colist A → Colist A → Set where
trans [] _ = []
trans (x ∷ xs≈) (.x ∷ ys≈) = x ∷ ♯ trans (♭ xs≈) (♭ ys≈)
- ⊑-resp-≈ˡ : {xs : Colist A} → (λ ys → xs ⊑ ys) Respects _≈_
- ⊑-resp-≈ˡ _ [] = []
- ⊑-resp-≈ˡ (x ∷ xs≈) (.x ∷ p) = x ∷ ♯ ⊑-resp-≈ˡ (♭ xs≈) (♭ p)
-
- ⊑-resp-≈ʳ : {ys : Colist A} → (λ xs → xs ⊑ ys) Respects _≈_
- ⊑-resp-≈ʳ [] _ = []
- ⊑-resp-≈ʳ (x ∷ xs≈) (.x ∷ p) = x ∷ ♯ ⊑-resp-≈ʳ (♭ xs≈) (♭ p)
-
antisym : Antisymmetric _≈_ _⊑_
antisym [] [] = []
antisym (x ∷ p₁) (.x ∷ p₂) = x ∷ ♯ antisym (♭ p₁) (♭ p₂)
diff --git a/src/Data/Container.agda b/src/Data/Container.agda
new file mode 100644
index 0000000..5f89f9c
--- /dev/null
+++ b/src/Data/Container.agda
@@ -0,0 +1,260 @@
+------------------------------------------------------------------------
+-- Containers, based on the work of Abbott and others
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Data.Container where
+
+open import Data.Product as Prod hiding (map)
+open import Function renaming (id to ⟨id⟩; _∘_ to _⟨∘⟩_)
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Inverse as Inv using (_⇿_; module Inverse)
+open import Level
+open import Relation.Binary using (Setoid; module Setoid)
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; _≗_; refl)
+open import Relation.Unary using (_⊆_)
+
+------------------------------------------------------------------------
+-- Containers
+
+-- A container is a set of shapes, and for every shape a set of
+-- positions.
+
+infix 5 _◃_
+
+record Container (ℓ : Level) : Set (suc ℓ) where
+ constructor _◃_
+ field
+ Shape : Set ℓ
+ Position : Shape → Set ℓ
+
+open Container public
+
+-- The semantics ("extension") of a container.
+
+⟦_⟧ : ∀ {ℓ} → Container ℓ → Set ℓ → Set ℓ
+⟦ C ⟧ X = Σ[ s ∶ Shape C ] (Position C s → X)
+
+-- Equality (based on propositional equality).
+
+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))
+
+private
+
+ -- Note that, if propositional equality were extensional, then _≈_
+ -- 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′
+ 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)) →
+ _≡_ {A = ⟦ C ⟧ X} (s , f) (s′ , f′)
+ helper refl f≈f′ = P.cong (_,_ _) (ext f≈f′)
+
+setoid : ∀ {ℓ} → Container ℓ → Set ℓ → Setoid ℓ ℓ
+setoid C X = record
+ { Carrier = ⟦ C ⟧ X
+ ; _≈_ = _≈_
+ ; isEquivalence = record
+ { refl = (refl , λ _ → refl)
+ ; sym = sym
+ ; trans = λ {_ _ zs} → trans zs
+ }
+ }
+ where
+ sym : {xs ys : ⟦ C ⟧ X} → 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)
+
+ trans : ∀ {xs ys : ⟦ C ⟧ X} 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))
+
+------------------------------------------------------------------------
+-- Functoriality
+
+-- Containers are functors.
+
+map : ∀ {c} {C : Container c} {X Y} → (X → Y) → ⟦ C ⟧ X → ⟦ C ⟧ Y
+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 _)
+
+ 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 _)
+
+------------------------------------------------------------------------
+-- Container morphisms
+
+-- Representation of container morphisms.
+
+record _⇒_ {c} (C₁ C₂ : Container c) : Set c where
+ field
+ shape : Shape C₁ → Shape C₂
+ position : ∀ {s} → Position C₂ (shape s) → Position C₁ s
+
+open _⇒_ public
+
+-- Interpretation of _⇒_.
+
+⟪_⟫ : ∀ {c} {C₁ C₂ : Container c} →
+ C₁ ⇒ C₂ → ∀ {X} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X
+⟪ m ⟫ xs = (shape m (proj₁ xs) , proj₂ xs ⟨∘⟩ position m)
+
+module Morphism where
+
+ -- Naturality.
+
+ 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 transformations.
+
+ NT : ∀ {c} (C₁ C₂ : Container c) → Set (suc c)
+ NT C₁ C₂ = ∃ λ (m : ∀ {X} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X) → Natural m
+
+ -- Container morphisms are natural.
+
+ natural : ∀ {c} {C₁ C₂ : Container c}
+ (m : C₁ ⇒ C₂) → Natural ⟪ m ⟫
+ natural {C₂ = C₂} m f xs = Setoid.refl (setoid C₂ _)
+
+ -- 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⟩))
+ where
+ m = record { shape = λ s → proj₁ (nt (s , ⟨id⟩))
+ ; position = λ {s} → proj₂ (nt (s , ⟨id⟩))
+ }
+
+ -- Identity.
+
+ id : ∀ {c} (C : Container c) → C ⇒ C
+ id _ = record {shape = ⟨id⟩; position = ⟨id⟩}
+
+ -- Composition.
+
+ infixr 9 _∘_
+
+ _∘_ : ∀ {c} {C₁ C₂ C₃ : Container c} → C₂ ⇒ C₃ → C₁ ⇒ C₂ → C₁ ⇒ C₃
+ f ∘ g = record
+ { shape = shape f ⟨∘⟩ shape g
+ ; position = position g ⟨∘⟩ position f
+ }
+
+ -- Identity and composition commute with ⟪_⟫.
+
+ id-correct : ∀ {c} {C : Container c} {X : Set c} →
+ ⟪ id C ⟫ {X} ≗ ⟨id⟩
+ id-correct xs = refl
+
+ ∘-correct : ∀ {c} {C₁ C₂ C₃ : Container c}
+ (f : C₂ ⇒ C₃) (g : C₁ ⇒ C₂) {X : Set c} →
+ ⟪ f ∘ g ⟫ {X} ≗ (⟪ f ⟫ ⟨∘⟩ ⟪ g ⟫)
+ ∘-correct f g xs = refl
+
+------------------------------------------------------------------------
+-- Linear container morphisms
+
+record _⊸_ {c} (C₁ C₂ : Container c) : Set c where
+ field
+ shape⊸ : Shape C₁ → Shape C₂
+ position⊸ : ∀ {s} → Position C₂ (shape⊸ s) ⇿ Position C₁ s
+
+ morphism : C₁ ⇒ C₂
+ morphism = record
+ { shape = shape⊸
+ ; position = _⟨$⟩_ (Inverse.to position⊸)
+ }
+
+ ⟪_⟫⊸ : ∀ {X} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X
+ ⟪_⟫⊸ = ⟪ morphism ⟫
+
+open _⊸_ public using (shape⊸; position⊸; ⟪_⟫⊸)
+
+------------------------------------------------------------------------
+-- All and any
+
+-- All.
+
+□ : ∀ {c} {C : Container c} {X : Set c} →
+ (X → Set c) → (⟦ C ⟧ X → Set c)
+□ P (s , f) = ∀ p → P (f p)
+
+□-map : ∀ {c} {C : Container c} {X : Set c} {P Q : X → Set c} →
+ P ⊆ Q → □ {C = C} P ⊆ □ Q
+□-map P⊆Q = _⟨∘⟩_ P⊆Q
+
+-- Any.
+
+◇ : ∀ {c} {C : Container c} {X : Set c} →
+ (X → Set c) → (⟦ C ⟧ X → Set c)
+◇ P (s , f) = ∃ λ p → P (f p)
+
+◇-map : ∀ {c} {C : Container c} {X : Set c} {P Q : X → Set c} →
+ P ⊆ Q → ◇ {C = C} P ⊆ ◇ Q
+◇-map P⊆Q = Prod.map ⟨id⟩ P⊆Q
+
+-- Membership.
+
+infix 4 _∈_
+
+_∈_ : ∀ {c} {C : Container c} {X : Set c} →
+ X → ⟦ C ⟧ X → Set c
+x ∈ xs = ◇ (_≡_ x) xs
+
+-- Bag and set equality. Two containers xs and ys are equal when
+-- viewed as sets if, whenever x ∈ xs, we also have x ∈ ys, and vice
+-- versa. They are equal when viewed as bags if, additionally, the
+-- sets x ∈ xs and x ∈ ys have the same size. For alternative but
+-- equivalent definitions of bag and set equality, see
+-- Data.Container.AlternativeBagAndSetEquality.
+
+open Inv public
+ using (Kind) renaming (inverse to bag; equivalent to set)
+
+[_]-Equality : ∀ {ℓ} → Kind → Container ℓ → Set ℓ → Setoid ℓ ℓ
+[ k ]-Equality C X = Inv.InducedEquivalence₂ k (_∈_ {C = C} {X = X})
+
+infix 4 _≈[_]_
+
+_≈[_]_ : ∀ {c} {C : Container c} {X : Set c} →
+ ⟦ C ⟧ X → Kind → ⟦ C ⟧ X → Set c
+xs ≈[ k ] ys = Setoid._≈_ ([ k ]-Equality _ _) xs ys
diff --git a/src/Data/Container/AlternativeBagAndSetEquality.agda b/src/Data/Container/AlternativeBagAndSetEquality.agda
new file mode 100644
index 0000000..6ddcc2e
--- /dev/null
+++ b/src/Data/Container/AlternativeBagAndSetEquality.agda
@@ -0,0 +1,145 @@
+------------------------------------------------------------------------
+-- 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
new file mode 100644
index 0000000..612ee4e
--- /dev/null
+++ b/src/Data/Container/Any.agda
@@ -0,0 +1,260 @@
+------------------------------------------------------------------------
+-- Properties related to ◇
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Data.Container.Any where
+
+open import Algebra
+open import Data.Container as C
+open import Data.Container.Combinator
+ using (module Composition) renaming (_∘_ to _⟨∘⟩_)
+open import Data.Product as Prod hiding (swap)
+open import Data.Sum
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Inverse as Inv
+ using (_⇿_; Isomorphism; module Inverse)
+open import Function.Inverse.TypeIsomorphisms
+import Relation.Binary.HeterogeneousEquality as H
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; _≗_; refl)
+import Relation.Binary.Sigma.Pointwise as Σ
+
+open Inv.EquationalReasoning
+private
+ module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
+
+-- ◇ can be expressed using _∈_.
+
+⇿∈ : ∀ {c} (C : Container c) {X : Set c}
+ {P : X → Set c} {xs : ⟦ C ⟧ X} →
+ ◇ P xs ⇿ (∃ λ x → x ∈ xs × P x)
+⇿∈ _ {P = P} {xs} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl
+ ; right-inverse-of = to∘from
+ }
+ }
+ where
+ to : ◇ P xs → ∃ λ x → x ∈ xs × P x
+ to (p , Px) = (proj₂ xs p , (p , refl) , Px)
+
+ from : (∃ λ x → x ∈ xs × P x) → ◇ P xs
+ from (.(proj₂ xs p) , (p , refl) , Px) = (p , Px)
+
+ to∘from : to ∘ from ≗ id
+ to∘from (.(proj₂ xs p) , (p , refl) , Px) = refl
+
+-- ◇ is a congruence for bag and set equality.
+
+cong : ∀ {k c} {C : Container c}
+ {X : Set c} {P₁ P₂ : X → Set c} {xs₁ xs₂ : ⟦ C ⟧ X} →
+ (∀ x → Isomorphism k (P₁ x) (P₂ x)) → xs₁ ≈[ k ] xs₂ →
+ 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) ⇿⟨ sym (⇿∈ C) ⟩
+ ◇ P₂ xs₂ ∎
+
+-- Nested occurrences of ◇ can sometimes be swapped.
+
+swap : ∀ {c} {C₁ C₂ : Container c} {X Y : Set c} {P : X → Y → Set c}
+ {xs : ⟦ C₁ ⟧ X} {ys : ⟦ C₂ ⟧ Y} →
+ let ◈ : ∀ {C : Container c} {X} → ⟦ C ⟧ X → (X → Set c) → Set c
+ ◈ = flip ◇ in
+ ◈ 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 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} →
+ (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 → y ∈ ys × ◇ (flip P y) xs) ⇿⟨ sym (⇿∈ C₂) ⟩
+ ◇ (λ y → ◇ (flip P y) xs) ys ∎
+
+-- Nested occurrences of ◇ can sometimes be flattened.
+
+flatten : ∀ {c} {C₁ C₂ : Container c} {X}
+ P (xss : ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)) →
+ ◇ (◇ P) xss ⇿
+ ◇ P (Inverse.from (Composition.correct C₁ C₂) ⟨$⟩ xss)
+flatten {C₁ = C₁} {C₂} {X} P xss = record
+ { to = P.→-to-⟶ t
+ ; from = P.→-to-⟶ f
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+ where
+ open Inverse
+
+ t : ◇ (◇ P) xss → ◇ P (from (Composition.correct C₁ C₂) ⟨$⟩ xss)
+ t (p₁ , p₂ , p) = ((p₁ , p₂) , p)
+
+ f : ◇ P (from (Composition.correct C₁ C₂) ⟨$⟩ xss) → ◇ (◇ P) xss
+ f ((p₁ , p₂) , p) = (p₁ , p₂ , p)
+
+-- Sums commute with ◇ (for a fixed instance of a given container).
+
+◇⊎⇿⊎◇ : ∀ {c} {C : Container c} {X : Set c} {xs : ⟦ C ⟧ X}
+ {P Q : X → Set c} →
+ ◇ (λ x → P x ⊎ Q x) xs ⇿ (◇ P xs ⊎ ◇ Q xs)
+◇⊎⇿⊎◇ {xs = xs} {P} {Q} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = from∘to
+ ; right-inverse-of = [ (λ _ → refl) , (λ _ → refl) ]
+ }
+ }
+ where
+ to : ◇ (λ x → P x ⊎ Q x) xs → ◇ P xs ⊎ ◇ Q xs
+ to (pos , inj₁ p) = inj₁ (pos , p)
+ to (pos , inj₂ q) = inj₂ (pos , q)
+
+ from : ◇ P xs ⊎ ◇ Q xs → ◇ (λ x → P x ⊎ Q x) xs
+ from = [ Prod.map id inj₁ , Prod.map id inj₂ ]
+
+ from∘to : from ∘ to ≗ id
+ from∘to (pos , inj₁ p) = refl
+ from∘to (pos , inj₂ q) = refl
+
+-- Products "commute" with ◇.
+
+×◇⇿◇◇× : ∀ {c} {C₁ C₂ : Container c}
+ {X Y} {P : X → Set c} {Q : Y → Set c}
+ {xs : ⟦ C₁ ⟧ X} {ys : ⟦ C₂ ⟧ Y} →
+ ◇ (λ x → ◇ (λ y → P x × Q y) ys) xs ⇿ (◇ P xs × ◇ Q ys)
+×◇⇿◇◇× {C₁ = C₁} {C₂} {P = P} {Q} {xs} {ys} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+ where
+ to : ◇ (λ x → ◇ (λ y → P x × Q y) ys) xs → ◇ P xs × ◇ Q ys
+ to (p₁ , p₂ , p , q) = ((p₁ , p) , (p₂ , q))
+
+ from : ◇ P xs × ◇ Q ys → ◇ (λ x → ◇ (λ y → P x × Q y) ys) xs
+ from ((p₁ , p) , (p₂ , q)) = (p₁ , p₂ , p , q)
+
+-- map can be absorbed by the predicate.
+
+map⇿∘ : ∀ {c} (C : Container c) {X Y : Set c}
+ (P : Y → Set c) {xs : ⟦ C ⟧ X} (f : X → Y) →
+ ◇ P (C.map f xs) ⇿ ◇ (P ∘ f) xs
+map⇿∘ _ _ _ = Inv.id
+
+-- Membership in a mapped container can be expressed without reference
+-- to map.
+
+∈map⇿∈×≡ : ∀ {c} (C : Container c) {X Y : Set c} {f : X → Y}
+ {xs : ⟦ C ⟧ X} {y} →
+ y ∈ C.map f xs ⇿ (∃ λ x → x ∈ xs × y ≡ f x)
+∈map⇿∈×≡ {c} C {f = f} {xs} {y} =
+ y ∈ C.map f xs ⇿⟨ map⇿∘ C (_≡_ y) f ⟩
+ ◇ (λ x → y ≡ f x) xs ⇿⟨ ⇿∈ C ⟩
+ (∃ λ x → x ∈ xs × y ≡ f x) ∎
+
+-- map is a congruence for bag and set equality.
+
+map-cong : ∀ {k c} {C : Container c} {X Y : Set c}
+ {f₁ f₂ : X → Y} {xs₁ xs₂ : ⟦ C ⟧ X} →
+ f₁ ≗ f₂ → xs₁ ≈[ k ] xs₂ →
+ C.map f₁ xs₁ ≈[ k ] C.map f₂ xs₂
+map-cong {c = c} {C} {f₁ = f₁} {f₂} {xs₁} {xs₂} f₁≗f₂ xs₁≈xs₂ {x} =
+ x ∈ C.map f₁ xs₁ ⇿⟨ map⇿∘ C (_≡_ x) f₁ ⟩
+ ◇ (λ y → x ≡ f₁ y) xs₁ ≈⟨ cong {xs₁ = xs₁} {xs₂ = xs₂} (Inv.⇿⇒ ∘ helper) xs₁≈xs₂ ⟩
+ ◇ (λ y → x ≡ f₂ y) xs₂ ⇿⟨ sym (map⇿∘ C (_≡_ x) f₂) ⟩
+ x ∈ C.map f₂ xs₂ ∎
+ where
+ helper : ∀ y → (x ≡ f₁ y) ⇿ (x ≡ f₂ y)
+ helper y = record
+ { to = P.→-to-⟶ (λ x≡f₁y → P.trans x≡f₁y ( f₁≗f₂ y))
+ ; from = P.→-to-⟶ (λ x≡f₂y → P.trans x≡f₂y (P.sym $ f₁≗f₂ y))
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.proof-irrelevance _ _
+ ; right-inverse-of = λ _ → P.proof-irrelevance _ _
+ }
+ }
+
+-- Uses of linear morphisms can be removed.
+
+remove-linear :
+ ∀ {c} {C₁ C₂ : Container c} {X} {xs : ⟦ C₁ ⟧ X}
+ (P : X → Set c) (m : C₁ ⊸ C₂) →
+ ◇ P (⟪ m ⟫⊸ xs) ⇿ ◇ P xs
+remove-linear {xs = xs} P m = record
+ { to = P.→-to-⟶ t
+ ; from = P.→-to-⟶ f
+ ; inverse-of = record
+ { left-inverse-of = f∘t
+ ; right-inverse-of = t∘f
+ }
+ }
+ where
+ open Inverse
+
+ t : ◇ P (⟪ m ⟫⊸ xs) → ◇ P xs
+ t = Prod.map (_⟨$⟩_ (to (position⊸ m))) id
+
+ f : ◇ P xs → ◇ P (⟪ m ⟫⊸ xs)
+ f = Prod.map (_⟨$⟩_ (from (position⊸ m)))
+ (P.subst (P ∘ proj₂ xs)
+ (P.sym $ right-inverse-of (position⊸ m) _))
+
+ f∘t : f ∘ t ≗ id
+ f∘t (p₂ , p) = H.≅-to-≡ $
+ H.cong₂ _,_ (H.≡-to-≅ $ left-inverse-of (position⊸ m) p₂)
+ (H.≡-subst-removable
+ (P ∘ proj₂ xs)
+ (P.sym (right-inverse-of (position⊸ m)
+ (to (position⊸ m) ⟨$⟩ p₂)))
+ p)
+
+ t∘f : t ∘ f ≗ id
+ t∘f (p₁ , p) = H.≅-to-≡ $
+ H.cong₂ _,_ (H.≡-to-≅ $ right-inverse-of (position⊸ m) p₁)
+ (H.≡-subst-removable
+ (P ∘ proj₂ xs)
+ (P.sym (right-inverse-of (position⊸ m) p₁))
+ p)
+
+-- Linear endomorphisms are identity functions if bag equality is
+-- used.
+
+linear-identity :
+ ∀ {c} {C : Container c} {X} {xs : ⟦ C ⟧ X} (m : C ⊸ C) →
+ ⟪ m ⟫⊸ xs ≈[ bag ] xs
+linear-identity {xs = xs} m {x} =
+ x ∈ ⟪ m ⟫⊸ xs ⇿⟨ remove-linear (_≡_ x) m ⟩
+ x ∈ xs ∎
+
+-- If join can be expressed using a linear morphism (in a certain
+-- way), then it can be absorbed by the predicate.
+
+join⇿◇ : ∀ {c} {C₁ C₂ C₃ : Container c} {X}
+ P (join′ : (C₁ ⟨∘⟩ C₂) ⊸ C₃) (xss : ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)) →
+ let join : ∀ {X} → ⟦ C₁ ⟧ (⟦ C₂ ⟧ X) → ⟦ C₃ ⟧ X
+ join = ⟪ join′ ⟫⊸ ∘
+ _⟨$⟩_ (Inverse.from (Composition.correct C₁ C₂)) in
+ ◇ P (join xss) ⇿ ◇ (◇ P) xss
+join⇿◇ {C₁ = C₁} {C₂} P join xss =
+ ◇ P (⟪ join ⟫⊸ xss′) ⇿⟨ remove-linear P join ⟩
+ ◇ P xss′ ⇿⟨ sym $ flatten P xss ⟩
+ ◇ (◇ P) xss ∎
+ where xss′ = Inverse.from (Composition.correct C₁ C₂) ⟨$⟩ xss
diff --git a/src/Data/Container/Combinator.agda b/src/Data/Container/Combinator.agda
new file mode 100644
index 0000000..eadfef0
--- /dev/null
+++ b/src/Data/Container/Combinator.agda
@@ -0,0 +1,222 @@
+------------------------------------------------------------------------
+-- Container combinators
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Data.Container.Combinator where
+
+open import Data.Container
+open import Data.Empty using (⊥; ⊥-elim)
+open import Data.Product as Prod hiding (Σ) renaming (_×_ to _⟨×⟩_)
+open import Data.Sum renaming (_⊎_ to _⟨⊎⟩_)
+open import Data.Unit using (⊤)
+open import Function as F hiding (id; const) renaming (_∘_ to _⟨∘⟩_)
+open import Function.Inverse using (_⇿_)
+open import Level
+open import Relation.Binary.PropositionalEquality as P
+ using (_≗_; refl)
+
+------------------------------------------------------------------------
+-- Combinators
+
+-- Identity.
+
+id : ∀ {c} → Container c
+id = Lift ⊤ ◃ F.const (Lift ⊤)
+
+-- Constant.
+
+const : ∀ {c} → Set c → Container c
+const X = X ◃ F.const (Lift ⊥)
+
+-- Composition.
+
+infixr 9 _∘_
+
+_∘_ : ∀ {c} → Container c → Container c → Container c
+C₁ ∘ C₂ = ⟦ C₁ ⟧ (Shape C₂) ◃ ◇ (Position C₂)
+
+-- Product. (Note that, up to isomorphism, this is a special case of
+-- indexed product.)
+
+infixr 2 _×_
+
+_×_ : ∀ {c} → Container c → Container c → Container c
+C₁ × C₂ =
+ (Shape C₁ ⟨×⟩ Shape C₂) ◃
+ uncurry (λ s₁ s₂ → Position C₁ s₁ ⟨⊎⟩ Position C₂ s₂)
+
+-- Indexed product.
+
+Π : ∀ {c} {I : Set c} → (I → Container c) → Container c
+Π C = (∀ i → Shape (C i)) ◃ λ s → ∃ λ i → Position (C i) (s i)
+
+-- Sum. (Note that, up to isomorphism, this is a special case of
+-- indexed sum.)
+
+infixr 1 _⊎_
+
+_⊎_ : ∀ {c} → Container c → Container c → Container c
+C₁ ⊎ C₂ = (Shape C₁ ⟨⊎⟩ Shape C₂) ◃ [ Position C₁ , Position C₂ ]
+
+-- Indexed sum.
+
+Σ : ∀ {c} {I : Set c} → (I → Container c) → Container c
+Σ C = (∃ λ i → Shape (C i)) ◃ λ s → Position (C (proj₁ s)) (proj₂ s)
+
+-- Constant exponentiation. (Note that this is a special case of
+-- indexed product.)
+
+infix 0 const[_]⟶_
+
+const[_]⟶_ : ∀ {c} → Set c → Container c → Container c
+const[ X ]⟶ C = Π {I = X} (F.const C)
+
+------------------------------------------------------------------------
+-- Correctness proofs
+
+-- I have proved some of the correctness statements under the
+-- assumption of functional extensionality. I could have reformulated
+-- the statements using suitable setoids, but I could not be bothered.
+
+module Identity where
+
+ correct : ∀ {c} {X : Set c} → ⟦ id ⟧ X ⇿ F.id X
+ correct {c} = record
+ { to = P.→-to-⟶ {a = c} λ xs → proj₂ xs _
+ ; from = P.→-to-⟶ {b₁ = c} λ x → (_ , λ _ → x)
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+
+module Constant (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
+
+ correct : ∀ {ℓ} (X : Set ℓ) {Y} → ⟦ const X ⟧ Y ⇿ F.const X Y
+ correct X {Y} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { right-inverse-of = λ _ → refl
+ ; left-inverse-of =
+ λ xs → P.cong (_,_ (proj₁ xs)) (ext (λ x → ⊥-elim (lower x)))
+ }
+ }
+ where
+ to : ⟦ const X ⟧ Y → X
+ to = proj₁
+
+ from : X → ⟦ const X ⟧ Y
+ from = < F.id , F.const (⊥-elim ∘′ lower) >
+
+module Composition where
+
+ correct : ∀ {c} (C₁ C₂ : Container c) {X : Set c} →
+ ⟦ C₁ ∘ C₂ ⟧ X ⇿ (⟦ C₁ ⟧ ⟨∘⟩ ⟦ C₂ ⟧) X
+ correct C₁ C₂ {X} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+ where
+ to : ⟦ C₁ ∘ C₂ ⟧ X → ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)
+ to ((s , f) , g) = (s , < f , curry g >)
+
+ from : ⟦ C₁ ⟧ (⟦ C₂ ⟧ X) → ⟦ C₁ ∘ C₂ ⟧ X
+ from (s , f) = ((s , proj₁ ⟨∘⟩ f) , uncurry (proj₂ ⟨∘⟩ f))
+
+module Product (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
+
+ correct : ∀ {c} (C₁ C₂ : Container c) {X : Set c} →
+ ⟦ C₁ × C₂ ⟧ X ⇿ (⟦ C₁ ⟧ X ⟨×⟩ ⟦ C₂ ⟧ X)
+ correct {c} C₁ C₂ {X} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = from∘to
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+ where
+ to : ⟦ C₁ × C₂ ⟧ X → ⟦ C₁ ⟧ X ⟨×⟩ ⟦ C₂ ⟧ X
+ to ((s₁ , s₂) , f) = ((s₁ , f ⟨∘⟩ inj₁) , (s₂ , f ⟨∘⟩ inj₂))
+
+ from : ⟦ C₁ ⟧ X ⟨×⟩ ⟦ C₂ ⟧ X → ⟦ C₁ × C₂ ⟧ X
+ from ((s₁ , f₁) , (s₂ , f₂)) = ((s₁ , s₂) , [ f₁ , f₂ ])
+
+ from∘to : from ⟨∘⟩ to ≗ F.id
+ from∘to (s , f) =
+ P.cong (_,_ s) (ext {ℓ = c} [ (λ _ → refl) , (λ _ → refl) ])
+
+module IndexedProduct where
+
+ correct : ∀ {c I} (C : I → Container c) {X : Set c} →
+ ⟦ Π C ⟧ X ⇿ (∀ i → ⟦ C i ⟧ X)
+ correct {I = I} C {X} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+ where
+ to : ⟦ Π C ⟧ X → ∀ i → ⟦ C i ⟧ X
+ to (s , f) = λ i → (s i , λ p → f (i , p))
+
+ from : (∀ i → ⟦ C i ⟧ X) → ⟦ Π C ⟧ X
+ from f = (proj₁ ⟨∘⟩ f , uncurry (proj₂ ⟨∘⟩ f))
+
+module Sum where
+
+ correct : ∀ {c} (C₁ C₂ : Container c) {X : Set c} →
+ ⟦ C₁ ⊎ C₂ ⟧ X ⇿ (⟦ C₁ ⟧ X ⟨⊎⟩ ⟦ C₂ ⟧ X)
+ correct C₁ C₂ {X} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = from∘to
+ ; right-inverse-of = [ (λ _ → refl) , (λ _ → refl) ]
+ }
+ }
+ where
+ to : ⟦ C₁ ⊎ C₂ ⟧ X → ⟦ C₁ ⟧ X ⟨⊎⟩ ⟦ C₂ ⟧ X
+ to (inj₁ s₁ , f) = inj₁ (s₁ , f)
+ to (inj₂ s₂ , f) = inj₂ (s₂ , f)
+
+ from : ⟦ C₁ ⟧ X ⟨⊎⟩ ⟦ C₂ ⟧ X → ⟦ C₁ ⊎ C₂ ⟧ X
+ from = [ Prod.map inj₁ F.id , Prod.map inj₂ F.id ]
+
+ from∘to : from ⟨∘⟩ to ≗ F.id
+ from∘to (inj₁ s₁ , f) = refl
+ from∘to (inj₂ s₂ , f) = refl
+
+module IndexedSum where
+
+ correct : ∀ {c I} (C : I → Container c) {X : Set c} →
+ ⟦ Σ C ⟧ X ⇿ (∃ λ i → ⟦ C i ⟧ X)
+ correct {I = I} C {X} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+ where
+ to : ⟦ Σ C ⟧ X → ∃ λ i → ⟦ C i ⟧ X
+ to ((i , s) , f) = (i , (s , f))
+
+ from : (∃ λ i → ⟦ C i ⟧ X) → ⟦ Σ C ⟧ X
+ from (i , (s , f)) = ((i , s) , f)
+
+module ConstantExponentiation where
+
+ correct : ∀ {c X} (C : Container c) {Y : Set c} →
+ ⟦ const[ X ]⟶ C ⟧ Y ⇿ (X → ⟦ C ⟧ Y)
+ correct C = IndexedProduct.correct (F.const C)
diff --git a/src/Data/Covec.agda b/src/Data/Covec.agda
index ee6757e..dc95a8f 100644
--- a/src/Data/Covec.agda
+++ b/src/Data/Covec.agda
@@ -127,7 +127,6 @@ poset A n = record
{ isEquivalence = Setoid.isEquivalence (setoid A n)
; reflexive = reflexive
; trans = trans
- ; ∼-resp-≈ = ((λ {_} → ⊑-resp-≈ˡ) , λ {_} → ⊑-resp-≈ʳ)
}
; antisym = antisym
}
@@ -141,16 +140,6 @@ poset A n = record
trans [] _ = []
trans (x ∷ xs≈) (.x ∷ ys≈) = x ∷ ♯ trans (♭ xs≈) (♭ ys≈)
- ⊑-resp-≈ˡ : ∀ {A n} {xs : Covec A n} →
- (λ ys → xs ⊑ ys) Respects _≈_ {A} {n}
- ⊑-resp-≈ˡ _ [] = []
- ⊑-resp-≈ˡ (x ∷ xs≈) (.x ∷ p) = x ∷ ♯ ⊑-resp-≈ˡ (♭ xs≈) (♭ p)
-
- ⊑-resp-≈ʳ : ∀ {A n} {ys : Covec A n} →
- (λ xs → xs ⊑ ys) Respects _≈_ {A} {n}
- ⊑-resp-≈ʳ [] _ = []
- ⊑-resp-≈ʳ (x ∷ xs≈) (.x ∷ p) = x ∷ ♯ ⊑-resp-≈ʳ (♭ xs≈) (♭ p)
-
antisym : ∀ {A n} → Antisymmetric (_≈_ {A} {n}) _⊑_
antisym [] [] = []
antisym (x ∷ p₁) (.x ∷ p₂) = x ∷ ♯ antisym (♭ p₁) (♭ p₂)
diff --git a/src/Data/DifferenceList.agda b/src/Data/DifferenceList.agda
index d51cfa3..41692b4 100644
--- a/src/Data/DifferenceList.agda
+++ b/src/Data/DifferenceList.agda
@@ -5,7 +5,7 @@
module Data.DifferenceList where
open import Data.List as L using (List)
-open import Data.Function
+open import Function
open import Data.Nat
infixr 5 _∷_ _++_
diff --git a/src/Data/DifferenceNat.agda b/src/Data/DifferenceNat.agda
index 1545385..5e8f259 100644
--- a/src/Data/DifferenceNat.agda
+++ b/src/Data/DifferenceNat.agda
@@ -6,7 +6,7 @@
module Data.DifferenceNat where
open import Data.Nat as N using (ℕ)
-open import Data.Function
+open import Function
infixl 6 _+_
diff --git a/src/Data/DifferenceVec.agda b/src/Data/DifferenceVec.agda
index 13fee17..2718e54 100644
--- a/src/Data/DifferenceVec.agda
+++ b/src/Data/DifferenceVec.agda
@@ -6,7 +6,7 @@ module Data.DifferenceVec where
open import Data.DifferenceNat
open import Data.Vec as V using (Vec)
-open import Data.Function
+open import Function
import Data.Nat as N
infixr 5 _∷_ _++_
diff --git a/src/Data/Digit.agda b/src/Data/Digit.agda
index ea5168d..9a2576c 100644
--- a/src/Data/Digit.agda
+++ b/src/Data/Digit.agda
@@ -16,7 +16,7 @@ open import Induction.Nat
open import Data.Nat.DivMod
open ≤-Reasoning
open import Relation.Binary.PropositionalEquality
-open import Data.Function
+open import Function
------------------------------------------------------------------------
-- A boring lemma
diff --git a/src/Data/Fin.agda b/src/Data/Fin.agda
index 869200d..9fba6ea 100644
--- a/src/Data/Fin.agda
+++ b/src/Data/Fin.agda
@@ -12,7 +12,7 @@ open import Data.Nat as Nat
using (ℕ; zero; suc; z≤n; s≤s)
renaming ( _+_ to _N+_; _∸_ to _N∸_
; _≤_ to _N≤_; _<_ to _N<_; _≤?_ to _N≤?_)
-open import Data.Function
+open import Function
open import Level hiding (lift)
open import Relation.Nullary.Decidable
open import Relation.Binary
diff --git a/src/Data/Fin/Dec.agda b/src/Data/Fin/Dec.agda
index 2381131..7b50302 100644
--- a/src/Data/Fin/Dec.agda
+++ b/src/Data/Fin/Dec.agda
@@ -6,7 +6,7 @@
module Data.Fin.Dec where
-open import Data.Function
+open import Function
open import Data.Nat hiding (_<_)
open import Data.Vec hiding (_∈_)
open import Data.Fin
@@ -40,7 +40,7 @@ private
any? : ∀ {n} {P : Fin n → Set} →
((f : Fin n) → Dec (P f)) →
Dec (∃ P)
-any? {zero} {P} dec = no ((¬ Fin 0 ∶ λ()) ∘ proj₁)
+any? {zero} {P} dec = no (((λ()) ∶ ¬ Fin 0) ∘ proj₁)
any? {suc n} {P} dec with dec zero | any? (restrict dec)
... | yes p | _ = yes (_ , p)
... | _ | yes (_ , p') = yes (_ , p')
diff --git a/src/Data/Fin/Props.agda b/src/Data/Fin/Props.agda
index 603b98c..d52926e 100644
--- a/src/Data/Fin/Props.agda
+++ b/src/Data/Fin/Props.agda
@@ -13,9 +13,9 @@ open import Data.Nat as N
renaming (_≤_ to _ℕ≤_; _<_ to _ℕ<_; _+_ to _ℕ+_)
open N.≤-Reasoning
import Data.Nat.Properties as N
-open import Data.Function
-open import Data.Function.Equality as FunS using (_⟨$⟩_)
-open import Data.Function.Injection
+open import Function
+open import Function.Equality as FunS using (_⟨$⟩_)
+open import Function.Injection
using (Injection; module Injection)
open import Relation.Nullary
open import Relation.Binary
@@ -29,7 +29,7 @@ open import Category.Applicative
private
drop-suc : ∀ {o} {m n : Fin o} →
- suc m ≡ (Fin (suc o) ∶ suc n) → m ≡ n
+ suc m ≡ (suc n ∶ Fin (suc o)) → m ≡ n
drop-suc refl = refl
preorder : ℕ → Preorder _ _ _
diff --git a/src/Data/Fin/Subset/Props.agda b/src/Data/Fin/Subset/Props.agda
index cbac082..7f8022a 100644
--- a/src/Data/Fin/Subset/Props.agda
+++ b/src/Data/Fin/Subset/Props.agda
@@ -7,7 +7,7 @@ module Data.Fin.Subset.Props where
open import Data.Nat
open import Data.Vec hiding (_∈_)
open import Data.Empty
-open import Data.Function
+open import Function
open import Data.Fin
open import Data.Fin.Subset
open import Data.Product
diff --git a/src/Data/Fin/Substitution.agda b/src/Data/Fin/Substitution.agda
index 97f4d07..be0ed71 100644
--- a/src/Data/Fin/Substitution.agda
+++ b/src/Data/Fin/Substitution.agda
@@ -14,7 +14,7 @@ module Data.Fin.Substitution where
open import Data.Nat
open import Data.Fin using (Fin; zero; suc)
open import Data.Vec
-open import Data.Function as Fun using (flip)
+open import Function as Fun using (flip)
open import Data.Star as Star using (Star; ε; _◅_)
------------------------------------------------------------------------
diff --git a/src/Data/Fin/Substitution/Lemmas.agda b/src/Data/Fin/Substitution/Lemmas.agda
index 20035a6..67ee5d1 100644
--- a/src/Data/Fin/Substitution/Lemmas.agda
+++ b/src/Data/Fin/Substitution/Lemmas.agda
@@ -9,7 +9,7 @@ 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 Data.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
diff --git a/src/Data/Fin/Substitution/List.agda b/src/Data/Fin/Substitution/List.agda
index 22c0219..0afa5c4 100644
--- a/src/Data/Fin/Substitution/List.agda
+++ b/src/Data/Fin/Substitution/List.agda
@@ -12,7 +12,7 @@ module Data.Fin.Substitution.List {T} (lemmas₄ : Lemmas₄ T) where
open import Data.List
open import Data.List.Properties
open import Data.Fin.Substitution
-import Data.Function as Fun
+import Function as Fun
open import Relation.Binary.PropositionalEquality
open ≡-Reasoning
diff --git a/src/Data/Function/Equality.agda b/src/Data/Function/Equality.agda
deleted file mode 100644
index 6a487a6..0000000
--- a/src/Data/Function/Equality.agda
+++ /dev/null
@@ -1,108 +0,0 @@
-------------------------------------------------------------------------
--- Function setoids and related constructions
-------------------------------------------------------------------------
-
-{-# OPTIONS --universe-polymorphism #-}
-
-module Data.Function.Equality where
-
-open import Data.Function as Fun using (_on_)
-open import Level
-open import Relation.Binary
-
-------------------------------------------------------------------------
--- Functions which preserve equality
-
-infixr 0 _⟶_
-
-record _⟶_ {f₁ f₂ t₁ t₂} (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) :
- Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
- open Setoid
- infixl 5 _⟨$⟩_
- field
- _⟨$⟩_ : Carrier From → Carrier To
- cong : _⟨$⟩_ Preserves _≈_ From ⟶ _≈_ To
-
-open _⟶_ public
-
-id : ∀ {a₁ a₂} {A : Setoid a₁ a₂} → A ⟶ A
-id = record { _⟨$⟩_ = Fun.id; cong = Fun.id }
-
-infixr 9 _∘_
-
-_∘_ : ∀ {a₁ a₂} {A : Setoid a₁ a₂}
- {b₁ b₂} {B : Setoid b₁ b₂}
- {c₁ c₂} {C : Setoid c₁ c₂} →
- B ⟶ C → A ⟶ B → A ⟶ C
-f ∘ g = record
- { _⟨$⟩_ = Fun._∘_ (_⟨$⟩_ f) (_⟨$⟩_ g)
- ; cong = Fun._∘_ (cong f) (cong g)
- }
-
-const : ∀ {a₁ a₂} {A : Setoid a₁ a₂}
- {b₁ b₂} {B : Setoid b₁ b₂} →
- Setoid.Carrier B → A ⟶ B
-const {B = B} b = record
- { _⟨$⟩_ = Fun.const b
- ; cong = Fun.const (Setoid.refl B)
- }
-
-------------------------------------------------------------------------
--- A logical relation (i.e. a relation which relates functions which
--- map related things to related things)
-
-infixr 0 _↝_
-
-_↝_ : ∀ {a b c d ℓ₁ ℓ₂}
- {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
- (∼₁ : REL A B ℓ₁) (∼₂ : REL C D ℓ₂) → REL (A → C) (B → D) _
-_∼₁_ ↝ _∼₂_ = λ f g → ∀ {x y} → x ∼₁ y → f x ∼₂ g y
-
-↝-isEquivalence : ∀ {a b c ℓ₁ ℓ₂} {A : Set a} {B : Set b} {C : Set c}
- {∼₁ : Rel A ℓ₁} {∼₂ : Rel B ℓ₂}
- (fun : C → (A → B)) →
- (∀ f → fun f Preserves ∼₁ ⟶ ∼₂) →
- IsEquivalence ∼₁ → IsEquivalence ∼₂ →
- IsEquivalence ((∼₁ ↝ ∼₂) on fun)
-↝-isEquivalence _ cong eq₁ eq₂ = record
- { refl = λ {f} x∼₁y → cong f x∼₁y
- ; sym = λ f∼g x∼y → sym eq₂ (f∼g (sym eq₁ x∼y))
- ; trans = λ f∼g g∼h x∼y → trans eq₂ (f∼g (refl eq₁)) (g∼h x∼y)
- } where open IsEquivalence
-
--- A generalised variant of (_↝_ _≡_).
-
-≡↝ : ∀ {a b c ℓ} {A : Set a} {B : A → Set b} {C : A → Set c} →
- (∀ x → REL (B x) (C x) ℓ) →
- REL ((x : A) → B x) ((x : A) → C x) _
-≡↝ R = λ f g → ∀ x → R x (f x) (g x)
-
-≡↝-isEquivalence :
- ∀ {a b ℓ} {A : Set a} {B : A → Set b} {R : ∀ x → Rel (B x) ℓ} →
- (∀ x → IsEquivalence (R x)) → IsEquivalence (≡↝ R)
-≡↝-isEquivalence eq = record
- { refl = λ _ → refl
- ; sym = λ f∼g x → sym (f∼g x)
- ; trans = λ f∼g g∼h x → trans (f∼g x) (g∼h x)
- } where open module Eq {x} = IsEquivalence (eq x)
-
-------------------------------------------------------------------------
--- Function setoids
-
-infixr 0 _⇨_ _≡⇨_
-
-_⇨_ : ∀ {f₁ f₂ t₁ t₂} → Setoid f₁ f₂ → Setoid t₁ t₂ → Setoid _ _
-S₁ ⇨ S₂ = record
- { Carrier = S₁ ⟶ S₂
- ; _≈_ = (_≈_ S₁ ↝ _≈_ S₂) on _⟨$⟩_
- ; isEquivalence =
- ↝-isEquivalence (_⟨$⟩_ {From = S₁} {To = S₂})
- cong (isEquivalence S₁) (isEquivalence S₂)
- } where open Setoid; open _⟶_
-
-_≡⇨_ : ∀ {a s₁ s₂} (A : Set a) → (A → Setoid s₁ s₂) → Setoid _ _
-A ≡⇨ S = record
- { Carrier = (x : A) → Carrier (S x)
- ; _≈_ = ≡↝ (λ x → _≈_ (S x))
- ; isEquivalence = ≡↝-isEquivalence (λ x → isEquivalence (S x))
- } where open Setoid
diff --git a/src/Data/Function/LeftInverse.agda b/src/Data/Function/LeftInverse.agda
deleted file mode 100644
index 061352f..0000000
--- a/src/Data/Function/LeftInverse.agda
+++ /dev/null
@@ -1,69 +0,0 @@
-------------------------------------------------------------------------
--- Left inverses
-------------------------------------------------------------------------
-
-{-# OPTIONS --universe-polymorphism #-}
-
-module Data.Function.LeftInverse where
-
-open import Data.Product
-open import Level
-import Relation.Binary.EqReasoning as EqReasoning
-open import Relation.Binary
-open import Data.Function.Equality as F
- using (_⟶_; _⇨_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_)
-open import Data.Function.Injection using (Injective; Injection)
-
-_LeftInverseOf_ : ∀ {a₁ a₂ b₁ b₂} {A : Setoid a₁ a₂} {B : Setoid b₁ b₂} →
- B ⟶ A → A ⟶ B → Set _
-_LeftInverseOf_ {A = A} f g = ∀ x → f ⟨$⟩ (g ⟨$⟩ x) ≈₁ x
- where open Setoid A renaming (_≈_ to _≈₁_)
-
-record LeftInverse {f₁ f₂ t₁ t₂}
- (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) :
- Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
- field
- to : From ⟶ To
- from : To ⟶ From
- left-inverse : from LeftInverseOf to
-
- open Setoid From
- open EqReasoning From
-
- injective : Injective to
- injective {x} {y} eq = begin
- x ≈⟨ sym (left-inverse x) ⟩
- from ⟨$⟩ (to ⟨$⟩ x) ≈⟨ F.cong from eq ⟩
- from ⟨$⟩ (to ⟨$⟩ y) ≈⟨ left-inverse y ⟩
- y ∎
-
- injection : Injection From To
- injection = record { to = to; injective = injective }
-
-Surjective : ∀ {f₁ f₂ t₁ t₂} {F : Setoid f₁ f₂} {T : Setoid t₁ t₂} →
- F ⟶ T → Set _
-Surjective f = ∃ λ g → f LeftInverseOf g
-
-infixr 9 _∘_
-
-id : ∀ {s₁ s₂} {S : Setoid s₁ s₂} → LeftInverse S S
-id {S = S} = record
- { to = F.id
- ; from = F.id
- ; left-inverse = λ _ → Setoid.refl S
- }
-
-_∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂}
- {F : Setoid f₁ f₂} {M : Setoid m₁ m₂} {T : Setoid t₁ t₂} →
- LeftInverse M T → LeftInverse F M → LeftInverse F T
-_∘_ {F = F} f g = record
- { to = to f ⟪∘⟫ to g
- ; from = from g ⟪∘⟫ from f
- ; left-inverse = λ x → begin
- from g ⟨$⟩ (from f ⟨$⟩ (to f ⟨$⟩ (to g ⟨$⟩ x))) ≈⟨ F.cong (from g) (left-inverse f (to g ⟨$⟩ x)) ⟩
- from g ⟨$⟩ (to g ⟨$⟩ x) ≈⟨ left-inverse g x ⟩
- x ∎
- }
- where
- open LeftInverse
- open EqReasoning F
diff --git a/src/Data/Graph/Acyclic.agda b/src/Data/Graph/Acyclic.agda
index aa6ea4e..68e0a70 100644
--- a/src/Data/Graph/Acyclic.agda
+++ b/src/Data/Graph/Acyclic.agda
@@ -15,7 +15,7 @@ open import Data.Fin as Fin
open import Data.Fin.Props as FP using (_≟_)
open import Data.Product as Prod using (∃; _×_; _,_)
open import Data.Maybe
-open import Data.Function
+open import Function
open import Data.Empty
open import Data.Unit using (⊤; tt)
open import Data.Vec as Vec using (Vec; []; _∷_)
diff --git a/src/Data/Integer.agda b/src/Data/Integer.agda
index de03f0b..daaade4 100644
--- a/src/Data/Integer.agda
+++ b/src/Data/Integer.agda
@@ -10,7 +10,7 @@ import Data.Nat.Show as ℕ
open import Data.Sign as Sign using (Sign) renaming (_*_ to _S*_)
open import Data.Product as Prod
open import Data.String using (String; _++_)
-open import Data.Function
+open import Function
open import Data.Sum as Sum
open import Relation.Nullary
import Relation.Nullary.Decidable as Dec
@@ -208,10 +208,10 @@ drop‿-≤- : ∀ {m n} → -[1+ m ] ≤ -[1+ n ] → ℕ._≤_ n m
drop‿-≤- (-≤- n≤m) = n≤m
_≤?_ : Decidable _≤_
--[1+ m ] ≤? -[1+ n ] = Dec.map (-≤- , drop‿-≤-) (ℕ._≤?_ n m)
+-[1+ m ] ≤? -[1+ n ] = Dec.map′ -≤- drop‿-≤- (ℕ._≤?_ n m)
-[1+ m ] ≤? + n = yes -≤+
+ m ≤? -[1+ n ] = no λ ()
-+ m ≤? + n = Dec.map (+≤+ , drop‿+≤+) (ℕ._≤?_ m n)
++ m ≤? + n = Dec.map′ +≤+ drop‿+≤+ (ℕ._≤?_ m n)
decTotalOrder : DecTotalOrder _ _ _
decTotalOrder = record
@@ -225,7 +225,6 @@ decTotalOrder = record
{ isEquivalence = PropEq.isEquivalence
; reflexive = refl′
; trans = trans
- ; ∼-resp-≈ = PropEq.resp₂ _≤_
}
; antisym = antisym
}
diff --git a/src/Data/Integer/Divisibility.agda b/src/Data/Integer/Divisibility.agda
index 374bc9d..7f033dc 100644
--- a/src/Data/Integer/Divisibility.agda
+++ b/src/Data/Integer/Divisibility.agda
@@ -4,7 +4,7 @@
module Data.Integer.Divisibility where
-open import Data.Function
+open import Function
open import Data.Integer
open import Data.Integer.Properties
import Data.Nat.Divisibility as ℕ
diff --git a/src/Data/Integer/Properties.agda b/src/Data/Integer/Properties.agda
index adc0e86..af023eb 100644
--- a/src/Data/Integer/Properties.agda
+++ b/src/Data/Integer/Properties.agda
@@ -7,7 +7,7 @@ module Data.Integer.Properties where
open import Algebra
import Algebra.Morphism as Morphism
open import Data.Empty
-open import Data.Function
+open import Function
open import Data.Integer hiding (suc)
open import Data.Nat as ℕ renaming (_*_ to _ℕ*_)
import Data.Nat.Properties as ℕ
diff --git a/src/Data/List.agda b/src/Data/List.agda
index f3f63f5..8f4c7ce 100644
--- a/src/Data/List.agda
+++ b/src/Data/List.agda
@@ -10,7 +10,7 @@ open import Data.Sum as Sum using (_⊎_; inj₁; inj₂)
open import Data.Bool
open import Data.Maybe using (Maybe; nothing; just)
open import Data.Product as Prod using (_×_; _,_)
-open import Data.Function
+open import Function
open import Algebra
import Relation.Binary.PropositionalEquality as PropEq
import Algebra.FunctionProperties as FunProp
@@ -252,7 +252,7 @@ partition p (x ∷ xs) with p x | partition p xs
------------------------------------------------------------------------
-- List monoid
-monoid : Set → Monoid
+monoid : ∀ {ℓ} → Set ℓ → Monoid _ _
monoid A = record
{ Carrier = List A
; _≈_ = _≡_
diff --git a/src/Data/List/All.agda b/src/Data/List/All.agda
index bc1cc70..5caedfd 100644
--- a/src/Data/List/All.agda
+++ b/src/Data/List/All.agda
@@ -4,11 +4,10 @@
module Data.List.All where
-open import Data.Function
+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 Data.Product as Prod using (_,_)
open import Relation.Nullary
import Relation.Nullary.Decidable as Dec
open import Relation.Unary using () renaming (_⊆_ to _⋐_)
@@ -45,5 +44,5 @@ all : ∀ {A} {P : A → Set} →
(∀ x → Dec (P x)) → (xs : List A) → Dec (All P xs)
all p [] = yes []
all p (x ∷ xs) with p x
-all p (x ∷ xs) | yes px = Dec.map (_∷_ px , tail) (all p xs)
+all p (x ∷ xs) | yes px = Dec.map′ (_∷_ px) tail (all p xs)
all p (x ∷ xs) | no ¬px = no (¬px ∘ head)
diff --git a/src/Data/List/All/Properties.agda b/src/Data/List/All/Properties.agda
index daadccb..accf2cb 100644
--- a/src/Data/List/All/Properties.agda
+++ b/src/Data/List/All/Properties.agda
@@ -6,7 +6,9 @@ module Data.List.All.Properties where
open import Data.Bool
open import Data.Bool.Properties
-open import Data.Function
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence using (module Equivalent)
open import Data.List as List
import Data.List.Any as Any
open Any.Membership-≡
@@ -37,12 +39,12 @@ gmap g = All-map ∘ All.map g
All-all : ∀ {A} (p : A → Bool) {xs} →
All (T ∘ p) xs → T (all p xs)
All-all p [] = _
-All-all p (px ∷ pxs) = proj₂ T-∧ (px , All-all p pxs)
+All-all p (px ∷ pxs) = Equivalent.from T-∧ ⟨$⟩ (px , All-all p pxs)
all-All : ∀ {A} (p : A → Bool) xs →
T (all p xs) → All (T ∘ p) xs
all-All p [] _ = []
-all-All p (x ∷ xs) px∷xs with proj₁ (T-∧ {p x}) px∷xs
+all-All p (x ∷ xs) px∷xs with Equivalent.to (T-∧ {p x}) ⟨$⟩ px∷xs
all-All p (x ∷ xs) px∷xs | (px , pxs) = px ∷ all-All p xs pxs
-- All is anti-monotone.
diff --git a/src/Data/List/Any.agda b/src/Data/List/Any.agda
index 99974aa..60f3237 100644
--- a/src/Data/List/Any.agda
+++ b/src/Data/List/Any.agda
@@ -6,7 +6,11 @@ module Data.List.Any where
open import Data.Empty
open import Data.Fin
-open import Data.Function
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence as Equiv using (module Equivalent)
+open import Function.Inverse as Inv
+ using (module Inverse)
open import Data.List as List using (List; []; _∷_)
open import Data.Product as Prod using (∃; _×_; _,_)
open import Level
@@ -18,7 +22,6 @@ import Relation.Binary.InducedPreorders as Ind
open import Relation.Binary.List.Pointwise as ListEq using ([]; _∷_)
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_)
-import Relation.Binary.Props.Preorder as PP
-- Any P xs means that at least one element in xs satisfies P.
@@ -40,12 +43,12 @@ tail ¬px (there pxs) = pxs
-- Decides Any.
-any : ∀ {A} {P : A → Set} →
+any : {A : Set} {P : A → Set} →
(∀ x → Dec (P x)) → (xs : List A) → Dec (Any P xs)
any p [] = no λ()
any p (x ∷ xs) with p x
any p (x ∷ xs) | yes px = yes (here px)
-any p (x ∷ xs) | no ¬px = Dec.map (there , tail ¬px) (any p xs)
+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.
@@ -118,12 +121,6 @@ module Membership (S : Setoid zero zero) where
_∈⟨_⟩_ : ∀ x {xs ys} → x ∈ xs → xs IsRelatedTo ys → x ∈ ys
x ∈⟨ x∈xs ⟩ xs⊆ys = (begin xs⊆ys) x∈xs
- -- Set equality, i.e. an equality which ignores order and
- -- multiplicity.
-
- set-equality : Setoid _ _
- set-equality = PP.inducedEquivalence ⊆-preorder
-
-- A variant of List.map.
map-with-∈ : ∀ {B : Set} (xs : List A) → (∀ {x} → x ∈ xs → B) → List B
@@ -140,40 +137,61 @@ module Membership (S : Setoid zero zero) where
lose resp x∈xs px = map (flip resp px) x∈xs
-- The code above instantiated (and slightly changed) for
--- propositional equality.
+-- propositional equality, along with some additional definitions.
-module Membership-≡ {A : Set} where
+module Membership-≡ where
private
- open module M = Membership (PropEq.setoid A) public
- hiding ( lift-resp; lose
- ; ⊆-preorder; module ⊆-Reasoning; set-equality
- )
+ open module M {A : Set} = Membership (PropEq.setoid A) public
+ hiding (lift-resp; lose; ⊆-preorder; module ⊆-Reasoning)
- lose : ∀ {P x xs} → x ∈ xs → P x → Any P xs
- lose {P} = M.lose (PropEq.subst P)
+ lose : ∀ {A} {P : A → Set} {x xs} → x ∈ xs → P x → Any P xs
+ lose {P = P} = M.lose (PropEq.subst P)
-- _⊆_ is a preorder.
- ⊆-preorder : Preorder _ _ _
- ⊆-preorder = Ind.InducedPreorder₂ (PropEq.setoid (List A)) _∈_
- (PropEq.subst (_∈_ _))
+ ⊆-preorder : Set → Preorder _ _ _
+ ⊆-preorder A = Ind.InducedPreorder₂ (PropEq.setoid (List A)) _∈_
+ (PropEq.subst (_∈_ _))
+
+ -- Set and bag equality.
+
+ open Inv public
+ using (Kind) renaming (equivalent to set; inverse to bag)
+
+ [_]-Equality : Kind → Set → Setoid _ _
+ [ k ]-Equality A = Inv.InducedEquivalence₂ k (_∈_ {A = A})
+
+ infix 4 _≈[_]_
+
+ _≈[_]_ : {A : Set} → List A → Kind → List A → Set
+ xs ≈[ k ] ys = Setoid._≈_ ([ k ]-Equality _) xs ys
+
+ -- Bag equality implies set equality.
+
+ bag-=⇒set-= : {A : Set} {xs ys : List A} →
+ xs ≈[ bag ] ys → xs ≈[ set ] ys
+ bag-=⇒set-= xs≈ys = Inv.⇿⇒ xs≈ys
+
+ -- "Equational" reasoning for _⊆_.
module ⊆-Reasoning where
import Relation.Binary.PreorderReasoning as PreR
- open PreR ⊆-preorder public
- renaming (_∼⟨_⟩_ to _⊆⟨_⟩_; _≈⟨_⟩_ to _≡⟨_⟩_)
+ private
+ open module PR {A : Set} = PreR (⊆-preorder A) public
+ renaming (_∼⟨_⟩_ to _⊆⟨_⟩_; _≈⟨_⟩_ to _≡⟨_⟩_)
- infix 1 _∈⟨_⟩_
+ infixr 2 _≈⟨_⟩_
+ infix 1 _∈⟨_⟩_
- _∈⟨_⟩_ : ∀ x {xs ys} → x ∈ xs → xs IsRelatedTo ys → x ∈ ys
+ _∈⟨_⟩_ : ∀ {A : Set} x {xs ys : List A} →
+ x ∈ xs → xs IsRelatedTo ys → x ∈ ys
x ∈⟨ x∈xs ⟩ xs⊆ys = (begin xs⊆ys) x∈xs
- -- Set equality, i.e. an equality which ignores order and
- -- multiplicity.
-
- set-equality : Setoid _ _
- set-equality = PP.inducedEquivalence ⊆-preorder
+ _≈⟨_⟩_ : ∀ {k} {A : Set} xs {ys zs : List A} →
+ xs ≈[ k ] ys → ys IsRelatedTo zs → xs IsRelatedTo zs
+ xs ≈⟨ xs≈ys ⟩ ys≈zs =
+ xs ⊆⟨ _⟨$⟩_ (Equivalent.to (Inv.⇒⇔ xs≈ys)) ⟩ ys≈zs
------------------------------------------------------------------------
-- Another function
diff --git a/src/Data/List/Any/BagAndSetEquality.agda b/src/Data/List/Any/BagAndSetEquality.agda
new file mode 100644
index 0000000..9a1a916
--- /dev/null
+++ b/src/Data/List/Any/BagAndSetEquality.agda
@@ -0,0 +1,167 @@
+------------------------------------------------------------------------
+-- Properties related to bag and set equality
+------------------------------------------------------------------------
+
+-- Bag and set equality are defined in Data.List.Any.
+
+module Data.List.Any.BagAndSetEquality where
+
+open import Algebra
+open import Algebra.FunctionProperties
+open import Category.Monad
+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.Properties
+open import Data.Product
+open import Data.Sum
+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 Relation.Binary
+import Relation.Binary.EqReasoning as EqR
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; _≗_)
+
+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 ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
+
+-- _++_ and [] form a commutative monoid, with either bag or set
+-- equality as the underlying equality.
+
+commutativeMonoid : Kind → Set → CommutativeMonoid _ _
+commutativeMonoid k A = record
+ { Carrier = List A
+ ; _≈_ = λ xs ys → xs ≈[ k ] ys
+ ; _∙_ = _++_
+ ; ε = []
+ ; isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = Eq.isEquivalence
+ ; 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₁ ⊎ x ∈ xs₃) ≈⟨ xs₁≈xs₂ ⟨ ×⊎.+-cong ⟩ xs₃≈xs₄ ⟩
+ (x ∈ xs₂ ⊎ x ∈ xs₄) ⇿⟨ ++⇿ ⟩
+ x ∈ xs₂ ++ xs₄ ∎
+ }
+ ; identityˡ = λ xs {x} → x ∈ xs ∎
+ ; comm = λ xs ys {x} →
+ x ∈ xs ++ ys ⇿⟨ ++⇿++ xs ys ⟩
+ x ∈ ys ++ xs ∎
+ }
+ }
+ where open Inv.EquationalReasoning
+
+-- _++_ is idempotent (under set equality).
+
+++-idempotent : {A : Set} →
+ Idempotent (λ (xs ys : List A) → xs ≈[ set ] ys) _++_
+++-idempotent xs {x} =
+ x ∈ xs ++ xs ≈⟨ FE.equivalent ([ id , id ]′ ∘ _⟨$⟩_ (Inverse.from ++⇿))
+ (_⟨$⟩_ (Inverse.to ++⇿) ∘ inj₁) ⟩
+ x ∈ xs ∎
+ where open Inv.EquationalReasoning
+
+-- List.map is a congruence.
+
+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⇿ ⟩
+ Any (λ y → x ≡ f₁ y) xs₁ ≈⟨ Any-cong (Inv.⇿⇒ ∘ helper) xs₁≈xs₂ ⟩
+ Any (λ y → x ≡ f₂ y) xs₂ ⇿⟨ map⇿ ⟩
+ x ∈ List.map f₂ xs₂ ∎
+ where
+ open Inv.EquationalReasoning
+
+ helper : ∀ y → x ≡ f₁ y ⇿ x ≡ f₂ y
+ helper y = record
+ { to = P.→-to-⟶ (λ x≡f₁y → P.trans x≡f₁y ( f₁≗f₂ y))
+ ; from = P.→-to-⟶ (λ x≡f₂y → P.trans x≡f₂y (P.sym $ f₁≗f₂ y))
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.proof-irrelevance _ _
+ ; right-inverse-of = λ _ → P.proof-irrelevance _ _
+ }
+ }
+
+-- concat is a congruence.
+
+concat-cong : ∀ {k} {A : Set} {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⇿ ⟩
+ Any (Any (_≡_ x)) xss₁ ≈⟨ Any-cong (λ _ → _ ∎) xss₁≈xss₂ ⟩
+ Any (Any (_≡_ x)) xss₂ ⇿⟨ concat⇿ ⟩
+ 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} →
+ 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 >>=⇿ ⟩
+ Any (λ y → x ∈ f₁ y) xs₁ ≈⟨ Any-cong (λ x → f₁≈f₂ x) xs₁≈xs₂ ⟩
+ Any (λ y → x ∈ f₂ y) xs₂ ⇿⟨ >>=⇿ ⟩
+ x ∈ (xs₂ >>= f₂) ∎
+ where open Inv.EquationalReasoning
+
+-- _⊛_ is a congruence.
+
+⊛-cong : ∀ {k A B} {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 →
+ >>=-cong xs₁≈xs₂ λ x →
+ _ ∎
+ where open Inv.EquationalReasoning
+
+-- _⊗_ is a congruence.
+
+⊗-cong : ∀ {k A B} {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₂
+
+-- The list monad's bind distributes from the left over _++_.
+
+>>=-left-distributive :
+ ∀ {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)) ⇿⟨ ++⇿ ⟩
+ y ∈ (xs >>= f) ++ (xs >>= g) ∎
+ where open Inv.EquationalReasoning
+
+-- The same applies to _⊛_.
+
+⊛-left-distributive :
+ ∀ {A B} (fs : List (A → B)) xs₁ xs₂ →
+ fs ⊛ (xs₁ ++ xs₂) ≈[ bag ] (fs ⊛ xs₁) ++ (fs ⊛ xs₂)
+⊛-left-distributive 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)) ⟩
+ (fs >>= λ f → (xs₁ >>= return ∘ f) ++
+ (xs₂ >>= return ∘ f)) ≈⟨ >>=-left-distributive fs ⟩
+
+ (fs >>= λ f → xs₁ >>= return ∘ f) ++
+ (fs >>= λ f → xs₂ >>= return ∘ f) ≡⟨ P.refl ⟩
+
+ (fs ⊛ xs₁) ++ (fs ⊛ xs₂) ∎
+ where open EqR ([ bag ]-Equality _)
diff --git a/src/Data/List/Any/Membership.agda b/src/Data/List/Any/Membership.agda
new file mode 100644
index 0000000..a7c5d19
--- /dev/null
+++ b/src/Data/List/Any/Membership.agda
@@ -0,0 +1,238 @@
+------------------------------------------------------------------------
+-- Properties related to list membership
+------------------------------------------------------------------------
+
+-- 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
+-- propositional equality.
+
+module Data.List.Any.Membership where
+
+open import Algebra
+open import Category.Monad
+open import Data.Bool
+open import Data.Empty
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence using (module Equivalent)
+import Function.Injection as Inj
+open import Function.Inverse as Inv using (_⇿_; 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.List.Any.Properties
+open import Data.Nat as Nat
+import Data.Nat.Properties as NatProp
+open import Data.Product as Prod
+open import Data.Sum as Sum
+open import Level
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; refl; _≗_)
+import Relation.Binary.Props.DecTotalOrder as DTOProps
+import Relation.Binary.Sigma.Pointwise as Σ
+open import Relation.Unary using (_⟨×⟩_)
+open import Relation.Nullary
+open import Relation.Nullary.Negation
+
+open Any.Membership-≡
+open RawMonad List.monad
+private
+ module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
+
+------------------------------------------------------------------------
+-- Properties relating _∈_ to various list functions
+
+map-∈⇿ : ∀ {A B : Set} {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⇿ ⟩
+ y ∈ List.map f xs ∎
+ where open Inv.EquationalReasoning
+
+concat-∈⇿ : ∀ {A : Set} {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⇿ ⟩
+ x ∈ concat xss ∎
+ where open Inv.EquationalReasoning
+
+>>=-∈⇿ : ∀ {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 ⇿⟨ >>=⇿ ⟩
+ y ∈ (xs >>= f) ∎
+ where open Inv.EquationalReasoning
+
+⊛-∈⇿ : ∀ {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⇿ ⟩
+ Any (λ f → Any (_≡_ y ∘ f) xs) fs ⇿⟨ ⊛⇿ ⟩
+ y ∈ fs ⊛ xs ∎
+ where open Inv.EquationalReasoning
+
+⊗-∈⇿ : ∀ {A B : Set} {xs ys} {x : A} {y : B} →
+ (x ∈ xs × y ∈ ys) ⇿ (x , y) ∈ (xs ⊗ ys)
+⊗-∈⇿ {A} {B} {xs} {ys} {x} {y} =
+ (x ∈ xs × y ∈ ys) ⇿⟨ ⊗⇿′ ⟩
+ Any (_≡_ x ⟨×⟩ _≡_ y) (xs ⊗ ys) ⇿⟨ Any-cong helper (_ ∎) ⟩
+ (x , y) ∈ (xs ⊗ ys) ∎
+ where
+ open Inv.EquationalReasoning
+
+ helper : (p : A × B) → (x ≡ proj₁ p × y ≡ proj₂ p) ⇿ (x , y) ≡ p
+ helper (x′ , y′) = record
+ { to = P.→-to-⟶ (uncurry $ P.cong₂ _,_)
+ ; from = P.→-to-⟶ < P.cong proj₁ , P.cong proj₂ >
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.cong₂ _,_ (P.proof-irrelevance _ _)
+ (P.proof-irrelevance _ _)
+ ; right-inverse-of = λ _ → P.proof-irrelevance _ _
+ }
+ }
+
+------------------------------------------------------------------------
+-- Various functions are monotone
+
+mono : ∀ {A : Set} {P : A → Set} {xs ys} →
+ xs ⊆ ys → Any P xs → Any P ys
+mono xs⊆ys =
+ _⟨$⟩_ (Inverse.to Any⇿) ∘
+ Prod.map id (Prod.map xs⊆ys id) ∘
+ _⟨$⟩_ (Inverse.from Any⇿)
+
+map-mono : ∀ {A B : Set} (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} →
+ 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)} →
+ xss ⊆ yss → concat xss ⊆ concat yss
+concat-mono xss⊆yss =
+ _⟨$⟩_ (Inverse.to concat-∈⇿) ∘
+ Prod.map id (Prod.map id xss⊆yss) ∘
+ _⟨$⟩_ (Inverse.from concat-∈⇿)
+
+>>=-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 >>=-∈⇿) ∘
+ Prod.map id (Prod.map xs⊆ys f⊆g) ∘
+ _⟨$⟩_ (Inverse.from >>=-∈⇿)
+
+_⊛-mono_ : ∀ {A B : Set} {fs gs : List (A → B)} {xs ys} →
+ fs ⊆ gs → xs ⊆ ys → (fs ⊛ xs) ⊆ (gs ⊛ ys)
+_⊛-mono_ {fs = fs} {gs} fs⊆gs xs⊆ys =
+ _⟨$⟩_ (Inverse.to $ ⊛-∈⇿ gs) ∘
+ Prod.map id (Prod.map id (Prod.map fs⊆gs (Prod.map xs⊆ys id))) ∘
+ _⟨$⟩_ (Inverse.from $ ⊛-∈⇿ fs)
+
+_⊗-mono_ : {A B : Set} {xs₁ ys₁ : List A} {xs₂ ys₂ : List B} →
+ xs₁ ⊆ ys₁ → xs₂ ⊆ ys₂ → (xs₁ ⊗ xs₂) ⊆ (ys₁ ⊗ ys₂)
+xs₁⊆ys₁ ⊗-mono xs₂⊆ys₂ =
+ _⟨$⟩_ (Inverse.to ⊗-∈⇿) ∘
+ Prod.map xs₁⊆ys₁ xs₂⊆ys₂ ∘
+ _⟨$⟩_ (Inverse.from ⊗-∈⇿)
+
+any-mono : ∀ {A : Set} (p : A → Bool) →
+ ∀ {xs ys} → xs ⊆ ys → T (any p xs) → T (any p ys)
+any-mono p xs⊆ys =
+ _⟨$⟩_ (Equivalent.to any⇔) ∘
+ mono xs⊆ys ∘
+ _⟨$⟩_ (Equivalent.from any⇔)
+
+map-with-∈-mono :
+ {A B : Set} {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} =
+ _⟨$⟩_ (Inverse.to map-with-∈⇿) ∘
+ Prod.map id (Prod.map xs⊆ys (λ {x∈xs} x≡fx∈xs → begin
+ x ≡⟨ x≡fx∈xs ⟩
+ f x∈xs ≡⟨ f≈g x∈xs ⟩
+ g (xs⊆ys x∈xs) ∎)) ∘
+ _⟨$⟩_ (Inverse.from map-with-∈⇿)
+ where open P.≡-Reasoning
+
+------------------------------------------------------------------------
+-- Other properties
+
+-- 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)) →
+ ∀ xs → ¬ (∀ i → Inj.Injection.to f ⟨$⟩ i ∈ xs)
+finite inj [] ∈[] with ∈[] zero
+... | ()
+finite {A} inj (x ∷ xs) ∈x∷xs = excluded-middle helper
+ where
+ open Inj.Injection inj
+
+ module DTO = DecTotalOrder Nat.decTotalOrder
+ module STO = StrictTotalOrder
+ (DTOProps.strictTotalOrder Nat.decTotalOrder)
+ module CS = CommutativeSemiring NatProp.commutativeSemiring
+
+ not-x : ∀ {i} → ¬ (to ⟨$⟩ i ≡ x) → to ⟨$⟩ i ∈ xs
+ not-x {i} ≢x with ∈x∷xs i
+ ... | here ≡x = ⊥-elim (≢x ≡x)
+ ... | there ∈xs = ∈xs
+
+ helper : ¬ Dec (∃ λ i → to ⟨$⟩ i ≡ x)
+ helper (no ≢x) = finite inj xs (λ i → not-x (≢x ∘ _,_ i))
+ helper (yes (i , ≡x)) = finite inj′ xs ∈xs
+ where
+ open P
+
+ f : ℕ → A
+ f j with STO.compare i j
+ f j | tri< _ _ _ = to ⟨$⟩ suc j
+ f j | tri≈ _ _ _ = to ⟨$⟩ suc j
+ f j | tri> _ _ _ = to ⟨$⟩ j
+
+ ∈-if-not-i : ∀ {j} → i ≢ j → to ⟨$⟩ j ∈ xs
+ ∈-if-not-i i≢j = not-x (i≢j ∘ injective ∘ trans ≡x ∘ sym)
+
+ lemma : ∀ {k j} → k ≤ j → suc j ≢ k
+ lemma 1+j≤j refl = NatProp.1+n≰n 1+j≤j
+
+ ∈xs : ∀ j → f j ∈ xs
+ ∈xs j with STO.compare i j
+ ∈xs j | tri< (i≤j , _) _ _ = ∈-if-not-i (lemma i≤j ∘ sym)
+ ∈xs j | tri> _ i≢j _ = ∈-if-not-i i≢j
+ ∈xs .i | tri≈ _ refl _ =
+ ∈-if-not-i (NatProp.m≢1+m+n i ∘
+ subst (_≡_ i ∘ suc) (sym $ proj₂ CS.+-identity i))
+
+ injective′ : Inj.Injective {B = P.setoid A} (→-to-⟶ f)
+ injective′ {j} {k} eq with STO.compare i j | STO.compare i k
+ ... | tri< _ _ _ | tri< _ _ _ = cong pred $ injective eq
+ ... | tri< _ _ _ | tri≈ _ _ _ = cong pred $ injective eq
+ ... | tri< (i≤j , _) _ _ | tri> _ _ (k≤i , _) = ⊥-elim (lemma (DTO.trans k≤i i≤j) $ injective eq)
+ ... | tri≈ _ _ _ | tri< _ _ _ = cong pred $ injective eq
+ ... | tri≈ _ _ _ | tri≈ _ _ _ = cong pred $ injective eq
+ ... | tri≈ _ i≡j _ | tri> _ _ (k≤i , _) = ⊥-elim (lemma (subst (_≤_ k) i≡j k≤i) $ injective eq)
+ ... | tri> _ _ (j≤i , _) | tri< (i≤k , _) _ _ = ⊥-elim (lemma (DTO.trans j≤i i≤k) $ sym $ injective eq)
+ ... | 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′ }
diff --git a/src/Data/List/Any/Properties.agda b/src/Data/List/Any/Properties.agda
index 080d606..f4b5d9b 100644
--- a/src/Data/List/Any/Properties.agda
+++ b/src/Data/List/Any/Properties.agda
@@ -1,7 +1,10 @@
------------------------------------------------------------------------
--- Properties relating Any to various list functions
+-- Properties related to Any
------------------------------------------------------------------------
+-- The other modules under Data.List.Any also contain properties
+-- related to Any.
+
module Data.List.Any.Properties where
open import Algebra
@@ -9,32 +12,32 @@ open import Category.Monad
open import Data.Bool
open import Data.Bool.Properties
open import Data.Empty
-open import Data.Function
-open import Data.Function.Equality as FunS
- using (_⟶_; _⟨$⟩_; _⇨_)
-import Data.Function.Injection as Inj
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence
+ using (_⇔_; equivalent; module Equivalent)
+open import Function.Inverse as Inv
+ using (Isomorphism; _⇿_; module Inverse)
+open import Function.Inverse.TypeIsomorphisms
open import Data.List as List
-open RawMonad List.monad
open import Data.List.Any as Any using (Any; here; there)
-open import Data.Nat as Nat
-import Data.Nat.Properties as NatProp
-open import Data.Product as Prod hiding (map)
+open import Data.Product as Prod hiding (swap)
open import Data.Sum as Sum using (_⊎_; inj₁; inj₂; [_,_]′)
-open import Level
-open import Relation.Unary using ( _⟨×⟩_; _⟨→⟩_) renaming (_⊆_ to _⋐_)
open import Relation.Binary
-import Relation.Binary.EqReasoning as EqReasoning
-import Relation.Binary.List.Pointwise as ListEq
-open import Relation.Binary.HeterogeneousEquality using (_≅_; refl)
-open import Relation.Binary.Product.Pointwise
+import Relation.Binary.HeterogeneousEquality as H
open import Relation.Binary.PropositionalEquality as P
- using (_≡_; _≢_; refl; inspect; _with-≡_)
-import Relation.Binary.Props.DecTotalOrder as DTOProps
-open import Relation.Nullary
-open import Relation.Nullary.Negation
+ using (_≡_; refl; inspect; _with-≡_)
+open import Relation.Unary using (_⟨×⟩_; _⟨→⟩_) renaming (_⊆_ to _⋐_)
+import Relation.Binary.Sigma.Pointwise as Σ
+
+open Any.Membership-≡
+open Inv.EquationalReasoning
+open RawMonad List.monad
+private
+ module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
------------------------------------------------------------------------
--- Lemmas related to Any
+-- Some lemmas related to map, find and lose
-- Any.map is functorial.
@@ -50,708 +53,539 @@ map-∘ : ∀ {A} {P Q R : A → Set} (f : Q ⋐ R) (g : P ⋐ Q)
map-∘ f g (here p) = refl
map-∘ f g (there p) = P.cong there $ map-∘ f g p
--- Introduction (⁺) and elimination (⁻) rules for map.
-
-map⁺ : ∀ {A B} {P : B → Set} {f : A → B} {xs} →
- Any (P ∘ f) xs → Any P (map f xs)
-map⁺ (here p) = here p
-map⁺ (there p) = there $ map⁺ p
-
-map⁻ : ∀ {A B} {P : B → Set} {f : A → B} {xs} →
- Any P (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
-
--- The rules are inverses.
-
-map⁺∘map⁻ : ∀ {A B : Set} {P : B → Set} {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} →
- (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)
-
--- Introduction and elimination rules for _++_.
-
-++⁺ˡ : ∀ {A} {P : A → Set} {xs ys} →
- Any P xs → Any P (xs ++ ys)
-++⁺ˡ (here p) = here p
-++⁺ˡ (there p) = there (++⁺ˡ p)
-
-++⁺ʳ : ∀ {A} {P : A → Set} xs {ys} →
- Any P ys → Any P (xs ++ ys)
-++⁺ʳ [] p = p
-++⁺ʳ (x ∷ xs) p = there (++⁺ʳ xs p)
-
-++⁻ : ∀ {A} {P : A → Set} 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)
-
--- The rules are inverses.
-
-++⁺∘++⁻ : ∀ {A} {P : A → Set} xs {ys}
- (p : Any P (xs ++ ys)) →
- [ ++⁺ˡ , ++⁺ʳ xs ]′ (++⁻ xs p) ≡ p
-++⁺∘++⁻ [] p = refl
-++⁺∘++⁻ (x ∷ xs) (here p) = refl
-++⁺∘++⁻ (x ∷ xs) (there p) with ++⁻ xs p | ++⁺∘++⁻ xs p
-++⁺∘++⁻ (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) →
- ++⁻ xs ([ ++⁺ˡ , ++⁺ʳ xs ]′ p) ≡ p
-++⁻∘++⁺ [] (inj₁ ())
-++⁻∘++⁺ [] (inj₂ p) = refl
-++⁻∘++⁺ (x ∷ xs) (inj₁ (here p)) = refl
-++⁻∘++⁺ (x ∷ xs) {ys} (inj₁ (there p)) rewrite ++⁻∘++⁺ xs {ys} (inj₁ p) = refl
-++⁻∘++⁺ (x ∷ xs) (inj₂ p) rewrite ++⁻∘++⁺ xs (inj₂ p) = refl
-
--- Introduction and elimination rules for return.
-
-return⁺ : ∀ {A} {P : A → Set} {x} →
- P x → Any P (return x)
-return⁺ = here
-
-return⁻ : ∀ {A} {P : A → Set} {x} →
- Any P (return x) → P x
-return⁻ (here p) = p
-return⁻ (there ())
-
--- The rules are inverses.
-
-return⁺∘return⁻ : ∀ {A} {P : A → Set} {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⁻ {P = P} (return⁺ p) ≡ p
-return⁻∘return⁺ p = refl
-
--- Introduction and elimination rules for concat.
-
-concat⁺ : ∀ {A} {P : A → Set} {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 →
- 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) (there p)
- with concat⁻ (xs ∷ xss) p
-... | here p′ = here (there p′)
-... | there p′ = there p′
-
--- The rules are inverses.
-
-concat⁻∘++⁺ˡ : ∀ {A} {P : A → Set} {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⁻ (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⁻ xss p) ≡ p
-concat⁺∘concat⁻ [] ()
-concat⁺∘concat⁻ ([] ∷ xss) p = concat⁺∘concat⁻ xss p
-concat⁺∘concat⁻ ((x ∷ xs) ∷ xss) (here p) = refl
-concat⁺∘concat⁻ ((x ∷ xs) ∷ xss) (there p)
- with concat⁻ (xs ∷ xss) p | concat⁺∘concat⁻ (xs ∷ xss) p
-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⁻ 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
-
--- Introduction and elimination rules for _>>=_.
-
->>=⁺ : ∀ {A B P xs} {f : A → List B} →
- Any (Any P ∘ f) xs → Any P (xs >>= f)
->>=⁺ = concat⁺ ∘ map⁺
-
->>=⁻ : ∀ {A B P} xs {f : A → List B} →
- Any P (xs >>= f) → Any (Any P ∘ f) xs
->>=⁻ _ = map⁻ ∘ concat⁻ _
-
--- The rules are inverses.
-
->>=⁺∘>>=⁻ : ∀ {A B P} xs {f : A → List B} (p : Any P (xs >>= f)) →
- >>=⁺ (>>=⁻ xs p) ≡ p
->>=⁺∘>>=⁻ xs {f} p rewrite map⁺∘map⁻ (concat⁻ (map f xs) p) =
- concat⁺∘concat⁻ (map f xs) p
-
->>=⁻∘>>=⁺ : ∀ {A B P xs} {f : A → List B} (p : Any (Any P ∘ f) xs) →
- >>=⁻ xs (>>=⁺ p) ≡ p
->>=⁻∘>>=⁺ {P = P} p rewrite concat⁻∘concat⁺ (map⁺ p) =
- map⁻∘map⁺ (Any P) p
-
--- Introduction and elimination rules for _⊛_.
-
-⊛⁺ : ∀ {A B P} {fs : List (A → B)} {xs} →
- Any (λ f → Any (P ∘ f) xs) fs → Any P (fs ⊛ xs)
-⊛⁺ = >>=⁺ ∘ Any.map (>>=⁺ ∘ Any.map return⁺)
-
-⊛⁺′ : ∀ {A B P Q} {fs : List (A → B)} {xs} →
- Any (P ⟨→⟩ Q) fs → Any P xs → Any Q (fs ⊛ xs)
-⊛⁺′ pq p = ⊛⁺ (Any.map (λ pq → Any.map (λ {x} → pq {x}) p) pq)
-
-⊛⁻ : ∀ {A B P} (fs : List (A → B)) xs →
- Any P (fs ⊛ xs) → Any (λ f → Any (P ∘ f) xs) fs
-⊛⁻ fs xs = Any.map (Any.map return⁻ ∘ >>=⁻ xs) ∘ >>=⁻ fs
-
--- The rules are inverses.
-
-⊛⁺∘⊛⁻ : ∀ {A B P} (fs : List (A → B))
- xs (p : Any P (fs ⊛ xs)) → ⊛⁺ (⊛⁻ fs xs p) ≡ p
-⊛⁺∘⊛⁻ {A} {B} {P} fs xs p = begin
- ⊛⁺ (⊛⁻ fs xs p) ≡⟨ P.cong >>=⁺ $ P.sym $
- map-∘ (>>=⁺ {P = P} ∘ Any.map return⁺)
- (Any.map return⁻ ∘ >>=⁻ {P = P} xs)
- (>>=⁻ fs p) ⟩
- >>=⁺ (Any.map (>>=⁺ {P = P} ∘ Any.map return⁺ ∘
- Any.map return⁻ ∘ >>=⁻ {P = P} xs)
- (>>=⁻ fs p)) ≡⟨ P.cong >>=⁺ (flip (map-id _) (>>=⁻ fs p) (λ p′ → begin
-
- >>=⁺ (Any.map return⁺ (Any.map (return⁻ {P = P}) (>>=⁻ xs p′))) ≡⟨ P.cong >>=⁺ $ P.sym $
- map-∘ return⁺ (return⁻ {P = P}) (>>=⁻ xs p′) ⟩
- >>=⁺ (Any.map (return⁺ ∘′ return⁻ {P = P}) (>>=⁻ xs p′)) ≡⟨ P.cong >>=⁺ (map-id _ return⁺∘return⁻ (>>=⁻ xs p′)) ⟩
- >>=⁺ (>>=⁻ xs p′) ≡⟨ >>=⁺∘>>=⁻ xs p′ ⟩
- p′ ∎)) ⟩
-
- >>=⁺ (>>=⁻ fs p) ≡⟨ >>=⁺∘>>=⁻ fs p ⟩
- p ∎
- where open P.≡-Reasoning
-
-⊛⁻∘⊛⁺ : ∀ {A B} {P : B → Set} {fs : List (A → B)} {xs}
- (p : Any (λ f → Any (P ∘ f) xs) fs) →
- ⊛⁻ {P = P} fs xs (⊛⁺ p) ≡ p
-⊛⁻∘⊛⁺ {P = P} {fs} {xs} p =
- helper₁ (>>=⁻∘>>=⁺ (Any.map (>>=⁺ {P = P} ∘ Any.map return⁺) p))
- where
- open P.≡-Reasoning
-
- helper₂ : ∀ {f} (p : Any (P ∘ f) xs) →
- Any.map return⁻ (>>=⁻ {P = P} xs
- (>>=⁺ (Any.map return⁺ p))) ≡ p
- helper₂ p rewrite >>=⁻∘>>=⁺ {P = P} (Any.map return⁺ p) = begin
- Any.map return⁻ (Any.map return⁺ p) ≡⟨ P.sym (map-∘ (return⁻ {P = P}) return⁺ p) ⟩
- Any.map id p ≡⟨ map-id id (λ _ → refl) p ⟩
- p ∎
-
- helper₁ : ∀ {p′} →
- p′ ≡ Any.map (>>=⁺ {P = P} ∘ Any.map return⁺) p →
- Any.map (Any.map return⁻ ∘ >>=⁻ xs) p′ ≡ p
- helper₁ refl
- rewrite P.sym (map-∘ (Any.map return⁻ ∘ >>=⁻ xs)
- (>>=⁺ {P = P} ∘ Any.map return⁺) p) =
- map-id _ helper₂ p
-
--- Introduction and elimination rules for _⊗_.
-
-⊗⁺ : ∀ {A B P} {xs : List A} {ys : List B} →
- Any (λ x → Any (λ y → P (x , y)) ys) xs → Any P (xs ⊗ ys)
-⊗⁺ = ⊛⁺ ∘′ ⊛⁺ ∘′ return⁺
-
-⊗⁺′ : ∀ {A B} {P : A → Set} {Q : B → Set} {xs ys} →
- Any P xs → Any Q ys → Any (P ⟨×⟩ Q) (xs ⊗ ys)
-⊗⁺′ p q = ⊗⁺ (Any.map (λ p → Any.map (λ q → (p , q)) q) p)
-
-⊗⁻ : ∀ {A B P} (xs : List A) (ys : List B) →
- Any P (xs ⊗ ys) → Any (λ x → Any (λ y → P (x , y)) ys) xs
-⊗⁻ _ _ = return⁻ ∘ ⊛⁻ _ _ ∘ ⊛⁻ _ _
-
-⊗⁻′ : ∀ {A B} {P : A → Set} {Q : B → Set} xs ys →
- Any (P ⟨×⟩ Q) (xs ⊗ ys) → Any P xs × Any Q ys
-⊗⁻′ _ _ pq =
- (Any.map (proj₁ ∘ proj₂ ∘ Any.satisfied) lem ,
- (Any.map proj₂ $ proj₂ (Any.satisfied lem)))
- where lem = ⊗⁻ _ _ pq
-
--- Any and any are related via T.
-
-any⁺ : ∀ {A} (p : A → Bool) {xs} →
- Any (T ∘ p) xs → T (any p xs)
-any⁺ p (here px) = proj₂ T-∨ (inj₁ px)
-any⁺ p (there {x = x} pxs) with p x
-... | true = _
-... | false = any⁺ p pxs
-
-any⁻ : ∀ {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)
-any⁻ p (x ∷ xs) px∷xs | true with-≡ eq = here (proj₂ T-≡ $ P.sym eq)
-any⁻ p (x ∷ xs) px∷xs | false with-≡ eq with p x
-any⁻ p (x ∷ xs) pxs | false with-≡ refl | .false =
- there (any⁻ p xs pxs)
+-- Lemmas relating map and find.
+
+map∘find : ∀ {A : Set} {P : A → Set} {xs}
+ (p : Any P xs) → let p′ = find p in
+ {f : _≡_ (proj₁ p′) ⋐ P} →
+ f refl ≡ proj₂ (proj₂ p′) →
+ Any.map f (proj₁ (proj₂ p′)) ≡ p
+map∘find (here p) hyp = 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 (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
+
+-- find satisfies a simple equality when the predicate is a
+-- propositional equality.
+
+find-∈ : ∀ {A : Set} {x : A} {xs} (x∈xs : x ∈ xs) →
+ find x∈xs ≡ (x , x∈xs , refl)
+find-∈ (here refl) = refl
+find-∈ (there x∈xs) rewrite find-∈ x∈xs = refl
+
+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 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 P x∈xs p
+ rewrite find∘map x∈xs (flip (P.subst P) p)
+ | find-∈ x∈xs
+ = refl
+
+-- Any can be expressed using _∈_.
+
+Any⇿ : ∀ {A : Set} {P : A → Set} {xs} →
+ (∃ λ x → x ∈ xs × P x) ⇿ Any P xs
+Any⇿ {P = P} = record
+ { to = P.→-to-⟶ (uncurry′ lose ∘ proj₂)
+ ; from = P.→-to-⟶ find
+ ; inverse-of = record
+ { left-inverse-of = λ p →
+ find∘lose P (proj₁ (proj₂ p)) (proj₂ (proj₂ p))
+ ; right-inverse-of = lose∘find
+ }
+ }
------------------------------------------------------------------------
--- Lemmas related to _∈_, parameterised on underlying equalities
+-- Any is a congruence
-module Membership₁ (S : Setoid zero zero) where
+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₂ ∎
- open Any.Membership S
- private
- open module S = Setoid S using (_≈_)
- open module [M] = Any.Membership (ListEq.setoid S)
- using () renaming (_∈_ to _[∈]_; _⊆_ to _[⊆]_)
- open module M≡ = Any.Membership-≡
- using () renaming (_∈_ to _∈≡_; _⊆_ to _⊆≡_)
+------------------------------------------------------------------------
+-- Swapping
+
+-- Nested occurrences of Any can sometimes be swapped. See also ×⇿.
+
+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} →
+ (x ∈ xs × y ∈ ys × P x y) ⇿⟨ sym $ ×⊎.*-assoc _ _ _ ⟩
+ ((x ∈ xs × y ∈ ys) × P x y) ⇿⟨ ×⊎.*-comm _ _ ⟨ ×⊎.*-cong ⟩ 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 → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ⇿⟨ Σ.cong (Inv.id ⟨ ×⊎.*-cong ⟩ Any⇿) ⟩
+ (∃ λ y → y ∈ ys × Any (flip P y) xs) ⇿⟨ Any⇿ ⟩
+ Any (λ y → Any (flip P y) xs) ys ∎
- -- Any is monotone.
+------------------------------------------------------------------------
+-- Lemmas relating Any to ⊥
+
+⊥⇿Any⊥ : {A : Set} {xs : List A} → ⊥ ⇿ Any (const ⊥) xs
+⊥⇿Any⊥ {A} = record
+ { to = P.→-to-⟶ (λ ())
+ ; from = P.→-to-⟶ (λ p → from p)
+ ; inverse-of = record
+ { left-inverse-of = λ ()
+ ; right-inverse-of = λ p → from p
+ }
+ }
+ where
+ from : {xs : List A} → Any (const ⊥) xs → {B : Set} → B
+ from (here ())
+ from (there p) = from p
+
+⊥⇿Any[] : {A : Set} {P : A → Set} → ⊥ ⇿ Any P []
+⊥⇿Any[] = record
+ { to = P.→-to-⟶ (λ ())
+ ; from = P.→-to-⟶ (λ ())
+ ; inverse-of = record
+ { left-inverse-of = λ ()
+ ; right-inverse-of = λ ()
+ }
+ }
- mono : ∀ {P xs ys} →
- P Respects _≈_ → xs ⊆ ys → Any P xs → Any P ys
- mono resp xs⊆ys pxs with find pxs
- ... | (x , x∈xs , px) = lose resp (xs⊆ys x∈xs) px
+------------------------------------------------------------------------
+-- Lemmas relating Any to sums and products
+
+-- Sums commute with Any (for a fixed list).
+
+⊎⇿ : ∀ {A : Set} {P Q : A → Set} {xs} →
+ (Any P xs ⊎ Any Q xs) ⇿ Any (λ x → P x ⊎ Q x) xs
+⊎⇿ {P = P} {Q} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = from∘to
+ ; right-inverse-of = to∘from
+ }
+ }
+ where
+ to : ∀ {xs} → Any P xs ⊎ Any Q xs → Any (λ x → P x ⊎ Q x) xs
+ to = [ Any.map inj₁ , Any.map inj₂ ]′
+
+ from : ∀ {xs} → Any (λ x → P x ⊎ Q x) xs → Any P xs ⊎ Any Q xs
+ from (here (inj₁ p)) = inj₁ (here p)
+ from (here (inj₂ q)) = inj₂ (here q)
+ from (there p) = Sum.map there there (from p)
+
+ from∘to : ∀ {xs} (p : Any P xs ⊎ Any Q xs) → from (to p) ≡ p
+ from∘to (inj₁ (here p)) = P.refl
+ from∘to (inj₁ (there p)) rewrite from∘to (inj₁ p) = P.refl
+ from∘to (inj₂ (here q)) = P.refl
+ from∘to (inj₂ (there q)) rewrite from∘to (inj₂ q) = P.refl
+
+ to∘from : ∀ {xs} (p : Any (λ x → P x ⊎ Q x) xs) →
+ to (from p) ≡ p
+ to∘from (here (inj₁ p)) = P.refl
+ to∘from (here (inj₂ q)) = P.refl
+ to∘from (there p) with from p | to∘from p
+ to∘from (there .(Any.map inj₁ p)) | inj₁ p | P.refl = P.refl
+ to∘from (there .(Any.map inj₂ q)) | inj₂ q | P.refl = P.refl
+
+-- Products "commute" with Any.
+
+×⇿ : ∀ {A B} {P : A → Set} {Q : B → Set} {xs} {ys} →
+ (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
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = from∘to
+ ; right-inverse-of = to∘from
+ }
+ }
+ where
+ to : Any P xs × Any Q ys → Any (λ x → Any (λ y → P x × Q y) ys) xs
+ 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
+ ... | (x , x∈xs , y , y∈ys , p , q) = (lose x∈xs p , lose y∈ys q)
+
+ from∘to : ∀ pq → from (to pq) ≡ pq
+ from∘to (p , q)
+ rewrite find∘map {Q = λ x → Any (λ y → P x × Q y) ys}
+ p (λ p → Any.map (λ q → (p , q)) q)
+ | find∘map {Q = λ y → P (proj₁ (find p)) × Q y}
+ q (λ q → proj₂ (proj₂ (find p)) , q)
+ | lose∘find p
+ | lose∘find q
+ = refl
+
+ to∘from : ∀ pq → to (from pq) ≡ pq
+ to∘from pq
+ with find pq
+ | (λ (f : _≡_ (proj₁ (find pq)) ⋐ _) → map∘find pq {f})
+ ... | (x , x∈xs , pq′) | lem₁
+ with find pq′
+ | (λ (f : _≡_ (proj₁ (find pq′)) ⋐ _) → map∘find pq′ {f})
+ ... | (y , y∈ys , p , q) | lem₂
+ rewrite P.sym $ map-∘ {R = λ x → Any (λ y → P x × Q y) ys}
+ (λ p → Any.map (λ q → p , q) (lose y∈ys q))
+ (λ y → P.subst P y p)
+ x∈xs
+ = lem₁ _ helper
+ where
+ helper : Any.map (λ q → p , q) (lose y∈ys q) ≡ pq′
+ helper rewrite P.sym $ map-∘ {R = λ y → P x × Q y}
+ (λ q → p , q)
+ (λ y → P.subst Q y q)
+ y∈ys
+ = lem₂ _ refl
- -- _++_ is monotone.
+------------------------------------------------------------------------
+-- Invertible introduction (⁺) and elimination (⁻) rules for various
+-- list functions
+
+-- map.
+
+private
+
+ map⁺ : ∀ {A B} {P : B → Set} {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} →
+ 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} →
+ (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} →
+ (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} →
+ Any (P ∘ f) xs ⇿ Any P (List.map f xs)
+map⇿ {P = P} = record
+ { to = P.→-to-⟶ $ map⁺ {P = P}
+ ; from = P.→-to-⟶ map⁻
+ ; inverse-of = record
+ { left-inverse-of = map⁻∘map⁺ P
+ ; right-inverse-of = map⁺∘map⁻
+ }
+ }
+
+-- _++_.
+
+private
+
+ ++⁺ˡ : ∀ {A} {P : A → Set} {xs ys} →
+ Any P xs → Any P (xs ++ ys)
+ ++⁺ˡ (here p) = here p
+ ++⁺ˡ (there p) = there (++⁺ˡ p)
+
+ ++⁺ʳ : ∀ {A} {P : A → Set} xs {ys} →
+ Any P ys → Any P (xs ++ ys)
+ ++⁺ʳ [] p = p
+ ++⁺ʳ (x ∷ xs) p = there (++⁺ʳ xs p)
+
+ ++⁻ : ∀ {A} {P : A → Set} 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}
+ (p : Any P (xs ++ ys)) →
+ [ ++⁺ˡ , ++⁺ʳ xs ]′ (++⁻ xs p) ≡ p
+ ++⁺∘++⁻ [] p = refl
+ ++⁺∘++⁻ (x ∷ xs) (here p) = refl
+ ++⁺∘++⁻ (x ∷ xs) (there p) with ++⁻ xs p | ++⁺∘++⁻ xs p
+ ++⁺∘++⁻ (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) →
+ ++⁻ xs ([ ++⁺ˡ , ++⁺ʳ xs ]′ p) ≡ p
+ ++⁻∘++⁺ [] (inj₁ ())
+ ++⁻∘++⁺ [] (inj₂ p) = refl
+ ++⁻∘++⁺ (x ∷ xs) (inj₁ (here p)) = refl
+ ++⁻∘++⁺ (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} →
+ (Any P xs ⊎ Any P ys) ⇿ Any P (xs ++ ys)
+++⇿ {xs = xs} = record
+ { to = P.→-to-⟶ [ ++⁺ˡ , ++⁺ʳ xs ]′
+ ; from = P.→-to-⟶ $ ++⁻ xs
+ ; inverse-of = record
+ { left-inverse-of = ++⁻∘++⁺ xs
+ ; right-inverse-of = ++⁺∘++⁻ xs
+ }
+ }
+
+-- return.
+
+private
+
+ return⁺ : ∀ {A} {P : A → Set} {x} →
+ P x → Any P (return x)
+ return⁺ = here
+
+ return⁻ : ∀ {A} {P : A → Set} {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⁻ p) ≡ p
+ return⁺∘return⁻ (here p) = refl
+ return⁺∘return⁻ (there ())
+
+ return⁻∘return⁺ : ∀ {A} (P : A → Set) {x} (p : P x) →
+ return⁻ {P = P} (return⁺ p) ≡ p
+ return⁻∘return⁺ P p = refl
+
+return⇿ : ∀ {A} {P : A → Set} {x} →
+ P x ⇿ Any P (return x)
+return⇿ {P = P} = record
+ { to = P.→-to-⟶ return⁺
+ ; from = P.→-to-⟶ return⁻
+ ; inverse-of = record
+ { left-inverse-of = return⁻∘return⁺ P
+ ; right-inverse-of = return⁺∘return⁻
+ }
+ }
+
+-- concat.
+
+private
+
+ concat⁺ : ∀ {A} {P : A → Set} {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 →
+ 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) (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⁻ (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⁻ (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⁻ xss p) ≡ p
+ concat⁺∘concat⁻ [] ()
+ concat⁺∘concat⁻ ([] ∷ xss) p = concat⁺∘concat⁻ xss p
+ concat⁺∘concat⁻ ((x ∷ xs) ∷ xss) (here p) = refl
+ concat⁺∘concat⁻ ((x ∷ xs) ∷ xss) (there p)
+ with concat⁻ (xs ∷ xss) p | concat⁺∘concat⁻ (xs ∷ xss) p
+ 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⁻ 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} →
+ Any (Any P) xss ⇿ Any P (concat xss)
+concat⇿ {xss = xss} = record
+ { to = P.→-to-⟶ concat⁺
+ ; from = P.→-to-⟶ $ concat⁻ xss
+ ; inverse-of = record
+ { left-inverse-of = concat⁻∘concat⁺
+ ; right-inverse-of = concat⁺∘concat⁻ xss
+ }
+ }
+
+-- _>>=_.
+
+>>=⇿ : ∀ {A B 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 P (xs >>= f) ∎
+
+-- _⊛_.
+
+⊛⇿ : ∀ {A B P} {fs : List (A → B)} {xs} →
+ 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 ⇿⟨ >>=⇿ ⟩
+ Any P (fs ⊛ xs) ∎
+
+-- An alternative introduction rule for _⊛_.
- _++-mono_ : ∀ {xs₁ xs₂ ys₁ ys₂} →
- xs₁ ⊆ ys₁ → xs₂ ⊆ ys₂ → xs₁ ++ xs₂ ⊆ ys₁ ++ ys₂
- _++-mono_ {ys₁ = ys₁} xs₁⊆ys₁ xs₂⊆ys₂ =
- [ ++⁺ˡ ∘ xs₁⊆ys₁ , ++⁺ʳ ys₁ ∘ xs₂⊆ys₂ ]′ ∘ ++⁻ _
+⊛⁺′ : ∀ {A B P Q} {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
+
+-- _⊗_.
+
+⊗⇿ : ∀ {A B P} {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⇿ ⟩
+ 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} →
+ (Any P xs × Any Q ys) ⇿ Any (P ⟨×⟩ Q) (xs ⊗ ys)
+⊗⇿′ {P = P} {Q} {xs} {ys} =
+ (Any P xs × Any Q ys) ⇿⟨ ×⇿ ⟩
+ Any (λ x → Any (λ y → P x × Q y) ys) xs ⇿⟨ ⊗⇿ ⟩
+ Any (P ⟨×⟩ Q) (xs ⊗ ys) ∎
+
+-- map-with-∈.
+
+map-with-∈⇿ :
+ ∀ {A B : Set} {P : B → Set} {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
+ { to = P.→-to-⟶ (map-with-∈⁺ _)
+ ; from = P.→-to-⟶ (map-with-∈⁻ _ _)
+ ; inverse-of = record
+ { left-inverse-of = from∘to _
+ ; right-inverse-of = to∘from _ _
+ }
+ }
+ where
+ map-with-∈⁺ : ∀ {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-∈⁺ f (_ , here refl , p) = here p
+ map-with-∈⁺ f (_ , there x∈xs , p) =
+ there $ map-with-∈⁺ (f ∘ there) (_ , x∈xs , p)
+
+ map-with-∈⁻ : ∀ (xs : List A)
+ (f : ∀ {x} → x ∈ xs → B) →
+ Any P (map-with-∈ xs f) →
+ ∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)
+ map-with-∈⁻ [] f ()
+ map-with-∈⁻ (y ∷ xs) f (here p) = (y , here refl , p)
+ map-with-∈⁻ (y ∷ xs) f (there p) =
+ Prod.map id (Prod.map there id) $ map-with-∈⁻ xs (f ∘ there) p
+
+ from∘to : ∀ {xs : List A} (f : ∀ {x} → x ∈ xs → B)
+ (p : ∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)) →
+ map-with-∈⁻ xs f (map-with-∈⁺ f p) ≡ p
+ from∘to f (_ , here refl , p) = refl
+ from∘to f (_ , there x∈xs , p)
+ rewrite from∘to (f ∘ there) (_ , x∈xs , p) = refl
+
+ to∘from : ∀ (xs : List A) (f : ∀ {x} → x ∈ xs → B)
+ (p : Any P (map-with-∈ xs f)) →
+ map-with-∈⁺ f (map-with-∈⁻ xs f p) ≡ p
+ to∘from [] f ()
+ to∘from (y ∷ xs) f (here p) = refl
+ to∘from (y ∷ xs) f (there p) =
+ P.cong there $ to∘from xs (f ∘ there) p
- -- _++_ is idempotent.
+------------------------------------------------------------------------
+-- Any and any are related via T
- ++-idempotent : ∀ {xs} → xs ++ xs ⊆ xs
- ++-idempotent = [ id , id ]′ ∘ ++⁻ _
+-- These introduction and elimination rules are not inverses, though.
- -- _++_ is commutative.
+private
- ++-comm : ∀ xs ys → xs ++ ys ⊆ ys ++ xs
- ++-comm xs ys = [ ++⁺ʳ ys , ++⁺ˡ ]′ ∘ ++⁻ xs
+ any⁺ : ∀ {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
- -- Introduction and elimination rules for concat.
-
- concat-∈⁺ : ∀ {x xs xss} →
- x ∈ xs → xs [∈] xss → x ∈ concat xss
- concat-∈⁺ x∈xs xs∈xss =
- concat⁺ (Any.map (λ xs≈ys → Pre.reflexive xs≈ys x∈xs) xs∈xss)
- where module Pre = Preorder ⊆-preorder
-
- concat-∈⁻ : ∀ {x} xss → x ∈ concat xss →
- ∃ λ xs → x ∈ xs × xs [∈] xss
- concat-∈⁻ xss x∈ = Prod.map id swap $ [M].find (concat⁻ xss x∈)
-
- -- concat is monotone.
-
- concat-mono : ∀ {xss yss} →
- xss [⊆] yss → concat xss ⊆ concat yss
- concat-mono {xss = xss} xss⊆yss x∈ with concat-∈⁻ xss x∈
- ... | (xs , x∈xs , xs∈xss) = concat-∈⁺ x∈xs (xss⊆yss xs∈xss)
-
- -- any is monotone.
-
- any-mono : ∀ p → (T ∘ p) Respects _≈_ →
- ∀ {xs ys} → xs ⊆ ys → T (any p xs) → T (any p ys)
- any-mono p resp xs⊆ys = any⁺ p ∘ mono resp xs⊆ys ∘ any⁻ p _
-
- -- Introduction and elimination rules for map-with-∈.
-
- map-with-∈-∈⁺ : ∀ {A} {xs : List A}
- (f : ∀ {x} → x ∈≡ xs → S.Carrier) {x} →
- (x∈xs : x ∈≡ xs) → f x∈xs ∈ M≡.map-with-∈ xs f
- map-with-∈-∈⁺ f (here refl) = here S.refl
- map-with-∈-∈⁺ f (there x∈xs) = there $ map-with-∈-∈⁺ (f ∘ there) x∈xs
-
- map-with-∈-∈⁻ : ∀ {A} {xs : List A}
- (f : ∀ {x} → x ∈≡ xs → S.Carrier) {fx∈xs} →
- fx∈xs ∈ M≡.map-with-∈ xs f →
- ∃ λ x → Σ (x ∈≡ xs) λ x∈xs → fx∈xs ≈ f x∈xs
- map-with-∈-∈⁻ {xs = []} f ()
- map-with-∈-∈⁻ {xs = y ∷ xs} f (here fx≈) = (y , here refl , fx≈)
- map-with-∈-∈⁻ {xs = y ∷ xs} f (there x∈xs) =
- Prod.map id (Prod.map there id) $ map-with-∈-∈⁻ (f ∘ there) x∈xs
-
- -- map-with-∈ is monotone.
-
- map-with-∈-mono :
- ∀ {A} {xs : List A} {f : ∀ {x} → x ∈≡ xs → S.Carrier}
- {ys : List A} {g : ∀ {x} → x ∈≡ ys → S.Carrier} →
- (xs⊆ys : xs ⊆≡ ys) →
- (∀ {x} (x∈xs : x ∈≡ xs) → f x∈xs ≈ g (xs⊆ys x∈xs)) →
- M≡.map-with-∈ xs f ⊆ M≡.map-with-∈ ys g
- map-with-∈-mono {f = f} {g = g} xs⊆ys f≈g {fx∈xs} fx∈xs∈
- with map-with-∈-∈⁻ f fx∈xs∈
- ... | (x , x∈xs , fx∈xs≈) =
- Any.map (λ {y} g[xs⊆ys-x∈xs]≈y → begin
- fx∈xs ≈⟨ fx∈xs≈ ⟩
- f x∈xs ≈⟨ f≈g x∈xs ⟩
- g (xs⊆ys x∈xs) ≈⟨ g[xs⊆ys-x∈xs]≈y ⟩
- y ∎) $
- map-with-∈-∈⁺ g (xs⊆ys x∈xs)
- where open EqReasoning S
-
- -- Only a finite number of distinct elements can be members of a
- -- given list.
-
- finite : (f : Inj.Injection (P.setoid ℕ) S) →
- ∀ xs → ¬ (∀ i → Inj.Injection.to f ⟨$⟩ i ∈ xs)
- finite inj [] ∈[] with ∈[] zero
- ... | ()
- finite inj (x ∷ xs) ∈x∷xs = excluded-middle helper
- where
- open Inj.Injection inj
-
- module DTO = DecTotalOrder Nat.decTotalOrder
- module STO = StrictTotalOrder
- (DTOProps.strictTotalOrder Nat.decTotalOrder)
- module CS = CommutativeSemiring NatProp.commutativeSemiring
-
- not-x : ∀ {i} → ¬ (to ⟨$⟩ i ≈ x) → to ⟨$⟩ i ∈ xs
- not-x {i} ≉x with ∈x∷xs i
- ... | here ≈x = ⊥-elim (≉x ≈x)
- ... | there ∈xs = ∈xs
-
- helper : ¬ Dec (∃ λ i → to ⟨$⟩ i ≈ x)
- helper (no ≉x) = finite inj xs (λ i → not-x (≉x ∘ _,_ i))
- helper (yes (i , ≈x)) = finite inj′ xs ∈xs
- where
- open P
-
- f : ℕ → S.Carrier
- f j with STO.compare i j
- f j | tri< _ _ _ = to ⟨$⟩ suc j
- f j | tri≈ _ _ _ = to ⟨$⟩ suc j
- f j | tri> _ _ _ = to ⟨$⟩ j
-
- ∈-if-not-i : ∀ {j} → i ≢ j → to ⟨$⟩ j ∈ xs
- ∈-if-not-i i≢j = not-x (i≢j ∘ injective ∘ S.trans ≈x ∘ S.sym)
-
- lemma : ∀ {k j} → k ≤ j → suc j ≢ k
- lemma 1+j≤j refl = NatProp.1+n≰n 1+j≤j
-
- ∈xs : ∀ j → f j ∈ xs
- ∈xs j with STO.compare i j
- ∈xs j | tri< (i≤j , _) _ _ = ∈-if-not-i (lemma i≤j ∘ sym)
- ∈xs j | tri> _ i≢j _ = ∈-if-not-i i≢j
- ∈xs .i | tri≈ _ refl _ =
- ∈-if-not-i (NatProp.m≢1+m+n i ∘
- subst (_≡_ i ∘ suc) (sym $ proj₂ CS.+-identity i))
-
- injective′ : Inj.Injective {B = S} (→-to-⟶ f)
- injective′ {j} {k} eq with STO.compare i j | STO.compare i k
- ... | tri< _ _ _ | tri< _ _ _ = cong pred $ injective eq
- ... | tri< _ _ _ | tri≈ _ _ _ = cong pred $ injective eq
- ... | tri< (i≤j , _) _ _ | tri> _ _ (k≤i , _) = ⊥-elim (lemma (DTO.trans k≤i i≤j) $ injective eq)
- ... | tri≈ _ _ _ | tri< _ _ _ = cong pred $ injective eq
- ... | tri≈ _ _ _ | tri≈ _ _ _ = cong pred $ injective eq
- ... | tri≈ _ i≡j _ | tri> _ _ (k≤i , _) = ⊥-elim (lemma (subst (_≤_ k) i≡j k≤i) $ injective eq)
- ... | tri> _ _ (j≤i , _) | tri< (i≤k , _) _ _ = ⊥-elim (lemma (DTO.trans j≤i i≤k) $ sym $ injective eq)
- ... | 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′ }
-
-module Membership₂ (S₁ S₂ : Setoid zero zero) where
-
- private
- open module S₁ = Setoid S₁ using () renaming (_≈_ to _≈₁_)
- open module S₂ = Setoid S₂ using () renaming (_≈_ to _≈₂_)
- LS₂ = ListEq.setoid S₂
- open module M₁ = Any.Membership S₁
- using () renaming (_∈_ to _∈₁_; _⊆_ to _⊆₁_)
- open module M₂ = Any.Membership S₂
- using () renaming (_∈_ to _∈₂_; _⊆_ to _⊆₂_)
- open module M₁₂ = Any.Membership (S₁ ⇨ S₂)
- using () renaming (_∈_ to _∈₁₂_; _⊆_ to _⊆₁₂_)
- open Any.Membership (S₁ ×-setoid S₂)
- using () renaming (_⊆_ to _⊆₁,₂_)
-
- -- Introduction and elimination rules for map.
-
- map-∈⁺ : ∀ (f : S₁ ⟶ S₂) {x xs} →
- x ∈₁ xs → f ⟨$⟩ x ∈₂ map (_⟨$⟩_ f) xs
- map-∈⁺ f = map⁺ ∘ Any.map (FunS.cong f)
-
- map-∈⁻ : ∀ {f fx} xs →
- fx ∈₂ map f xs → ∃ λ x → x ∈₁ xs × fx ≈₂ f x
- map-∈⁻ _ fx∈ = M₁.find (map⁻ fx∈)
-
- -- map is monotone.
-
- map-mono : ∀ (f : S₁ ⟶ S₂) {xs ys} →
- xs ⊆₁ ys → map (_⟨$⟩_ f) xs ⊆₂ map (_⟨$⟩_ f) ys
- map-mono f xs⊆ys fx∈ with map-∈⁻ _ fx∈
- ... | (x , x∈ , eq) = Any.map (S₂.trans eq) (map-∈⁺ f (xs⊆ys x∈))
-
- -- Introduction and elimination rules for _>>=_.
-
- >>=-∈⁺ : ∀ (f : S₁ ⟶ LS₂) {x y xs} →
- x ∈₁ xs → y ∈₂ f ⟨$⟩ x → y ∈₂ (xs >>= _⟨$⟩_ f)
- >>=-∈⁺ f x∈xs y∈fx =
- >>=⁺ (Any.map (flip M₂.∈-resp-list-≈ y∈fx ∘ FunS.cong f) x∈xs)
-
- >>=-∈⁻ : ∀ (f : S₁ ⟶ LS₂) {y} xs →
- y ∈₂ (xs >>= _⟨$⟩_ f) → ∃ λ x → x ∈₁ xs × y ∈₂ f ⟨$⟩ x
- >>=-∈⁻ f xs y∈ = M₁.find (>>=⁻ xs y∈)
-
- -- _>>=_ is monotone.
-
- >>=-mono : ∀ (f g : S₁ ⟶ LS₂) {xs ys} →
- xs ⊆₁ ys → (∀ {x} → f ⟨$⟩ x ⊆₂ g ⟨$⟩ x) →
- (xs >>= _⟨$⟩_ f) ⊆₂ (ys >>= _⟨$⟩_ g)
- >>=-mono f g {xs} xs⊆ys f⊆g z∈ with >>=-∈⁻ f xs z∈
- ... | (x , x∈xs , z∈fx) = >>=-∈⁺ g (xs⊆ys x∈xs) (f⊆g z∈fx)
-
- -- Introduction and elimination rules for _⊛_.
-
- private
-
- infixl 4 _⟨⊛⟩_
-
- _⟨⊛⟩_ : List (S₁ ⟶ S₂) → List S₁.Carrier → List S₂.Carrier
- fs ⟨⊛⟩ xs = map _⟨$⟩_ fs ⊛ xs
-
- ⊛-∈⁺ : ∀ f {fs x xs} →
- f ∈₁₂ fs → x ∈₁ xs → f ⟨$⟩ x ∈₂ fs ⟨⊛⟩ xs
- ⊛-∈⁺ _ f∈fs x∈xs =
- ⊛⁺′ (map⁺ (Any.map (λ f≈g x≈y → f≈g x≈y) f∈fs)) x∈xs
-
- ⊛-∈⁻ : ∀ fs xs {fx} → fx ∈₂ fs ⟨⊛⟩ xs →
- ∃₂ λ f x → f ∈₁₂ fs × x ∈₁ xs × fx ≈₂ f ⟨$⟩ x
- ⊛-∈⁻ fs xs fx∈ with M₁₂.find $ map⁻ (⊛⁻ (map _⟨$⟩_ fs) xs fx∈)
- ... | (f , f∈fs , x∈) with M₁.find x∈
- ... | (x , x∈xs , fx≈fx) = (f , x , f∈fs , x∈xs , fx≈fx)
-
- -- _⊛_ is monotone.
-
- _⊛-mono_ : ∀ {fs gs xs ys} →
- fs ⊆₁₂ gs → xs ⊆₁ ys → fs ⟨⊛⟩ xs ⊆₂ gs ⟨⊛⟩ ys
- _⊛-mono_ {fs = fs} {xs = xs} fs⊆gs xs⊆ys fx∈ with ⊛-∈⁻ fs xs fx∈
- ... | (f , x , f∈fs , x∈xs , fx≈fx) =
- Any.map (S₂.trans fx≈fx) $ ⊛-∈⁺ f (fs⊆gs {f} f∈fs) (xs⊆ys x∈xs)
-
- -- _⊗_ is monotone.
-
- _⊗-mono_ : ∀ {xs₁ xs₂ ys₁ ys₂} →
- xs₁ ⊆₁ ys₁ → xs₂ ⊆₂ ys₂ → (xs₁ ⊗ xs₂) ⊆₁,₂ (ys₁ ⊗ ys₂)
- _⊗-mono_ {xs₁ = xs₁} {xs₂} xs₁⊆ys₁ xs₂⊆ys₂ {x , y} x,y∈
- with ⊗⁻′ {P = _≈₁_ x} {Q = _≈₂_ y} xs₁ xs₂ x,y∈
- ... | (x∈ , y∈) = ⊗⁺′ (xs₁⊆ys₁ x∈) (xs₂⊆ys₂ y∈)
+ any⁻ : ∀ {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)
+ any⁻ p (x ∷ xs) px∷xs | true with-≡ eq = here (Equivalent.from T-≡ ⟨$⟩ eq)
+ any⁻ p (x ∷ xs) px∷xs | false with-≡ eq with p x
+ any⁻ p (x ∷ xs) pxs | false with-≡ refl | .false =
+ there (any⁻ p xs pxs)
+
+any⇔ : ∀ {A} {p : A → Bool} {xs} →
+ Any (T ∘ p) xs ⇔ T (any p xs)
+any⇔ = equivalent (any⁺ _) (any⁻ _ _)
------------------------------------------------------------------------
--- Lemmas related to the variant of _∈_ which is defined using
--- propositional equality
-
-module Membership-≡ where
-
- open Any.Membership-≡
- private
- module S {A} = Setoid (ListEq.setoid (P.setoid A))
- module M {A} = Any.Membership (P.setoid A)
- open module M₁ {A} = Membership₁ (P.setoid A) public
- using (_++-mono_; ++-idempotent; ++-comm;
- map-with-∈-∈⁺; map-with-∈-∈⁻; map-with-∈-mono;
- finite)
- open module M₂ {A} {B} =
- Membership₂ (P.setoid A) (P.setoid B) public
- using (map-∈⁻)
-
- -- Any is monotone.
-
- mono : ∀ {A xs ys} {P : A → Set} → xs ⊆ ys → Any P xs → Any P ys
- mono {P = P} = M₁.mono (P.subst P)
-
- -- Any.map is related to find.
-
- map∘find : ∀ {A : Set} {P : A → Set} {xs}
- (p : Any P xs) → let p′ = find p in
- {f : _≡_ (proj₁ p′) ⋐ P} →
- f refl ≡ proj₂ (proj₂ p′) →
- Any.map f (proj₁ (proj₂ p′)) ≡ p
- map∘find (here p) hyp = P.cong here hyp
- map∘find (there p) hyp = P.cong there (map∘find p hyp)
-
- find∘map : ∀ {A : Set} {P : A → Set} {x : A} {xs}
- (x∈xs : x ∈ xs) (f : _≡_ x ⋐ P) →
- find (Any.map f x∈xs) ≡ (x , x∈xs , f refl)
- find∘map (here refl) f = refl
- find∘map (there x∈xs) f rewrite find∘map x∈xs f = refl
-
- -- Introduction and elimination rules for concat.
-
- concat-∈⁺ : ∀ {A} {x : A} {xs xss} →
- x ∈ xs → xs ∈ xss → x ∈ concat xss
- concat-∈⁺ x∈xs = M₁.concat-∈⁺ x∈xs ∘ Any.map S.reflexive
-
- concat-∈⁻ : ∀ {A} {x : A} xss →
- x ∈ concat xss → ∃ λ xs → x ∈ xs × xs ∈ xss
- concat-∈⁻ xss x∈ =
- Prod.map id (Prod.map id (Any.map ListEq.Rel≡⇒≡)) $
- M₁.concat-∈⁻ xss x∈
-
- -- concat is monotone.
-
- concat-mono : ∀ {A} {xss yss : List (List A)} →
- xss ⊆ yss → concat xss ⊆ concat yss
- concat-mono xss⊆yss =
- M₁.concat-mono (Any.map S.reflexive ∘ xss⊆yss ∘
- Any.map ListEq.Rel≡⇒≡)
-
- -- any is monotone.
-
- any-mono : ∀ {A} (p : A → Bool) {xs ys} →
- xs ⊆ ys → T (any p xs) → T (any p ys)
- any-mono p = M₁.any-mono p (P.subst (T ∘ p))
-
- -- Introduction rule for map.
-
- map-∈⁺ : ∀ {A B} {f : A → B} {x xs} →
- x ∈ xs → f x ∈ map f xs
- map-∈⁺ {f = f} = M₂.map-∈⁺ (P.→-to-⟶ f)
-
- -- The introduction and elimination rules are inverses (more or
- -- less).
-
- map-∈⁺∘map-∈⁻ : ∀ {A B : Set} {f : A → B} {fx xs}
- (p : fx ∈ List.map f xs) →
- map-∈⁺ {f = f} (proj₁ (proj₂ (map-∈⁻ xs p))) ≅ p
- map-∈⁺∘map-∈⁻ {xs = []} ()
- map-∈⁺∘map-∈⁻ {xs = x ∷ xs} (here refl) = refl
- map-∈⁺∘map-∈⁻ {xs = x ∷ xs} (there p) = there-cong (map-∈⁺∘map-∈⁻ p)
- where
- there-cong : ∀ {A : Set} {P Q : A → Set} {x xs}
- {p : Any P xs} {q : Any Q xs} →
- p ≅ q → there {x = x} p ≅ there {x = x} q
- there-cong refl = refl
-
- map-∈⁻∘map-∈⁺ : ∀ {A B : Set} (f : A → B) {x xs} (x∈xs : x ∈ xs) →
- map-∈⁻ xs (map-∈⁺ {f = f} x∈xs) ≡ (x , x∈xs , refl)
- map-∈⁻∘map-∈⁺ {B = B} f {x} x∈xs = begin
- find (map⁻ {P = _≡_ _} $
- map⁺ (Any.map f-cong x∈xs)) ≡⟨ P.cong find $
- map⁻∘map⁺ (_≡_ _) (Any.map f-cong x∈xs) ⟩
- find (Any.map f-cong x∈xs) ≡⟨ find∘map x∈xs f-cong ⟩
- (x , x∈xs , refl) ∎
- where
- open P.≡-Reasoning
-
- f-cong : _≡_ =[ f ]⇒ _≡_
- f-cong = FunS.cong (P.→-to-⟶ {B = P.setoid B} f)
+-- _++_ is commutative
+
+private
- -- map is monotone.
+ ++-comm : ∀ {A} {P : A → Set} xs ys →
+ Any P (xs ++ ys) → Any P (ys ++ xs)
+ ++-comm xs ys = [ ++⁺ʳ ys , ++⁺ˡ ]′ ∘ ++⁻ xs
- map-mono : ∀ {A B} {f : A → B} {xs ys} →
- xs ⊆ ys → map f xs ⊆ map f ys
- map-mono {f = f} = M₂.map-mono (P.→-to-⟶ f)
-
- -- Introduction and elimination rules for _>>=_.
-
- >>=-∈⁺ : ∀ {A B} (f : A → List B) {x y xs} →
- x ∈ xs → y ∈ f x → y ∈ (xs >>= f)
- >>=-∈⁺ f = M₂.>>=-∈⁺ (P.→-to-⟶ f)
-
- >>=-∈⁻ : ∀ {A B} (f : A → List B) {y} xs →
- y ∈ (xs >>= f) → ∃ λ x → x ∈ xs × y ∈ f x
- >>=-∈⁻ f = M₂.>>=-∈⁻ (P.→-to-⟶ f)
-
- -- The introduction and elimination rules are inverses (more or
- -- less).
-
- private
-
- lift-resp-id : ∀ {A : Set} {x : A} {xs} xs≈ys (p : x ∈ xs) →
- M.∈-resp-list-≈ xs≈ys p ≡ p
- lift-resp-id ListEq.[] ()
- lift-resp-id (ListEq._∷_ refl xs≈ys) (here refl) = refl
- lift-resp-id (ListEq._∷_ refl xs≈ys) (there p) =
- P.cong there (lift-resp-id xs≈ys p)
-
- >>=-∈⁺∘>>=-∈⁻ : ∀ {A B : Set} (f : A → List B) {y} xs
- (p : y ∈ (xs >>= f)) →
- let p′ = proj₂ $ >>=-∈⁻ f xs p in
- >>=-∈⁺ f (proj₁ p′) (proj₂ p′) ≡ p
- >>=-∈⁺∘>>=-∈⁻ f xs p = begin
- >>=⁺ (Any.map _ (proj₁ (proj₂ (find (>>=⁻ xs p))))) ≡⟨ P.cong >>=⁺ $
- map∘find (>>=⁻ xs p) (lift-resp-id _ _) ⟩
- >>=⁺ (>>=⁻ xs p) ≡⟨ >>=⁺∘>>=⁻ xs p ⟩
- p ∎
- where open P.≡-Reasoning
-
- >>=-∈⁻∘>>=-∈⁺ : ∀ {A B : Set} (f : A → List B) {x y xs}
- (x∈xs : x ∈ xs) (y∈fx : y ∈ f x) →
- >>=-∈⁻ f xs (>>=-∈⁺ f x∈xs y∈fx) ≡ (x , x∈xs , y∈fx)
- >>=-∈⁻∘>>=-∈⁺ f {x = x} {xs = xs} x∈xs y∈fx = begin
- find (>>=⁻ xs (>>=⁺ (Any.map _ x∈xs))) ≡⟨ P.cong find (>>=⁻∘>>=⁺ (Any.map _ x∈xs)) ⟩
- find (Any.map _ x∈xs) ≡⟨ find∘map x∈xs _ ⟩
- (x , x∈xs , _) ≡⟨ P.cong (λ p → (x , x∈xs , p)) (lift-resp-id _ _) ⟩
- (x , x∈xs , y∈fx) ∎
- where open P.≡-Reasoning
-
- -- _>>=_ is monotone.
-
- _>>=-mono_ : ∀ {A B} {f g : A → List B} {xs ys} →
- xs ⊆ ys → (∀ {x} → f x ⊆ g x) →
- (xs >>= f) ⊆ (ys >>= g)
- _>>=-mono_ {f = f} {g} =
- M₂.>>=-mono (P.→-to-⟶ f) (P.→-to-⟶ g)
-
- -- Introduction and elimination rules for _⊛_.
-
- ⊛-∈⁺ : ∀ {A B} {fs : List (A → B)} {xs f x} →
- f ∈ fs → x ∈ xs → f x ∈ fs ⊛ xs
- ⊛-∈⁺ f∈fs x∈xs =
- ⊛⁺′ (Any.map (λ f≡g x≡y → P.cong₂ _$_ f≡g x≡y) f∈fs) x∈xs
-
- ⊛-∈⁻ : ∀ {A B} (fs : List (A → B)) xs {fx} →
- fx ∈ fs ⊛ xs → ∃₂ λ f x → f ∈ fs × x ∈ xs × fx ≡ f x
- ⊛-∈⁻ fs xs fx∈ with find $ ⊛⁻ fs xs fx∈
- ... | (f , f∈fs , x∈) with find x∈
- ... | (x , x∈xs , fx≡fx) = (f , x , f∈fs , x∈xs , fx≡fx)
-
- -- _⊛_ is monotone.
-
- _⊛-mono_ : ∀ {A B} {fs gs : List (A → B)} {xs ys} →
- fs ⊆ gs → xs ⊆ ys → fs ⊛ xs ⊆ gs ⊛ ys
- _⊛-mono_ {fs = fs} {xs = xs} fs⊆gs xs⊆ys fx∈ with ⊛-∈⁻ fs xs fx∈
- ... | (f , x , f∈fs , x∈xs , refl) =
- ⊛-∈⁺ (fs⊆gs f∈fs) (xs⊆ys x∈xs)
-
- -- Introduction and elimination rules for _⊗_.
-
- private
-
- lemma₁ : {A B : Set} {p q : A × B} →
- (p ⟨ _≡_ on proj₁ ⟩ q) × (p ⟨ _≡_ on proj₂ ⟩ q) → p ≡ q
- lemma₁ {p = (x , y)} {q = (.x , .y)} (refl , refl) = refl
-
- lemma₂ : {A B : Set} {p q : A × B} →
- p ≡ q → (p ⟨ _≡_ on proj₁ ⟩ q) × (p ⟨ _≡_ on proj₂ ⟩ q)
- lemma₂ = < P.cong proj₁ , P.cong proj₂ >
-
- ⊗-∈⁺ : ∀ {A B} {xs : List A} {ys : List B} {x y} →
- x ∈ xs → y ∈ ys → (x , y) ∈ (xs ⊗ ys)
- ⊗-∈⁺ x∈xs y∈ys = Any.map lemma₁ (⊗⁺′ x∈xs y∈ys)
-
- ⊗-∈⁻ : ∀ {A B} (xs : List A) (ys : List B) {p} →
- p ∈ (xs ⊗ ys) → proj₁ p ∈ xs × proj₂ p ∈ ys
- ⊗-∈⁻ xs ys = ⊗⁻′ xs ys ∘ Any.map lemma₂
-
- -- _⊗_ is monotone.
-
- _⊗-mono_ : ∀ {A B} {xs₁ ys₁ : List A} {xs₂ ys₂ : List B} →
- xs₁ ⊆ ys₁ → xs₂ ⊆ ys₂ → (xs₁ ⊗ xs₂) ⊆ (ys₁ ⊗ ys₂)
- _⊗-mono_ xs₁⊆ys₁ xs₂⊆ys₂ {p} =
- Any.map lemma₁ ∘ M₂._⊗-mono_ xs₁⊆ys₁ xs₂⊆ys₂ {p} ∘ Any.map lemma₂
+ ++-comm∘++-comm : ∀ {A} {P : A → Set} 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
+ ++-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
+ ++-comm∘++-comm (x ∷ xs) {ys} (there .([ ++⁺ʳ xs , ++⁺ˡ ]′ (++⁻ ys (++⁺ʳ ys p))))
+ | inj₁ p | P.refl
+ rewrite ++⁻∘++⁺ ys (inj₂ p)
+ | ++⁻∘++⁺ ys (inj₂ $ there {x = x} p) = P.refl
+ ++-comm∘++-comm (x ∷ xs) {ys} (there .([ ++⁺ʳ xs , ++⁺ˡ ]′ (++⁻ ys (++⁺ˡ p))))
+ | inj₂ p | P.refl
+ rewrite ++⁻∘++⁺ ys {ys = xs} (inj₁ p)
+ | ++⁻∘++⁺ ys {ys = x ∷ xs} (inj₁ p) = P.refl
+
+++⇿++ : ∀ {A} {P : A → Set} 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
+ ; inverse-of = record
+ { left-inverse-of = ++-comm∘++-comm xs
+ ; right-inverse-of = ++-comm∘++-comm ys
+ }
+ }
diff --git a/src/Data/List/Countdown.agda b/src/Data/List/Countdown.agda
index 80ba0de..3bbb50c 100644
--- a/src/Data/List/Countdown.agda
+++ b/src/Data/List/Countdown.agda
@@ -10,9 +10,9 @@ module Data.List.Countdown (D : DecSetoid zero zero) where
open import Data.Empty
open import Data.Fin using (Fin; zero; suc)
-open import Data.Function
-open import Data.Function.Equality using (_⟨$⟩_)
-open import Data.Function.Injection
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Injection
using (Injection; module Injection)
open import Data.List
open import Data.List.Any as Any using (here; there)
@@ -223,9 +223,9 @@ insert {counted} {n} counted⊕1+n x x∉counted =
inj _ _ | no _ | no _ | inj₁ x∈counted | _ | _ | _ | _ | _ = ⊥-elim (x∉counted x∈counted)
inj () _ | no _ | no _ | inj₂ _ | inj₁ _ | _ | _ | _ | _
inj _ () | no _ | no _ | inj₂ _ | _ | inj₁ _ | _ | _ | _
- inj eq₁ eq₂ | no _ | no _ | inj₂ _ | inj₂ _ | inj₂ _ | _ | _ | hlp =
+ inj eq₁ eq₂ | no _ | no _ | inj₂ i | inj₂ _ | inj₂ _ | _ | _ | hlp =
hlp _ refl refl $
- thick-injective _ _ _ $
+ thick-injective i _ _ $
PropEq.trans (drop-inj₂ eq₁) (PropEq.sym (drop-inj₂ eq₂))
-- Counts an element if it has not already been counted.
diff --git a/src/Data/List/NonEmpty.agda b/src/Data/List/NonEmpty.agda
index 2515687..4b8fdf5 100644
--- a/src/Data/List/NonEmpty.agda
+++ b/src/Data/List/NonEmpty.agda
@@ -6,7 +6,7 @@ module Data.List.NonEmpty where
open import Data.Product hiding (map)
open import Data.Nat
-open import Data.Function
+open import Function
open import Data.Vec as Vec using (Vec; []; _∷_)
open import Data.List as List using (List; []; _∷_)
open import Category.Monad
diff --git a/src/Data/List/NonEmpty/Properties.agda b/src/Data/List/NonEmpty/Properties.agda
index 34e71ff..1a8b9ae 100644
--- a/src/Data/List/NonEmpty/Properties.agda
+++ b/src/Data/List/NonEmpty/Properties.agda
@@ -2,15 +2,17 @@
-- Properties of non-empty lists
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Data.List.NonEmpty.Properties where
open import Algebra
open import Category.Monad
-open import Data.Function
+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} = Monoid (List.monoid A)
+private module LM {a} {A : Set a} = Monoid (List.monoid A)
open import Data.List.NonEmpty as List⁺
open RawMonad List⁺.monad
open import Relation.Binary.PropositionalEquality
diff --git a/src/Data/List/Properties.agda b/src/Data/List/Properties.agda
index 95089d3..724dfb5 100644
--- a/src/Data/List/Properties.agda
+++ b/src/Data/List/Properties.agda
@@ -10,21 +10,38 @@
module Data.List.Properties where
open import Algebra
+open import Category.Monad
open import Data.List as List
-private module LM {A} = Monoid (List.monoid A)
open import Data.Nat
open import Data.Nat.Properties
open import Data.Bool
-open import Data.Function
+open import Function
open import Data.Product as Prod hiding (map)
open import Data.Maybe
-open import Relation.Binary.PropositionalEquality
-import Relation.Binary.EqReasoning as Eq
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; _≢_; _≗_; refl)
+import Relation.Binary.EqReasoning as EqR
+
+open RawMonadPlus List.monadPlus
+private
+ module LM {a} {A : Set a} = Monoid (List.monoid A)
∷-injective : ∀ {a} {A : Set a} {x y xs ys} →
- (List A ∶ x ∷ xs) ≡ (y ∷ ys) → (x ≡ y) × (xs ≡ ys)
+ (x ∷ xs ∶ List A) ≡ y ∷ ys → x ≡ y × xs ≡ ys
∷-injective refl = (refl , refl)
+∷ʳ-injective : ∀ {a} {A : Set a} {x y} xs ys →
+ (xs ∷ʳ x ∶ List A) ≡ ys ∷ʳ y → xs ≡ ys × x ≡ y
+∷ʳ-injective [] [] refl = (refl , refl)
+∷ʳ-injective (x ∷ xs) (y ∷ ys) eq with ∷-injective eq
+∷ʳ-injective (x ∷ xs) (.x ∷ ys) eq | (refl , eq′) =
+ Prod.map (P.cong (_∷_ x)) id $ ∷ʳ-injective xs ys eq′
+
+∷ʳ-injective [] (_ ∷ []) ()
+∷ʳ-injective [] (_ ∷ _ ∷ _) ()
+∷ʳ-injective (_ ∷ []) [] ()
+∷ʳ-injective (_ ∷ _ ∷ _) [] ()
+
right-identity-unique : ∀ {a} {A : Set a} (xs : List A) {ys} →
xs ≡ xs ++ ys → ys ≡ []
right-identity-unique [] refl = refl
@@ -38,9 +55,9 @@ left-identity-unique {xs = []} (y ∷ ys) ()
left-identity-unique {xs = x ∷ xs} (y ∷ ys) eq
with left-identity-unique (ys ++ [ x ]) (begin
xs ≡⟨ proj₂ (∷-injective eq) ⟩
- ys ++ x ∷ xs ≡⟨ sym (LM.assoc ys [ x ] xs) ⟩
+ ys ++ x ∷ xs ≡⟨ P.sym (LM.assoc ys [ x ] xs) ⟩
(ys ++ [ x ]) ++ xs ∎)
- where open ≡-Reasoning
+ where open P.≡-Reasoning
left-identity-unique {xs = x ∷ xs} (y ∷ [] ) eq | ()
left-identity-unique {xs = x ∷ xs} (y ∷ _ ∷ _) eq | ()
@@ -50,20 +67,20 @@ map-++-commute : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) xs ys →
map f (xs ++ ys) ≡ map f xs ++ map f ys
map-++-commute f [] ys = refl
map-++-commute f (x ∷ xs) ys =
- cong (_∷_ (f x)) (map-++-commute f xs ys)
+ P.cong (_∷_ (f x)) (map-++-commute f xs ys)
sum-++-commute : ∀ xs ys → sum (xs ++ ys) ≡ sum xs + sum ys
sum-++-commute [] ys = refl
sum-++-commute (x ∷ xs) ys = begin
x + sum (xs ++ ys)
- ≡⟨ cong (_+_ x) (sum-++-commute xs ys) ⟩
+ ≡⟨ P.cong (_+_ x) (sum-++-commute xs ys) ⟩
x + (sum xs + sum ys)
- ≡⟨ sym $ +-assoc x _ _ ⟩
+ ≡⟨ P.sym $ +-assoc x _ _ ⟩
(x + sum xs) + sum ys
where
- open CommutativeSemiring commutativeSemiring hiding (_+_; sym)
- open ≡-Reasoning
+ open CommutativeSemiring commutativeSemiring hiding (_+_)
+ open P.≡-Reasoning
-- Various properties about folds.
@@ -77,10 +94,10 @@ foldr-universal h f e base step (x ∷ xs) = begin
h (x ∷ xs)
≡⟨ step x xs ⟩
f x (h xs)
- ≡⟨ cong (f x) (foldr-universal h f e base step xs) ⟩
+ ≡⟨ P.cong (f x) (foldr-universal h f e base step xs) ⟩
f x (foldr f e xs)
- where open ≡-Reasoning
+ where open P.≡-Reasoning
foldr-fusion : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
(h : B → C) {f : A → B → B} {g : A → C → C} (e : B) →
@@ -98,65 +115,65 @@ idIsFold = foldr-universal id _∷_ [] refl (λ _ _ → refl)
++IsFold xs ys =
begin
xs ++ ys
- ≡⟨ cong (λ xs → xs ++ ys) (idIsFold xs) ⟩
+ ≡⟨ P.cong (λ xs → xs ++ ys) (idIsFold xs) ⟩
foldr _∷_ [] xs ++ ys
≡⟨ foldr-fusion (λ xs → xs ++ ys) [] (λ _ _ → refl) xs ⟩
foldr _∷_ ([] ++ ys) xs
≡⟨ refl ⟩
foldr _∷_ ys xs
- where open ≡-Reasoning
+ where open P.≡-Reasoning
mapIsFold : ∀ {a b} {A : Set a} {B : Set b} {f : A → B} →
map f ≗ foldr (λ x ys → f x ∷ ys) []
mapIsFold {f = f} =
begin
map f
- ≈⟨ cong (map f) ∘ idIsFold ⟩
+ ≈⟨ P.cong (map f) ∘ idIsFold ⟩
map f ∘ foldr _∷_ []
≈⟨ foldr-fusion (map f) [] (λ _ _ → refl) ⟩
foldr (λ x ys → f x ∷ ys) []
- where open Eq (_ →-setoid _)
+ where open EqR (P._→-setoid_ _ _)
concat-map : ∀ {a b} {A : Set a} {B : Set b} {f : A → B} →
concat ∘ map (map f) ≗ map f ∘ concat
concat-map {b = b} {f = f} =
begin
concat ∘ map (map f)
- ≈⟨ cong concat ∘ mapIsFold {b = b} ⟩
+ ≈⟨ P.cong concat ∘ mapIsFold {b = b} ⟩
concat ∘ foldr (λ xs ys → map f xs ∷ ys) []
≈⟨ foldr-fusion {b = b} concat [] (λ _ _ → refl) ⟩
foldr (λ ys zs → map f ys ++ zs) []
- ≈⟨ sym ∘
+ ≈⟨ P.sym ∘
foldr-fusion (map f) [] (λ ys zs → map-++-commute f ys zs) ⟩
map f ∘ concat
- where open Eq (_ →-setoid _)
+ where open EqR (P._→-setoid_ _ _)
map-id : ∀ {a} {A : Set a} → map id ≗ id {A = List A}
-map-id = begin
+map-id {A = A} = begin
map id ≈⟨ mapIsFold ⟩
- foldr _∷_ [] ≈⟨ sym ∘ idIsFold ⟩
+ foldr _∷_ [] ≈⟨ P.sym ∘ idIsFold {A = A} ⟩
id ∎
- where open Eq (_ →-setoid _)
+ where open EqR (P._→-setoid_ _ _)
map-compose : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
{g : B → C} {f : A → B} →
map (g ∘ f) ≗ map g ∘ map f
-map-compose {g = g} {f} =
+map-compose {A = A} {B} {g = g} {f} =
begin
map (g ∘ f)
- ≈⟨ cong (map (g ∘ f)) ∘ idIsFold ⟩
+ ≈⟨ P.cong (map (g ∘ f)) ∘ idIsFold ⟩
map (g ∘ f) ∘ foldr _∷_ []
≈⟨ foldr-fusion (map (g ∘ f)) [] (λ _ _ → refl) ⟩
foldr (λ a y → g (f a) ∷ y) []
- ≈⟨ sym ∘ foldr-fusion (map g) [] (λ _ _ → refl) ⟩
+ ≈⟨ P.sym ∘ foldr-fusion (map g) [] (λ _ _ → refl) ⟩
map g ∘ foldr (λ a y → f a ∷ y) []
- ≈⟨ cong (map g) ∘ sym ∘ mapIsFold ⟩
+ ≈⟨ P.cong (map g) ∘ P.sym ∘ mapIsFold {A = A} {B = B} ⟩
map g ∘ map f
- where open Eq (_ →-setoid _)
+ where open EqR (P._→-setoid_ _ _)
foldr-cong : ∀ {a b} {A : Set a} {B : Set b}
{f₁ f₂ : A → B → B} {e₁ e₂ : B} →
@@ -165,26 +182,26 @@ foldr-cong : ∀ {a b} {A : Set a} {B : Set b}
foldr-cong {f₁ = f₁} {f₂} {e} f₁≗₂f₂ refl =
begin
foldr f₁ e
- ≈⟨ cong (foldr f₁ e) ∘ idIsFold ⟩
+ ≈⟨ P.cong (foldr f₁ e) ∘ idIsFold ⟩
foldr f₁ e ∘ foldr _∷_ []
≈⟨ foldr-fusion (foldr f₁ e) [] (λ x xs → f₁≗₂f₂ x (foldr f₁ e xs)) ⟩
foldr f₂ e
- where open Eq (_ →-setoid _)
+ where open EqR (P._→-setoid_ _ _)
map-cong : ∀ {a b} {A : Set a} {B : Set b} {f g : A → B} →
f ≗ g → map f ≗ map g
-map-cong {f = f} {g} f≗g =
+map-cong {A = A} {B} {f} {g} f≗g =
begin
map f
≈⟨ mapIsFold ⟩
foldr (λ x ys → f x ∷ ys) []
- ≈⟨ foldr-cong (λ x ys → cong₂ _∷_ (f≗g x) refl) refl ⟩
+ ≈⟨ foldr-cong (λ x ys → P.cong₂ _∷_ (f≗g x) refl) refl ⟩
foldr (λ x ys → g x ∷ ys) []
- ≈⟨ sym ∘ mapIsFold ⟩
+ ≈⟨ P.sym ∘ mapIsFold {A = A} {B = B} ⟩
map g
- where open Eq (_ →-setoid _)
+ where open EqR (P._→-setoid_ _ _)
-- Take, drop, and splitAt.
@@ -193,14 +210,14 @@ take++drop : ∀ {a} {A : Set a}
take++drop zero xs = refl
take++drop (suc n) [] = refl
take++drop (suc n) (x ∷ xs) =
- cong (λ xs → x ∷ xs) (take++drop n xs)
+ P.cong (λ xs → x ∷ xs) (take++drop n xs)
splitAt-defn : ∀ {a} {A : Set a} n →
splitAt {A = A} n ≗ < take n , drop n >
splitAt-defn zero xs = refl
splitAt-defn (suc n) [] = refl
splitAt-defn (suc n) (x ∷ xs) with splitAt n xs | splitAt-defn n xs
-... | (ys , zs) | ih = cong (Prod.map (_∷_ x) id) ih
+... | (ys , zs) | ih = P.cong (Prod.map (_∷_ x) id) ih
-- TakeWhile, dropWhile, and span.
@@ -208,14 +225,14 @@ takeWhile++dropWhile : ∀ {a} {A : Set a} (p : A → Bool) (xs : List A) →
takeWhile p xs ++ dropWhile p xs ≡ xs
takeWhile++dropWhile p [] = refl
takeWhile++dropWhile p (x ∷ xs) with p x
-... | true = cong (_∷_ x) (takeWhile++dropWhile p xs)
+... | true = P.cong (_∷_ x) (takeWhile++dropWhile p xs)
... | false = refl
span-defn : ∀ {a} {A : Set a} (p : A → Bool) →
span p ≗ < takeWhile p , dropWhile p >
span-defn p [] = refl
span-defn p (x ∷ xs) with p x
-... | true = cong (Prod.map (_∷_ x) id) (span-defn p xs)
+... | true = P.cong (Prod.map (_∷_ x) id) (span-defn p xs)
... | false = refl
-- Partition.
@@ -225,8 +242,8 @@ partition-defn : ∀ {a} {A : Set a} (p : A → Bool) →
partition-defn p [] = refl
partition-defn p (x ∷ xs)
with p x | partition p xs | partition-defn p xs
-... | true | (ys , zs) | eq = cong (Prod.map (_∷_ x) id) eq
-... | false | (ys , zs) | eq = cong (Prod.map id (_∷_ x)) eq
+... | true | (ys , zs) | eq = P.cong (Prod.map (_∷_ x) id) eq
+... | false | (ys , zs) | eq = P.cong (Prod.map id (_∷_ x)) eq
-- Inits, tails, and scanr.
@@ -239,13 +256,13 @@ scanr-defn f e (x₁ ∷ x₂ ∷ xs)
with scanr f e (x₂ ∷ xs) | scanr-defn f e (x₂ ∷ xs)
... | [] | ()
... | y ∷ ys | eq with ∷-injective eq
-... | y≡fx₂⦇f⦈xs , _ = cong₂ (λ z zs → f x₁ z ∷ zs) y≡fx₂⦇f⦈xs eq
+... | y≡fx₂⦇f⦈xs , _ = P.cong₂ (λ z zs → f x₁ z ∷ zs) y≡fx₂⦇f⦈xs eq
scanl-defn : ∀ {a b} {A : Set a} {B : Set b}
(f : A → B → A) (e : A) →
scanl f e ≗ map (foldl f e) ∘ inits
scanl-defn f e [] = refl
-scanl-defn f e (x ∷ xs) = cong (_∷_ e) (begin
+scanl-defn f e (x ∷ xs) = P.cong (_∷_ e) (begin
scanl f (f e x) xs
≡⟨ scanl-defn f (f e x) xs ⟩
map (foldl f (f e x)) (inits xs)
@@ -254,19 +271,19 @@ scanl-defn f e (x ∷ xs) = cong (_∷_ e) (begin
≡⟨ map-compose (inits xs) ⟩
map (foldl f e) (map (_∷_ x) (inits xs))
∎)
- where open ≡-Reasoning
+ where open P.≡-Reasoning
-- Length.
length-map : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) xs →
length (map f xs) ≡ length xs
length-map f [] = refl
-length-map f (x ∷ xs) = cong suc (length-map f xs)
+length-map f (x ∷ xs) = P.cong suc (length-map f xs)
length-++ : ∀ {a} {A : Set a} (xs : List A) {ys} →
length (xs ++ ys) ≡ length xs + length ys
length-++ [] = refl
-length-++ (x ∷ xs) = cong suc (length-++ xs)
+length-++ (x ∷ xs) = P.cong suc (length-++ xs)
length-gfilter : ∀ {a b} {A : Set a} {B : Set b} (p : A → Maybe B) xs →
length (gfilter p xs) ≤ length xs
@@ -274,3 +291,167 @@ length-gfilter p [] = z≤n
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)
+
+-- The list monad.
+
+module Monad where
+
+ left-zero : {A B : Set} (f : A → List B) → (∅ >>= f) ≡ ∅
+ left-zero f = refl
+
+ right-zero : {A B : Set} (xs : List A) →
+ (xs >>= const ∅) ≡ (∅ ∶ List B)
+ right-zero [] = refl
+ right-zero (x ∷ xs) = right-zero xs
+
+ private
+
+ not-left-distributive :
+ let xs = true ∷ false ∷ []; f = return; g = return in
+ (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) →
+ (xs ∣ ys >>= f) ≡ ((xs >>= f) ∣ (ys >>= f))
+ right-distributive [] ys f = refl
+ right-distributive (x ∷ xs) ys f = begin
+ f x ∣ (xs ∣ ys >>= f) ≡⟨ P.cong (_∣_ (f x)) $ right-distributive xs ys f ⟩
+ f x ∣ ((xs >>= f) ∣ (ys >>= f)) ≡⟨ P.sym $ LM.assoc (f x) _ _ ⟩
+ (f x ∣ (xs >>= f)) ∣ (ys >>= f) ∎
+ where open P.≡-Reasoning
+
+ 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)
+
+ right-identity : {A : Set} (xs : List A) →
+ (xs >>= return) ≡ xs
+ right-identity [] = refl
+ right-identity (x ∷ xs) = P.cong (_∷_ x) (right-identity xs)
+
+ 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
+ associative (x ∷ xs) f g = begin
+ (f x >>= g) ∣ (xs >>= λ x → f x >>= g) ≡⟨ P.cong (_∣_ (f x >>= g)) $ associative xs f g ⟩
+ (f x >>= g) ∣ (xs >>= f >>= g) ≡⟨ P.sym $ right-distributive (f x) (xs >>= f) g ⟩
+ (f x ∣ (xs >>= f) >>= g) ∎
+ where open P.≡-Reasoning
+
+ 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)
+
+-- The applicative functor derived from the list monad.
+
+-- Note that these proofs (almost) show that RawIMonad.rawIApplicative
+-- is correctly defined. The proofs can be reused if proof components
+-- are ever added to RawIMonad and RawIApplicative.
+
+module Applicative where
+
+ open P.≡-Reasoning
+
+ private
+
+ -- A variant of flip map.
+
+ 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 xs = begin
+ ∅ ⊛ xs ≡⟨ refl ⟩
+ (∅ >>= pam xs) ≡⟨ Monad.left-zero (pam xs) ⟩
+ ∅ ∎
+
+ -- ∅ is a right zero for _⊛_.
+
+ right-zero : ∀ {A B} (fs : List (A → B)) → fs ⊛ ∅ ≡ ∅
+ right-zero fs = begin
+ fs ⊛ ∅ ≡⟨ refl ⟩
+ (fs >>= pam ∅) ≡⟨ (Monad.cong (refl {x = fs}) λ 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 →
+ (fs₁ ∣ fs₂) ⊛ xs ≡ (fs₁ ⊛ xs ∣ fs₂ ⊛ xs)
+ right-distributive fs₁ fs₂ xs = begin
+ (fs₁ ∣ fs₂) ⊛ xs ≡⟨ refl ⟩
+ (fs₁ ∣ fs₂ >>= pam xs) ≡⟨ Monad.right-distributive fs₁ fs₂ (pam xs) ⟩
+ (fs₁ >>= pam xs) ∣ (fs₂ >>= pam xs) ≡⟨ refl ⟩
+ fs₁ ⊛ xs ∣ fs₂ ⊛ xs ∎
+
+ -- _⊛_ does not distribute over _∣_ from the left.
+
+ private
+
+ not-left-distributive :
+ let fs = id ∷ id ∷ []; xs₁ = true ∷ []; xs₂ = true ∷ false ∷ [] in
+ fs ⊛ (xs₁ ∣ xs₂) ≢ (fs ⊛ xs₁ ∣ fs ⊛ xs₂)
+ not-left-distributive ()
+
+ -- Applicative functor laws.
+
+ identity : ∀ {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) ⟩
+ (xs >>= return) ≡⟨ Monad.right-identity xs ⟩
+ xs ∎
+
+ private
+
+ 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
+ (pam xs f >>= fs) ≡⟨ P.sym $ Monad.associative xs (return ∘ f) fs ⟩
+ (xs >>= λ x → return (f x) >>= fs) ≡⟨ Monad.cong (refl {x = xs}) (λ x → Monad.left-identity (f x) fs) ⟩
+ (xs >>= λ x → fs (f x)) ∎
+
+ composition :
+ ∀ {A B C} (fs : List (B → C)) (gs : List (A → B)) xs →
+ return _∘′_ ⊛ fs ⊛ gs ⊛ xs ≡ fs ⊛ (gs ⊛ xs)
+ 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) ⟩
+ (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 →
+ pam-lemma gs (_∘′_ f) (pam xs)) ⟩
+ (fs >>= λ f → gs >>= λ g → pam xs (f ∘′ g)) ≡⟨ (Monad.cong (refl {x = fs}) λ f →
+ Monad.cong (refl {x = gs}) λ g →
+ P.sym $ pam-lemma xs g (return ∘ f)) ⟩
+ (fs >>= λ f → gs >>= λ g → pam (pam xs g) f) ≡⟨ (Monad.cong (refl {x = fs}) λ f →
+ Monad.associative gs (pam xs) (return ∘ f)) ⟩
+ (fs >>= pam (gs >>= pam xs)) ≡⟨ refl ⟩
+ fs ⊛ (gs ⊛ xs) ∎
+
+ homomorphism : ∀ {A B : Set} (f : A → B) x →
+ return f ⊛ return x ≡ return (f x)
+ homomorphism f x = begin
+ return f ⊛ return x ≡⟨ refl ⟩
+ (return f >>= pam (return x)) ≡⟨ Monad.left-identity f (pam (return x)) ⟩
+ pam (return x) f ≡⟨ Monad.left-identity x (return ∘ f) ⟩
+ return (f x) ∎
+
+ interchange : ∀ {A B} (fs : List (A → B)) {x} →
+ fs ⊛ return x ≡ return (λ f → f x) ⊛ fs
+ interchange fs {x} = begin
+ fs ⊛ return x ≡⟨ refl ⟩
+ (fs >>= pam (return x)) ≡⟨ (Monad.cong (refl {x = fs}) λ f →
+ Monad.left-identity x (return ∘ f)) ⟩
+ (fs >>= λ f → return (f x)) ≡⟨ refl ⟩
+ (pam fs (λ f → f x)) ≡⟨ P.sym $ Monad.left-identity (λ f → f x) (pam fs) ⟩
+ (return (λ f → f x) >>= pam fs) ≡⟨ refl ⟩
+ return (λ f → f x) ⊛ fs ∎
diff --git a/src/Data/Map.agda b/src/Data/Map.agda
deleted file mode 100644
index d17dc2f..0000000
--- a/src/Data/Map.agda
+++ /dev/null
@@ -1,85 +0,0 @@
-------------------------------------------------------------------------
--- Finite maps, i.e. lookup tables (currently only some type
--- signatures)
-------------------------------------------------------------------------
-
-module Data.Map where
-
-open import Relation.Nullary
-open import Relation.Binary
-open import Data.List as L using (List)
-open import Data.Product
-open import Level
-
-module Map₁ (key-dto : DecTotalOrder zero zero zero)
- (elem-s : Setoid zero zero) where
-
- open DecTotalOrder key-dto renaming (Carrier to key)
- open Setoid elem-s renaming (Carrier to elem; _≈_ to _≗_)
-
- infixr 6 _∪_
- infix 5 _∈?_
- infix 4 _∈_ _|≈|_
-
- abstract postulate decSetoid : DecSetoid _ _
-
- Map : Set
- Map = Setoid.Carrier (DecSetoid.setoid decSetoid)
-
- _|≈|_ : Rel Map zero
- _|≈|_ = Setoid._≈_ (DecSetoid.setoid decSetoid)
-
- abstract
- postulate
- empty : Map
- insert : key → elem → Map → Map
- _∪_ : Map → Map → Map
- _∈_ : key → Map → Set
- _↦_∈_ : key → elem → Map → Set
-
- data LookupResult (k : key) (s : Map) : Set where
- found : (e : elem) (k↦e∈s : k ↦ e ∈ s) → LookupResult k s
- notFound : (k∉s : ¬ k ∈ s) → LookupResult k s
-
- abstract
- postulate
- _∈?_ : (k : key) → (s : Map) → LookupResult k s
- toList : Map → List (key × elem)
-
- postulate
- prop-∈₁ : ∀ {x v s} → x ↦ v ∈ s → x ∈ s
- prop-∈₂ : ∀ {x s} → x ∈ s → Σ elem (λ v → x ↦ v ∈ s)
- prop-∈₃ : ∀ {x v w s} → x ↦ v ∈ s → x ↦ w ∈ s → v ≗ w
-
- prop-∈-insert₁ : ∀ {x y v w s} →
- x ≈ y → v ≗ w → x ↦ v ∈ insert y w s
- prop-∈-insert₂ : ∀ {x y v w s} →
- ¬ x ≈ y → x ↦ v ∈ s → x ↦ v ∈ insert y w s
- prop-∈-insert₃ : ∀ {x y v w s} →
- ¬ x ≈ y → x ↦ v ∈ insert y w s → x ↦ v ∈ s
-
- prop-∈-empty : ∀ {x} → ¬ x ∈ empty
-
- prop-∈-∪ : ∀ {x s₁ s₂} → x ∈ s₁ → x ∈ s₁ ∪ s₂
-
- prop-∪₁ : ∀ {s₁ s₂} → s₁ ∪ s₂ |≈| s₂ ∪ s₁
- prop-∪₂ : ∀ {s₁ s₂ s₃} → s₁ ∪ (s₂ ∪ s₃) |≈| (s₁ ∪ s₂) ∪ s₃
-
- prop-∈-|≈|₁ : ∀ {x} → (λ s → x ∈ s) Respects _|≈|_
- prop-∈-|≈|₂ : ∀ {x v} → (λ s → x ↦ v ∈ s) Respects _|≈|_
- prop-∈-≈₁ : ∀ {s} → (λ x → x ∈ s) Respects _≈_
- prop-∈-≈₂ : ∀ {v s} → (λ x → x ↦ v ∈ s) Respects _≈_
- prop-∈-≗ : ∀ {x s} → (λ v → x ↦ v ∈ s) Respects _≗_
-
- -- TODO: Postulates for toList.
-
- singleton : key → elem → Map
- singleton k v = insert k v empty
-
- ⋃_ : List Map → Map
- ⋃_ = L.foldr _∪_ empty
-
- fromList : List (key × elem) → Map
- fromList = L.foldr (uncurry insert) empty
-
-open Map₁ public renaming (Map to _⇰_)
diff --git a/src/Data/Maybe.agda b/src/Data/Maybe.agda
index f007b2f..2e7e8fc 100644
--- a/src/Data/Maybe.agda
+++ b/src/Data/Maybe.agda
@@ -36,7 +36,7 @@ maybe j n nothing = n
------------------------------------------------------------------------
-- Maybe monad
-open import Data.Function
+open import Function
open import Category.Functor
open import Category.Monad
open import Category.Monad.Identity
@@ -77,27 +77,60 @@ monadPlus = record
open import Relation.Nullary
open import Relation.Binary
-open import Relation.Binary.PropositionalEquality as PropEq
- using (_≡_; refl)
-drop-just : ∀ {A : Set} {x y : A} → just x ≡ just y → x ≡ y
-drop-just refl = refl
-
-decSetoid : {A : Set} → Decidable (_≡_ {A = A}) → DecSetoid _ _
-decSetoid {A} _A-≟_ = PropEq.decSetoid _≟_
+data Eq {a ℓ} {A : Set a} (_≈_ : Rel A ℓ) : Rel (Maybe A) (a ⊔ ℓ) where
+ just : ∀ {x y} (x≈y : x ≈ y) → Eq _≈_ (just x) (just y)
+ nothing : Eq _≈_ nothing nothing
+
+drop-just : ∀ {a ℓ} {A : Set a} {_≈_ : Rel A ℓ} {x y : A} →
+ just x ⟨ Eq _≈_ ⟩ just y → x ≈ y
+drop-just (just x≈y) = x≈y
+
+setoid : ∀ {ℓ₁ ℓ₂} → Setoid ℓ₁ ℓ₂ → Setoid _ _
+setoid S = record
+ { Carrier = Maybe S.Carrier
+ ; _≈_ = _≈_
+ ; isEquivalence = record
+ { refl = refl
+ ; sym = sym
+ ; trans = trans
+ }
+ }
+ where
+ module S = Setoid S
+ _≈_ = Eq S._≈_
+
+ refl : ∀ {x} → x ≈ x
+ refl {just x} = just S.refl
+ refl {nothing} = nothing
+
+ sym : ∀ {x y} → x ≈ y → y ≈ x
+ sym (just x≈y) = just (S.sym x≈y)
+ sym nothing = nothing
+
+ trans : ∀ {x y z} → x ≈ y → y ≈ z → x ≈ z
+ trans (just x≈y) (just y≈z) = just (S.trans x≈y y≈z)
+ trans nothing nothing = nothing
+
+decSetoid : ∀ {ℓ₁ ℓ₂} → DecSetoid ℓ₁ ℓ₂ → DecSetoid _ _
+decSetoid D = record
+ { isDecEquivalence = record
+ { isEquivalence = Setoid.isEquivalence (setoid (DecSetoid.setoid D))
+ ; _≟_ = _≟_
+ }
+ }
where
- _≟_ : Decidable (_≡_ {A = Maybe A})
- just x ≟ just y with x A-≟ y
- just x ≟ just .x | yes refl = yes refl
- just x ≟ just y | no x≢y = no (x≢y ∘ drop-just)
+ _≟_ : Decidable (Eq (DecSetoid._≈_ D))
+ just x ≟ just y with DecSetoid._≟_ D x y
+ just x ≟ just y | yes x≈y = yes (just x≈y)
+ just x ≟ just y | no x≉y = no (x≉y ∘ drop-just)
just x ≟ nothing = no λ()
nothing ≟ just y = no λ()
- nothing ≟ nothing = yes refl
+ nothing ≟ nothing = yes nothing
------------------------------------------------------------------------
-- Any and All
-open import Data.Product using (_,_)
open import Data.Empty using (⊥)
import Relation.Nullary.Decidable as Dec
@@ -108,11 +141,11 @@ data All {a} {A : Set a} (P : A → Set a) : Maybe A → Set a where
just : ∀ {x} (px : P x) → All P (just x)
nothing : All P nothing
-IsJust : ∀ {A : Set} → Maybe A → Set
-IsJust = Any (λ _ → ⊤)
+IsJust : ∀ {a} {A : Set a} → Maybe A → Set a
+IsJust = Any (λ _ → Lift ⊤)
-IsNothing : ∀ {A} → Maybe A → Set
-IsNothing = All (λ _ → ⊥)
+IsNothing : ∀ {a} {A : Set a} → Maybe A → Set a
+IsNothing = All (λ _ → Lift ⊥)
private
@@ -125,9 +158,9 @@ private
anyDec : ∀ {A} {P : A → Set} →
(∀ x → Dec (P x)) → (x : Maybe A) → Dec (Any P x)
anyDec p nothing = no λ()
-anyDec p (just x) = Dec.map (just , drop-just-Any) (p x)
+anyDec p (just x) = Dec.map′ just drop-just-Any (p x)
allDec : ∀ {A} {P : A → Set} →
(∀ x → Dec (P x)) → (x : Maybe A) → Dec (All P x)
allDec p nothing = yes nothing
-allDec p (just x) = Dec.map (just , drop-just-All) (p x)
+allDec p (just x) = Dec.map′ just drop-just-All (p x)
diff --git a/src/Data/Nat.agda b/src/Data/Nat.agda
index e5be372..feacbd7 100644
--- a/src/Data/Nat.agda
+++ b/src/Data/Nat.agda
@@ -6,9 +6,9 @@
module Data.Nat where
-open import Data.Function
-open import Data.Function.Equality as F using (_⟨$⟩_)
-open import Data.Function.Injection
+open import Function
+open import Function.Equality as F using (_⟨$⟩_)
+open import Function.Injection
using (Injection; module Injection)
open import Data.Sum
open import Data.Empty
@@ -52,9 +52,9 @@ m > n = n < m
infix 4 _≤′_ _<′_ _≥′_ _>′_
-data _≤′_ : Rel ℕ zero where
- ≤′-refl : ∀ {n} → n ≤′ n
- ≤′-step : ∀ {m n} (m≤′n : m ≤′ n) → m ≤′ suc n
+data _≤′_ (m : ℕ) : ℕ → Set where
+ ≤′-refl : m ≤′ m
+ ≤′-step : ∀ {n} (m≤′n : m ≤′ n) → m ≤′ suc n
_<′_ : Rel ℕ zero
m <′ n = suc m ≤′ n
@@ -204,7 +204,6 @@ decTotalOrder = record
{ isEquivalence = PropEq.isEquivalence
; reflexive = refl′
; trans = trans
- ; ∼-resp-≈ = PropEq.resp₂ _≤_
}
; antisym = antisym
}
@@ -236,5 +235,11 @@ decTotalOrder = record
... | inj₂ n≤m = inj₂ (s≤s n≤m)
import Relation.Binary.PartialOrderReasoning as POR
-module ≤-Reasoning = POR (DecTotalOrder.poset decTotalOrder)
- renaming (_≈⟨_⟩_ to _≡⟨_⟩_)
+module ≤-Reasoning where
+ open POR (DecTotalOrder.poset decTotalOrder) public
+ renaming (_≈⟨_⟩_ to _≡⟨_⟩_)
+
+ infixr 2 _<⟨_⟩_
+
+ _<⟨_⟩_ : ∀ x {y z} → x < y → y IsRelatedTo z → suc x IsRelatedTo z
+ x <⟨ x<y ⟩ y≤z = suc x ≤⟨ x<y ⟩ y≤z
diff --git a/src/Data/Nat/Coprimality.agda b/src/Data/Nat/Coprimality.agda
index 875d805..45853cf 100644
--- a/src/Data/Nat/Coprimality.agda
+++ b/src/Data/Nat/Coprimality.agda
@@ -10,7 +10,7 @@ open import Data.Nat.Divisibility as Div
open import Data.Nat.GCD
open import Data.Nat.GCD.Lemmas
open import Data.Product as Prod
-open import Data.Function
+open import Function
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; _≢_; refl)
open import Relation.Nullary
diff --git a/src/Data/Nat/DivMod.agda b/src/Data/Nat/DivMod.agda
index a1dbdc5..70a6c01 100644
--- a/src/Data/Nat/DivMod.agda
+++ b/src/Data/Nat/DivMod.agda
@@ -13,7 +13,7 @@ open import Induction.Nat
open import Relation.Nullary.Decidable
open import Relation.Binary.PropositionalEquality
open ≡-Reasoning
-open import Data.Function
+open import Function
------------------------------------------------------------------------
-- Some boring lemmas
@@ -65,7 +65,7 @@ data DivMod' (dividend divisor : ℕ) : Set where
-- Integer division with remainder.
--- Note that Induction.Nat.<-rec is used to ensure termination of
+-- Note that Induction.Nat.<-rec is used to establish termination of
-- division. The run-time complexity of this implementation of integer
-- division should be linear in the size of the dividend, assuming
-- that well-founded recursion and the equality type are optimised
@@ -87,7 +87,7 @@ _divMod'_ m n {≢0} = <-rec Pred dm m n {≢0}
dm m rec zero {≢0 = ()}
dm zero rec (suc n) = result 0 zero refl
dm (suc m) rec (suc n) with compare m n
- dm (suc m) rec (suc .(suc m + k)) | less .m k = result 0 r (lem₁ m k)
+ dm (suc m) rec (suc .(suc m + k)) | less .m k = result 0 r (lem₁ m k)
where r = suc (Fin.inject+ k (fromℕ m))
dm (suc m) rec (suc .m) | equal .m = result 1 zero (lem₂ m)
dm (suc .(suc n + k)) rec (suc n) | greater .n k =
diff --git a/src/Data/Nat/Divisibility.agda b/src/Data/Nat/Divisibility.agda
index ea94377..262e920 100644
--- a/src/Data/Nat/Divisibility.agda
+++ b/src/Data/Nat/Divisibility.agda
@@ -19,7 +19,7 @@ open import Relation.Binary
import Relation.Binary.PartialOrderReasoning as PartialOrderReasoning
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; _≢_; refl; sym; cong; subst)
-open import Data.Function
+open import Function
-- m ∣ n is inhabited iff m divides n. Some sources, like Hardy and
-- Wright's "An Introduction to the Theory of Numbers", require m to
@@ -59,7 +59,6 @@ poset = record
{ isEquivalence = PropEq.isEquivalence
; reflexive = reflexive
; trans = trans
- ; ∼-resp-≈ = PropEq.resp₂ _∣_
}
; antisym = antisym
}
diff --git a/src/Data/Nat/GCD.agda b/src/Data/Nat/GCD.agda
index 13d6479..29d8bad 100644
--- a/src/Data/Nat/GCD.agda
+++ b/src/Data/Nat/GCD.agda
@@ -13,7 +13,7 @@ open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
open import Induction
open import Induction.Nat
open import Induction.Lexicographic
-open import Data.Function
+open import Function
open import Data.Nat.GCD.Lemmas
------------------------------------------------------------------------
@@ -76,18 +76,18 @@ open GCD public using (GCD)
module Bézout where
- -- If m and n have greatest common divisor d, then one of the
- -- following two equations is satisfied, for some numbers x and y.
- -- The proof is "lemma" below (Bézout's lemma).
- --
- -- (If this identity was stated using integers instead of natural
- -- numbers, then it would not be necessary to have two equations.)
+ module Identity where
- data Identity (d m n : ℕ) : Set where
- +- : (x y : ℕ) (eq : d + y * n ≡ x * m) → Identity d m n
- -+ : (x y : ℕ) (eq : d + x * m ≡ y * n) → Identity d m n
+ -- If m and n have greatest common divisor d, then one of the
+ -- following two equations is satisfied, for some numbers x and y.
+ -- The proof is "lemma" below (Bézout's lemma).
+ --
+ -- (If this identity was stated using integers instead of natural
+ -- numbers, then it would not be necessary to have two equations.)
- module Identity where
+ data Identity (d m n : ℕ) : Set where
+ +- : (x y : ℕ) (eq : d + y * n ≡ x * m) → Identity d m n
+ -+ : (x y : ℕ) (eq : d + x * m ≡ y * n) → Identity d m n
-- Various properties about Identity.
@@ -117,14 +117,16 @@ module Bézout where
step {d} (-+ .x .(x ⊕ i) eq) | less x i = -+ (2 * x ⊕ i) (x ⊕ i) (lem₆ d x eq)
step {d} {n} (-+ .(y ⊕ i) .y eq) | greater y i = -+ (2 * y ⊕ i) y (lem₇ d y n eq)
- -- This type packs up the gcd, the proof that it is a gcd, and the
- -- proof that it satisfies Bézout's identity.
-
- data Lemma (m n : ℕ) : Set where
- result : (d : ℕ) (g : GCD m n d) (b : Identity d m n) → Lemma m n
+ open Identity public using (Identity; +-; -+)
module Lemma where
+ -- This type packs up the gcd, the proof that it is a gcd, and the
+ -- proof that it satisfies Bézout's identity.
+
+ data Lemma (m n : ℕ) : Set where
+ result : (d : ℕ) (g : GCD m n d) (b : Identity d m n) → Lemma m n
+
-- Various properties about Lemma.
sym : Symmetric Lemma
@@ -144,6 +146,8 @@ module Bézout where
stepʳ : ∀ {n k} → Lemma (suc k) n → Lemma (suc (n + k)) n
stepʳ = sym ∘ stepˡ ∘ sym
+ open Lemma public using (Lemma; result)
+
-- Bézout's lemma proved using some variant of the extended
-- Euclidean algorithm.
diff --git a/src/Data/Nat/GCD/Lemmas.agda b/src/Data/Nat/GCD/Lemmas.agda
index 767732b..ddef011 100644
--- a/src/Data/Nat/GCD/Lemmas.agda
+++ b/src/Data/Nat/GCD/Lemmas.agda
@@ -9,7 +9,7 @@ import Data.Nat.Properties as NatProp
open NatProp.SemiringSolver
open import Relation.Binary.PropositionalEquality
open ≡-Reasoning
-open import Data.Function
+open import Function
lem₀ = solve 2 (λ n k → n :+ (con 1 :+ k) := con 1 :+ n :+ k) refl
diff --git a/src/Data/Nat/InfinitelyOften.agda b/src/Data/Nat/InfinitelyOften.agda
index 735f8d9..efafabd 100644
--- a/src/Data/Nat/InfinitelyOften.agda
+++ b/src/Data/Nat/InfinitelyOften.agda
@@ -7,7 +7,7 @@ module Data.Nat.InfinitelyOften where
open import Algebra
open import Category.Monad
open import Data.Empty
-open import Data.Function
+open import Function
open import Data.Nat
import Data.Nat.Properties as NatProp
open import Data.Product as Prod hiding (map)
diff --git a/src/Data/Nat/LCM.agda b/src/Data/Nat/LCM.agda
index 7bf11f6..4e41e8a 100644
--- a/src/Data/Nat/LCM.agda
+++ b/src/Data/Nat/LCM.agda
@@ -11,7 +11,7 @@ open import Data.Nat.GCD
open import Data.Nat.Divisibility as Div
open import Data.Nat.Coprimality as Coprime
open import Data.Product
-open import Data.Function
+open import Function
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; refl; inspect; _with-≡_)
open import Algebra
diff --git a/src/Data/Nat/Properties.agda b/src/Data/Nat/Properties.agda
index e2ff36e..4e22dd7 100644
--- a/src/Data/Nat/Properties.agda
+++ b/src/Data/Nat/Properties.agda
@@ -12,7 +12,7 @@ open ≤-Reasoning
renaming (begin_ to start_; _∎ to _□; _≡⟨_⟩_ to _≡⟨_⟩'_)
open import Relation.Binary
open DecTotalOrder Nat.decTotalOrder using () renaming (refl to ≤-refl)
-open import Data.Function
+open import Function
open import Algebra
open import Algebra.Structures
open import Relation.Nullary
@@ -96,43 +96,20 @@ private
n * suc m
- distrib-*-+ : _*_ DistributesOver _+_
- distrib-*-+ = distˡ , distʳ
- where
- distˡ : _*_ DistributesOverˡ _+_
- distˡ zero n o = refl
- distˡ (suc m) n o =
- begin
- suc m * (n + o)
- ≡⟨ refl ⟩
- (n + o) + m * (n + o)
- ≡⟨ cong (λ x → (n + o) + x) (distˡ m n o) ⟩
- (n + o) + (m * n + m * o)
- ≡⟨ sym $ +-assoc (n + o) (m * n) (m * o) ⟩
- ((n + o) + m * n) + m * o
- ≡⟨ cong (λ x → x + (m * o)) $ +-assoc n o (m * n) ⟩
- (n + (o + m * n)) + m * o
- ≡⟨ cong (λ x → (n + x) + m * o) $ +-comm o (m * n) ⟩
- (n + (m * n + o)) + m * o
- ≡⟨ cong (λ x → x + (m * o)) $ sym $ +-assoc n (m * n) o ⟩
- ((n + m * n) + o) + m * o
- ≡⟨ +-assoc (n + m * n) o (m * o) ⟩
- (n + m * n) + (o + m * o)
- ≡⟨ refl ⟩
- suc m * n + suc m * o
- ∎
-
- distʳ : _*_ DistributesOverʳ _+_
- distʳ m n o =
- begin
- (n + o) * m
- ≡⟨ *-comm (n + o) m ⟩
- m * (n + o)
- ≡⟨ distˡ m n o ⟩
- m * n + m * o
- ≡⟨ cong₂ _+_ (*-comm m n) (*-comm m o) ⟩
- n * m + o * m
- ∎
+ distribʳ-*-+ : _*_ DistributesOverʳ _+_
+ distribʳ-*-+ m zero o = refl
+ distribʳ-*-+ m (suc n) o =
+ begin
+ (suc n + o) * m
+ ≡⟨ refl ⟩
+ m + (n + o) * m
+ ≡⟨ cong (_+_ m) $ distribʳ-*-+ m n o ⟩
+ m + (n * m + o * m)
+ ≡⟨ sym $ +-assoc m (n * m) (o * m) ⟩
+ m + n * m + o * m
+ ≡⟨ refl ⟩
+ suc n * m + o * m
+ ∎
*-assoc : Associative _*_
*-assoc zero n o = refl
@@ -141,7 +118,7 @@ private
(suc m * n) * o
≡⟨ refl ⟩
(n + m * n) * o
- ≡⟨ proj₂ distrib-*-+ o n (m * n) ⟩
+ ≡⟨ distribʳ-*-+ o n (m * n) ⟩
n * o + (m * n) * o
≡⟨ cong (λ x → n * o + x) $ *-assoc m n o ⟩
n * o + m * (n * o)
@@ -149,52 +126,31 @@ private
suc m * (n * o)
- *-identity : Identity 1 _*_
- *-identity = proj₂ +-identity , n*1≡n
- where
- n*1≡n : RightIdentity 1 _*_
- n*1≡n n =
- begin
- n * 1
- ≡⟨ *-comm n 1 ⟩
- 1 * n
- ≡⟨ refl ⟩
- n + 0
- ≡⟨ proj₂ +-identity n ⟩
- n
- ∎
-
isCommutativeSemiring : IsCommutativeSemiring _≡_ _+_ _*_ 0 1
isCommutativeSemiring = record
- { isSemiring = record
- { isSemiringWithoutAnnihilatingZero = record
- { +-isCommutativeMonoid = record
- { isMonoid = record
- { isSemigroup = record
- { isEquivalence = PropEq.isEquivalence
- ; assoc = +-assoc
- ; ∙-cong = cong₂ _+_
- }
- ; identity = +-identity
- }
- ; comm = +-comm
- }
- ; *-isMonoid = record
- { isSemigroup = record
- { isEquivalence = PropEq.isEquivalence
- ; assoc = *-assoc
- ; ∙-cong = cong₂ _*_
- }
- ; identity = *-identity
- }
- ; distrib = distrib-*-+
+ { +-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = PropEq.isEquivalence
+ ; assoc = +-assoc
+ ; ∙-cong = cong₂ _+_
}
- ; zero = *-zero
+ ; identityˡ = proj₁ +-identity
+ ; comm = +-comm
}
- ; *-comm = *-comm
+ ; *-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = PropEq.isEquivalence
+ ; assoc = *-assoc
+ ; ∙-cong = cong₂ _*_
+ }
+ ; identityˡ = proj₂ +-identity
+ ; comm = *-comm
+ }
+ ; distribʳ = distribʳ-*-+
+ ; zeroˡ = proj₁ *-zero
}
-commutativeSemiring : CommutativeSemiring
+commutativeSemiring : CommutativeSemiring _ _
commutativeSemiring = record
{ _+_ = _+_
; _*_ = _*_
@@ -291,29 +247,27 @@ private
: IsCommutativeSemiringWithoutOne _≡_ _⊔_ _⊓_ 0
⊔-⊓-0-isCommutativeSemiringWithoutOne = record
{ isSemiringWithoutOne = record
- { +-isCommutativeMonoid = record
- { isMonoid = record
- { isSemigroup = record
- { isEquivalence = PropEq.isEquivalence
- ; assoc = ⊔-assoc
- ; ∙-cong = cong₂ _⊔_
- }
- ; identity = ⊔-identity
- }
- ; comm = ⊔-comm
- }
- ; *-isSemigroup = record
- { isEquivalence = PropEq.isEquivalence
- ; assoc = ⊓-assoc
- ; ∙-cong = cong₂ _⊓_
- }
- ; distrib = distrib-⊓-⊔
- ; zero = ⊓-zero
+ { +-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = PropEq.isEquivalence
+ ; assoc = ⊔-assoc
+ ; ∙-cong = cong₂ _⊔_
+ }
+ ; identityˡ = proj₁ ⊔-identity
+ ; comm = ⊔-comm
}
+ ; *-isSemigroup = record
+ { isEquivalence = PropEq.isEquivalence
+ ; assoc = ⊓-assoc
+ ; ∙-cong = cong₂ _⊓_
+ }
+ ; distrib = distrib-⊓-⊔
+ ; zero = ⊓-zero
+ }
; *-comm = ⊓-comm
}
-⊔-⊓-0-commutativeSemiringWithoutOne : CommutativeSemiringWithoutOne
+⊔-⊓-0-commutativeSemiringWithoutOne : CommutativeSemiringWithoutOne _ _
⊔-⊓-0-commutativeSemiringWithoutOne = record
{ _+_ = _⊔_
; _*_ = _⊓_
@@ -362,7 +316,7 @@ isDistributiveLattice = record
; ∨-∧-distribʳ = proj₂ distrib-⊓-⊔
}
-distributiveLattice : DistributiveLattice
+distributiveLattice : DistributiveLattice _ _
distributiveLattice = record
{ _∨_ = _⊓_
; _∧_ = _⊔_
diff --git a/src/Data/Nat/Show.agda b/src/Data/Nat/Show.agda
index 93adef6..65a31a1 100644
--- a/src/Data/Nat/Show.agda
+++ b/src/Data/Nat/Show.agda
@@ -8,7 +8,7 @@ open import Data.Nat
open import Relation.Nullary.Decidable using (True)
open import Data.String as String using (String)
open import Data.Digit
-open import Data.Function
+open import Function
open import Data.List
-- showInBase b n is a string containing the representation of n in
diff --git a/src/Data/Product.agda b/src/Data/Product.agda
index 85bd4c6..0ce4716 100644
--- a/src/Data/Product.agda
+++ b/src/Data/Product.agda
@@ -6,9 +6,9 @@
module Data.Product where
-open import Data.Function
+open import Function
open import Level
-open import Relation.Nullary.Core
+open import Relation.Nullary
infixr 4 _,_ _,′_
infix 4 ,_
@@ -17,8 +17,15 @@ infixr 2 _×_ _-×-_ _-,-_
------------------------------------------------------------------------
-- Definition
-data Σ {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where
- _,_ : (x : A) (y : B x) → Σ A B
+record Σ {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where
+ constructor _,_
+ field
+ proj₁ : A
+ proj₂ : B proj₁
+
+open Σ public
+
+syntax Σ A (λ x → B) = Σ[ x ∶ A ] B
∃ : ∀ {a b} {A : Set a} → (A → Set b) → Set (a ⊔ b)
∃ = Σ _
@@ -31,7 +38,7 @@ data Σ {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where
∃₂ C = ∃ λ a → ∃ λ b → C a b
_×_ : ∀ {a b} (A : Set a) (B : Set b) → Set (a ⊔ b)
-A × B = Σ A (λ _ → B)
+A × B = Σ[ x ∶ A ] B
_,′_ : ∀ {a b} {A : Set a} {B : Set b} → A → B → A × B
_,′_ = _,_
@@ -53,13 +60,6 @@ _,′_ = _,_
,_ : ∀ {a b} {A : Set a} {B : A → Set b} {x} → B x → ∃ B
, y = _ , y
-proj₁ : ∀ {a b} {A : Set a} {B : A → Set b} → Σ A B → A
-proj₁ (x , y) = x
-
-proj₂ : ∀ {a b} {A : Set a} {B : A → Set b} →
- (p : Σ A B) → B (proj₁ p)
-proj₂ (x , y) = y
-
<_,_> : ∀ {a b c} {A : Set a} {B : A → Set b} {C : ∀ {x} → B x → Set c}
(f : (x : A) → B x) → ((x : A) → C (f x)) →
((x : A) → Σ (B x) C)
@@ -69,7 +69,7 @@ map : ∀ {a b p q}
{A : Set a} {B : Set b} {P : A → Set p} {Q : B → Set q} →
(f : A → B) → (∀ {x} → P x → Q (f x)) →
Σ A P → Σ B Q
-map f g = < f ∘ proj₁ , g ∘ proj₂ >
+map f g (x , y) = (f x , g y)
zip : ∀ {a b c p q r}
{A : Set a} {B : Set b} {C : Set c}
@@ -77,10 +77,10 @@ zip : ∀ {a b c p q r}
(_∙_ : A → B → C) →
(∀ {x y} → P x → Q y → R (x ∙ y)) →
Σ A P → Σ B Q → Σ C R
-zip _∙_ _∘_ p₁ p₂ = (proj₁ p₁ ∙ proj₁ p₂ , proj₂ p₁ ∘ proj₂ p₂)
+zip _∙_ _∘_ (a , p) (b , q) = (a ∙ b , p ∘ q)
swap : ∀ {a b} {A : Set a} {B : Set b} → A × B → B × A
-swap = < proj₂ , proj₁ >
+swap (x , y) = (y , x)
_-×-_ : ∀ {a b i j} {A : Set a} {B : Set b} →
(A → B → Set i) → (A → B → Set j) → (A → B → Set _)
@@ -98,4 +98,8 @@ curry f x y = f (x , y)
uncurry : ∀ {a b c} {A : Set a} {B : A → Set b} {C : Σ A B → Set c} →
((x : A) → (y : B x) → C (x , y)) →
((p : Σ A B) → C p)
-uncurry f (p₁ , p₂) = f p₁ p₂
+uncurry f (x , y) = f x y
+
+uncurry′ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} →
+ (A → B → C) → (A × B → C)
+uncurry′ = uncurry
diff --git a/src/Data/Product/Record.agda b/src/Data/Product/Record.agda
deleted file mode 100644
index b510cbc..0000000
--- a/src/Data/Product/Record.agda
+++ /dev/null
@@ -1,65 +0,0 @@
-------------------------------------------------------------------------
--- Products implemented using records
-------------------------------------------------------------------------
-
-{-# OPTIONS --universe-polymorphism #-}
-
--- It it ever becomes convenient to pattern match on records I might
--- make this the default implementation of products.
-
-module Data.Product.Record where
-
-open import Data.Function
-open import Level
-
-infixr 4 _,_
-infixr 2 _×_ _-×-_ _-,-_
-
-------------------------------------------------------------------------
--- Definition
-
-record Σ {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where
- constructor _,_
- field
- proj₁ : A
- proj₂ : B proj₁
-
-open Σ public
-
-_×_ : ∀ {a b} (A : Set a) (B : Set b) → Set (a ⊔ b)
-A × B = Σ A (λ _ → B)
-
-------------------------------------------------------------------------
--- Functions
-
-<_,_> : ∀ {a b c} {A : Set a} {B : A → Set b} {C : ∀ {x} → B x → Set c}
- (f : (x : A) → B x) → ((x : A) → C (f x)) →
- ((x : A) → Σ (B x) C)
-< f , g > x = (f x , g x)
-
-map : ∀ {a b p q}
- {A : Set a} {B : Set b} {P : A → Set p} {Q : B → Set q} →
- (f : A → B) → (∀ {x} → P x → Q (f x)) →
- Σ A P → Σ B Q
-map f g = < f ∘ proj₁ , g ∘ proj₂ >
-
-swap : ∀ {a b} {A : Set a} {B : Set b} → A × B → B × A
-swap = < proj₂ , proj₁ >
-
-_-×-_ : ∀ {a b i j} {A : Set a} {B : Set b} →
- (A → B → Set i) → (A → B → Set j) → (A → B → Set _)
-f -×- g = f -[ _×_ ]- g
-
-_-,-_ : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
- (A → B → C) → (A → B → D) → (A → B → C × D)
-f -,- g = f -[ _,_ ]- g
-
-curry : ∀ {a b c} {A : Set a} {B : A → Set b} {C : Σ A B → Set c} →
- ((p : Σ A B) → C p) →
- ((x : A) → (y : B x) → C (x , y))
-curry f x y = f (x , y)
-
-uncurry : ∀ {a b c} {A : Set a} {B : A → Set b} {C : Σ A B → Set c} →
- ((x : A) → (y : B x) → C (x , y)) →
- ((p : Σ A B) → C p)
-uncurry f p = f (proj₁ p) (proj₂ p)
diff --git a/src/Data/Rational.agda b/src/Data/Rational.agda
index 7d463d3..0f11fb3 100644
--- a/src/Data/Rational.agda
+++ b/src/Data/Rational.agda
@@ -5,7 +5,7 @@
module Data.Rational where
import Data.Bool.Properties as Bool
-open import Data.Function
+open import Function
open import Data.Integer hiding (suc) renaming (_*_ to _ℤ*_)
open import Data.Integer.Divisibility as ℤDiv using (Coprime)
import Data.Integer.Properties as ℤ
diff --git a/src/Data/Sets.agda b/src/Data/Sets.agda
deleted file mode 100644
index b473350..0000000
--- a/src/Data/Sets.agda
+++ /dev/null
@@ -1,95 +0,0 @@
-------------------------------------------------------------------------
--- Finite sets (currently only some type signatures)
-------------------------------------------------------------------------
-
-module Data.Sets where
-
-open import Relation.Nullary
-open import Relation.Binary
-open import Relation.Binary.OrderMorphism
-open import Data.Function
-open import Data.List as L using (List)
-open import Data.Product using (∃; _×_)
-open import Level
-
-module Sets₁ (dto : DecTotalOrder zero zero zero) where
-
- open DecTotalOrder dto public using (_≈_)
- open DecTotalOrder dto hiding (_≈_)
-
- infixr 6 _∪_
- infix 5 _∈?_
- infix 4 _∈_ _|≈|_
-
- abstract postulate decSetoid : DecSetoid _ _
-
- <Set> : Set
- <Set> = DecSetoid.Carrier decSetoid
-
- _|≈|_ : Rel <Set> zero
- _|≈|_ = DecSetoid._≈_ decSetoid
-
- abstract
- postulate
- empty : <Set>
- insert : Carrier → <Set> → <Set>
- _∪_ : <Set> → <Set> → <Set>
- _∈_ : Carrier → <Set> → Set
- _∈?_ : (x : Carrier) → (s : <Set>) → Dec (x ∈ s)
- toList : <Set> → List Carrier
-
- postulate
- prop-∈-insert₁ : ∀ {x y s} → x ≈ y → x ∈ insert y s
- prop-∈-insert₂ : ∀ {x y s} → x ∈ s → x ∈ insert y s
- prop-∈-insert₃ : ∀ {x y s} → ¬ x ≈ y → x ∈ insert y s → x ∈ s
-
- prop-∈-empty : ∀ {x} → ¬ x ∈ empty
-
- prop-∈-∪ : ∀ {x s₁ s₂} → x ∈ s₁ → x ∈ s₁ ∪ s₂
-
- prop-∪₁ : ∀ {s₁ s₂} → s₁ ∪ s₂ |≈| s₂ ∪ s₁
- prop-∪₂ : ∀ {s₁ s₂ s₃} → s₁ ∪ (s₂ ∪ s₃) |≈| (s₁ ∪ s₂) ∪ s₃
-
- prop-∈-|≈| : ∀ {x} → (λ s → x ∈ s) Respects _|≈|_
- prop-∈-≈ : ∀ {s} → (λ x → x ∈ s) Respects _≈_
-
- -- TODO: Postulates for toList.
-
- singleton : Carrier → <Set>
- singleton x = insert x empty
-
- ⋃_ : List <Set> → <Set>
- ⋃_ = L.foldr _∪_ empty
-
- fromList : List Carrier → <Set>
- fromList = L.foldr insert empty
-
- _⊆_ : <Set> → <Set> → Set
- s₁ ⊆ s₂ = ∀ x → x ∈ s₁ → x ∈ s₂
-
-open Sets₁ public
-open DecTotalOrder hiding (_≈_)
-open _⇒-Poset_
-
-abstract
- postulate
- map : ∀ {do₁ do₂} → do₁ ⇒-DTO do₂ → <Set> do₁ → <Set> do₂
- mapToSet : ∀ {do₁ do₂} →
- (Carrier do₁ → <Set> do₂) →
- <Set> do₁ → <Set> do₂
-
- prop-map-∈₁ : ∀ {do₁ do₂ f x s} →
- x ⟨ _∈_ do₁ ⟩ s →
- fun f x ⟨ _∈_ do₂ ⟩ map f s
- prop-map-∈₂ : ∀ {do₁ do₂ f y s} →
- y ⟨ _∈_ do₂ ⟩ map f s →
- ∃ λ x → (fun f x ⟨ _≈_ do₂ ⟩ y) ×
- ( x ⟨ _∈_ do₁ ⟩ s)
-
- prop-mapToSet₁ : ∀ {do₁ do₂ f x s} →
- x ⟨ _∈_ do₁ ⟩ s →
- f x ⟨ _⊆_ do₂ ⟩ mapToSet f s
- prop-mapToSet₂ : ∀ {do₁ do₂ f y s} →
- y ⟨ _∈_ do₂ ⟩ mapToSet f s →
- ∃ λ x → (y ⟨ _∈_ do₂ ⟩ f x) ×
- (x ⟨ _∈_ do₁ ⟩ s)
diff --git a/src/Data/Sign/Properties.agda b/src/Data/Sign/Properties.agda
index b05902f..b143bc3 100644
--- a/src/Data/Sign/Properties.agda
+++ b/src/Data/Sign/Properties.agda
@@ -5,7 +5,7 @@
module Data.Sign.Properties where
open import Data.Empty
-open import Data.Function
+open import Function
open import Data.Sign
open import Relation.Binary.PropositionalEquality
diff --git a/src/Data/Star.agda b/src/Data/Star.agda
index c18b7e5..f2d716a 100644
--- a/src/Data/Star.agda
+++ b/src/Data/Star.agda
@@ -9,7 +9,7 @@
module Data.Star where
open import Relation.Binary
-open import Data.Function
+open import Function
open import Level
infixr 5 _◅_
@@ -60,11 +60,6 @@ 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 = gmap id
--- TransFlip is used to state the type signature of gfold.
-
-TransFlip : {A : Set} → Rel A zero → Rel A zero → Rel A zero → Set
-TransFlip P Q R = ∀ {i j k} → Q j k → P i j → R i k
-
-- A generalised variant of fold.
gfold : ∀ {I J T} (f : I → J) P →
diff --git a/src/Data/Star/BoundedVec.agda b/src/Data/Star/BoundedVec.agda
index 7d99a29..454007f 100644
--- a/src/Data/Star/BoundedVec.agda
+++ b/src/Data/Star/BoundedVec.agda
@@ -12,7 +12,7 @@ open import Data.Star.Decoration
open import Data.Star.Pointer
open import Data.Star.List using (List)
open import Data.Unit
-open import Data.Function
+open import Function
import Data.Maybe as Maybe
open import Relation.Binary
open import Relation.Binary.Consequences
diff --git a/src/Data/Star/Decoration.agda b/src/Data/Star/Decoration.agda
index 3b4da23..8a831da 100644
--- a/src/Data/Star/Decoration.agda
+++ b/src/Data/Star/Decoration.agda
@@ -6,7 +6,7 @@ module Data.Star.Decoration where
open import Data.Star
open import Relation.Binary
-open import Data.Function
+open import Function
open import Data.Unit
open import Level
diff --git a/src/Data/Star/Nat.agda b/src/Data/Star/Nat.agda
index 2830871..57f0b0d 100644
--- a/src/Data/Star/Nat.agda
+++ b/src/Data/Star/Nat.agda
@@ -8,7 +8,7 @@ open import Data.Star
open import Data.Unit
open import Relation.Binary
open import Relation.Binary.Simple
-open import Data.Function
+open import Function
import Level as L
-- Natural numbers.
diff --git a/src/Data/Star/Pointer.agda b/src/Data/Star/Pointer.agda
index 0f675e3..ccddcca 100644
--- a/src/Data/Star/Pointer.agda
+++ b/src/Data/Star/Pointer.agda
@@ -9,7 +9,7 @@ open import Data.Star.Decoration
open import Relation.Binary
open import Data.Maybe using (Maybe; nothing; just)
open import Data.Unit
-open import Data.Function
+open import Function
open import Level
-- Pointers into star-lists. The edge pointed to is decorated with Q,
diff --git a/src/Data/Star/Properties.agda b/src/Data/Star/Properties.agda
index 3a793de..52465ba 100644
--- a/src/Data/Star/Properties.agda
+++ b/src/Data/Star/Properties.agda
@@ -5,7 +5,7 @@
module Data.Star.Properties where
open import Data.Star
-open import Data.Function
+open import Function
open import Level
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as PropEq
@@ -73,7 +73,6 @@ preorder T = record
{ isEquivalence = PropEq.isEquivalence
; reflexive = reflexive
; trans = _◅◅_
- ; ∼-resp-≈ = PropEq.resp₂ (Star T)
}
}
where
diff --git a/src/Data/Star/Vec.agda b/src/Data/Star/Vec.agda
index ffeb200..7f03db9 100644
--- a/src/Data/Star/Vec.agda
+++ b/src/Data/Star/Vec.agda
@@ -12,7 +12,7 @@ open import Data.Star.Pointer as Pointer hiding (lookup)
open import Data.Star.List using (List)
open import Relation.Binary
open import Relation.Binary.Consequences
-open import Data.Function
+open import Function
open import Data.Unit
-- The vector type. Vectors are natural numbers decorated with extra
diff --git a/src/Data/Stream.agda b/src/Data/Stream.agda
index f675c23..ea6a270 100644
--- a/src/Data/Stream.agda
+++ b/src/Data/Stream.agda
@@ -21,28 +21,28 @@ data Stream (A : Set) : Set where
------------------------------------------------------------------------
-- Some operations
-head : forall {A} -> Stream A -> A
+head : forall {A} → Stream A → A
head (x ∷ xs) = x
-tail : forall {A} -> Stream A -> Stream A
+tail : forall {A} → Stream A → Stream A
tail (x ∷ xs) = ♭ xs
map : ∀ {A B} → (A → B) → Stream A → Stream B
map f (x ∷ xs) = f x ∷ ♯ map f (♭ xs)
-zipWith : forall {A B C} ->
- (A -> B -> C) -> Stream A -> Stream B -> Stream C
+zipWith : forall {A B C} →
+ (A → B → C) → Stream A → Stream B → Stream C
zipWith _∙_ (x ∷ xs) (y ∷ ys) = (x ∙ y) ∷ ♯ zipWith _∙_ (♭ xs) (♭ ys)
take : ∀ {A} (n : ℕ) → Stream A → Vec A n
take zero xs = []
take (suc n) (x ∷ xs) = x ∷ take n (♭ xs)
-drop : ∀ {A} -> ℕ -> Stream A -> Stream A
+drop : ∀ {A} → ℕ → Stream A → Stream A
drop zero xs = xs
drop (suc n) (x ∷ xs) = drop n (♭ xs)
-repeat : forall {A} -> A -> Stream A
+repeat : forall {A} → A → Stream A
repeat x = x ∷ ♯ repeat x
iterate : ∀ {A} → (A → A) → A → Stream A
@@ -55,6 +55,20 @@ infixr 5 _⋎_
_⋎_ : ∀ {A} → Stream A → Stream A → Stream A
(x ∷ xs) ⋎ ys = x ∷ ♯ (ys ⋎ ♭ xs)
+mutual
+
+ -- Takes every other element from the stream, starting with the
+ -- first one.
+
+ evens : {A : Set} → Stream A → Stream A
+ evens (x ∷ xs) = x ∷ ♯ odds (♭ xs)
+
+ -- Takes every other element from the stream, starting with the
+ -- second one.
+
+ odds : {A : Set} → Stream A → Stream A
+ odds (x ∷ xs) = evens (♭ xs)
+
toColist : ∀ {A} → Stream A → Colist A
toColist (x ∷ xs) = x ∷ ♯ toColist (♭ xs)
@@ -120,3 +134,15 @@ setoid A = record
map-cong : ∀ {A B} (f : A → B) {xs ys : Stream A} →
xs ≈ ys → map f xs ≈ map f ys
map-cong f (x ∷ xs≈) = f x ∷ ♯ map-cong f (♭ xs≈)
+
+zipWith-cong : ∀ {A B C} (_∙_ : A → B → C) {xs xs′ ys ys′} →
+ xs ≈ xs′ → ys ≈ ys′ →
+ zipWith _∙_ xs ys ≈ zipWith _∙_ xs′ ys′
+zipWith-cong _∙_ (x ∷ xs≈) (y ∷ ys≈) =
+ (x ∙ y) ∷ ♯ zipWith-cong _∙_ (♭ xs≈) (♭ ys≈)
+
+infixr 5 _⋎-cong_
+
+_⋎-cong_ : ∀ {A} {xs xs′ ys ys′ : Stream A} →
+ xs ≈ xs′ → ys ≈ ys′ → xs ⋎ ys ≈ xs′ ⋎ ys′
+(x ∷ xs≈) ⋎-cong ys≈ = x ∷ ♯ (ys≈ ⋎-cong ♭ xs≈)
diff --git a/src/Data/String.agda b/src/Data/String.agda
index 7acaae6..41ba5e4 100644
--- a/src/Data/String.agda
+++ b/src/Data/String.agda
@@ -4,12 +4,12 @@
module Data.String where
-open import Data.List as List using (List)
+open import Data.List as List using (_∷_; []; List)
open import Data.Vec as Vec using (Vec)
open import Data.Colist as Colist using (Colist)
open import Data.Char using (Char)
open import Data.Bool using (Bool; true; false)
-open import Data.Function
+open import Function
open import Relation.Nullary
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
@@ -58,6 +58,10 @@ toVec s = Vec.fromList (toList s)
toCostring : String → Costring
toCostring = Colist.fromList ∘ toList
+unlines : List String → String
+unlines [] = ""
+unlines (x ∷ xs) = x ++ "\n" ++ unlines xs
+
infix 4 _==_
_==_ : String → String → Bool
diff --git a/src/Data/Sum.agda b/src/Data/Sum.agda
index f931649..6c54002 100644
--- a/src/Data/Sum.agda
+++ b/src/Data/Sum.agda
@@ -6,7 +6,7 @@
module Data.Sum where
-open import Data.Function
+open import Function
open import Data.Maybe.Core
open import Level
diff --git a/src/Data/Unit.agda b/src/Data/Unit.agda
index 0db2f9f..92fc219 100644
--- a/src/Data/Unit.agda
+++ b/src/Data/Unit.agda
@@ -53,7 +53,6 @@ decTotalOrder = record
{ isEquivalence = PropEq.isEquivalence
; reflexive = λ _ → _
; trans = λ _ _ → _
- ; ∼-resp-≈ = PropEq.resp₂ _≤_
}
; antisym = antisym
}
diff --git a/src/Data/Vec/Equality.agda b/src/Data/Vec/Equality.agda
index 373990c..0892021 100644
--- a/src/Data/Vec/Equality.agda
+++ b/src/Data/Vec/Equality.agda
@@ -8,7 +8,7 @@ module Data.Vec.Equality where
open import Data.Vec
open import Data.Nat using (suc)
-open import Data.Function
+open import Function
open import Level
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
@@ -93,4 +93,7 @@ module HeterogeneousEquality {a} {A : Set a} where
from-≅ : ∀ {n m} {xs : Vec A n} {ys : Vec A m} →
xs ≅ ys → xs ≈ ys
- from-≅ HetEq.refl = refl _
+ from-≅ {xs = []} {ys = []} _ = refl _
+ from-≅ {xs = x ∷ xs} {ys = .x ∷ .xs} HetEq.refl = refl _
+ from-≅ {xs = _ ∷ _} {ys = []} ()
+ from-≅ {xs = []} {ys = _ ∷ _} ()
diff --git a/src/Data/Vec/N-ary.agda b/src/Data/Vec/N-ary.agda
index e465f20..bc5512d 100644
--- a/src/Data/Vec/N-ary.agda
+++ b/src/Data/Vec/N-ary.agda
@@ -6,7 +6,7 @@
module Data.Vec.N-ary where
-open import Data.Function
+open import Function
open import Data.Nat
open import Data.Vec
open import Level
diff --git a/src/Data/Vec/Properties.agda b/src/Data/Vec/Properties.agda
index cf1a71a..02fabc6 100644
--- a/src/Data/Vec/Properties.agda
+++ b/src/Data/Vec/Properties.agda
@@ -11,7 +11,7 @@ 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.Function
+open import Function
open import Level
open import Relation.Binary
diff --git a/src/Foreign/Haskell.agda b/src/Foreign/Haskell.agda
index 3db7d11..d3ebae4 100644
--- a/src/Foreign/Haskell.agda
+++ b/src/Foreign/Haskell.agda
@@ -7,7 +7,7 @@ module Foreign.Haskell where
open import Coinduction
open import Data.Char using (Char)
open import Data.Colist as C using ([]; _∷_)
-open import Data.Function using (_∘_)
+open import Function using (_∘_)
open import Data.String as String using (String)
------------------------------------------------------------------------
@@ -24,19 +24,19 @@ data Unit : Set where
infixr 5 _∷_
-codata Colist (A : Set) : Set where
+data Colist (A : Set) : Set where
[] : Colist A
- _∷_ : (x : A) (xs : Colist A) → Colist A
+ _∷_ : (x : A) (xs : ∞ (Colist A)) → Colist A
{-# COMPILED_DATA Colist [] [] (:) #-}
fromColist : ∀ {A} → C.Colist A → Colist A
fromColist [] = []
-fromColist (x ∷ xs) = x ∷ fromColist (♭ xs)
+fromColist (x ∷ xs) = x ∷ ♯ fromColist (♭ xs)
toColist : ∀ {A} → Colist A → C.Colist A
toColist [] = []
-toColist (x ∷ xs) = x ∷ ♯ toColist xs
+toColist (x ∷ xs) = x ∷ ♯ toColist (♭ xs)
fromString : String → Colist Char
fromString = fromColist ∘ String.toCostring
diff --git a/src/Data/Function.agda b/src/Function.agda
index 100e7fe..0f1187e 100644
--- a/src/Data/Function.agda
+++ b/src/Function.agda
@@ -4,7 +4,7 @@
{-# OPTIONS --universe-polymorphism #-}
-module Data.Function where
+module Function where
open import Level
@@ -12,7 +12,7 @@ infixr 9 _∘_ _∘′_
infixl 1 _on_
infixl 1 _⟨_⟩_
infixr 0 _-[_]-_ _$_
-infix 0 _∶_
+infix 0 type-signature
------------------------------------------------------------------------
-- Types
@@ -74,9 +74,8 @@ f -[ _*_ ]- g = λ x y → f x y * g x y
-- In Agda you cannot annotate every subexpression with a type
-- signature. This function can be used instead.
---
--- You should think of the colon as being mirrored around its vertical
--- axis; the type comes first.
-_∶_ : ∀ {a} (A : Set a) → A → A
-_ ∶ x = x
+type-signature : ∀ {a} (A : Set a) → A → A
+type-signature A x = x
+
+syntax type-signature A x = x ∶ A
diff --git a/src/Function/Bijection.agda b/src/Function/Bijection.agda
new file mode 100644
index 0000000..ad66dc9
--- /dev/null
+++ b/src/Function/Bijection.agda
@@ -0,0 +1,88 @@
+------------------------------------------------------------------------
+-- Bijections
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Function.Bijection where
+
+open import Data.Product
+open import Level
+open import Relation.Binary
+open import Function.Equality as F
+ using (_⟶_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_)
+open import Function.Injection as Inj hiding (id; _∘_)
+open import Function.Surjection as Surj hiding (id; _∘_)
+open import Function.LeftInverse as Left hiding (id; _∘_)
+
+-- Bijective functions.
+
+record Bijective {f₁ f₂ t₁ t₂}
+ {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
+ (to : From ⟶ To) :
+ Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
+ field
+ injective : Injective to
+ surjective : Surjective to
+
+ open Surjective surjective public
+
+ left-inverse-of : from LeftInverseOf to
+ left-inverse-of x = injective (right-inverse-of (to ⟨$⟩ x))
+
+-- The set of all bijections between two setoids.
+
+record Bijection {f₁ f₂ t₁ t₂}
+ (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) :
+ Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
+ field
+ to : From ⟶ To
+ bijective : Bijective to
+
+ open Bijective bijective public
+
+ injection : Injection From To
+ injection = record
+ { to = to
+ ; injective = injective
+ }
+
+ surjection : Surjection From To
+ surjection = record
+ { to = to
+ ; surjective = surjective
+ }
+
+ open Surjection surjection public using (equivalent; right-inverse)
+
+ left-inverse : LeftInverse From To
+ left-inverse = record
+ { to = to
+ ; from = from
+ ; left-inverse-of = left-inverse-of
+ }
+
+-- Identity and composition. (Note that these proofs are superfluous,
+-- given that Bijection is equivalent to Function.Inverse.Inverse.)
+
+id : ∀ {s₁ s₂} {S : Setoid s₁ s₂} → Bijection S S
+id {S = S} = record
+ { to = F.id
+ ; bijective = record
+ { injective = Injection.injective (Inj.id {S = S})
+ ; surjective = Surjection.surjective (Surj.id {S = S})
+ }
+ }
+
+infixr 9 _∘_
+
+_∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂}
+ {F : Setoid f₁ f₂} {M : Setoid m₁ m₂} {T : Setoid t₁ t₂} →
+ Bijection M T → Bijection F M → Bijection F T
+f ∘ g = record
+ { to = to f ⟪∘⟫ to g
+ ; bijective = record
+ { injective = Injection.injective (Inj._∘_ (injection f) (injection g))
+ ; surjective = Surjection.surjective (Surj._∘_ (surjection f) (surjection g))
+ }
+ } where open Bijection
diff --git a/src/Function/Equality.agda b/src/Function/Equality.agda
new file mode 100644
index 0000000..23011ed
--- /dev/null
+++ b/src/Function/Equality.agda
@@ -0,0 +1,101 @@
+------------------------------------------------------------------------
+-- Function setoids and related constructions
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Function.Equality where
+
+open import Function as Fun using (_on_)
+open import Level
+import Relation.Binary as B
+import Relation.Binary.Indexed as I
+
+------------------------------------------------------------------------
+-- Functions which preserve equality
+
+record Π {f₁ f₂ t₁ t₂}
+ (From : B.Setoid f₁ f₂)
+ (To : I.Setoid (B.Setoid.Carrier From) t₁ t₂) :
+ Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
+ open I using (_=[_]⇒_)
+ infixl 5 _⟨$⟩_
+ field
+ _⟨$⟩_ : (x : B.Setoid.Carrier From) → I.Setoid.Carrier To x
+ cong : B.Setoid._≈_ From =[ _⟨$⟩_ ]⇒ I.Setoid._≈_ To
+
+open Π public
+
+infixr 0 _⟶_
+
+_⟶_ : ∀ {f₁ f₂ t₁ t₂} → B.Setoid f₁ f₂ → B.Setoid t₁ t₂ → Set _
+From ⟶ To = Π From (B.Setoid.indexedSetoid To)
+
+-- Identity and composition.
+
+id : ∀ {a₁ a₂} {A : B.Setoid a₁ a₂} → A ⟶ A
+id = record { _⟨$⟩_ = Fun.id; cong = Fun.id }
+
+infixr 9 _∘_
+
+_∘_ : ∀ {a₁ a₂} {A : B.Setoid a₁ a₂}
+ {b₁ b₂} {B : B.Setoid b₁ b₂}
+ {c₁ c₂} {C : B.Setoid c₁ c₂} →
+ B ⟶ C → A ⟶ B → A ⟶ C
+f ∘ g = record
+ { _⟨$⟩_ = Fun._∘_ (_⟨$⟩_ f) (_⟨$⟩_ g)
+ ; cong = Fun._∘_ (cong f) (cong g)
+ }
+
+-- Constant equality-preserving function.
+
+const : ∀ {a₁ a₂} {A : B.Setoid a₁ a₂}
+ {b₁ b₂} {B : B.Setoid b₁ b₂} →
+ B.Setoid.Carrier B → A ⟶ B
+const {B = B} b = record
+ { _⟨$⟩_ = Fun.const b
+ ; cong = Fun.const (B.Setoid.refl B)
+ }
+
+------------------------------------------------------------------------
+-- Function setoids
+
+-- Dependent.
+
+setoid : ∀ {f₁ f₂ t₁ t₂}
+ (From : B.Setoid f₁ f₂) →
+ I.Setoid (B.Setoid.Carrier From) t₁ t₂ →
+ B.Setoid _ _
+setoid From To = record
+ { Carrier = Π From To
+ ; _≈_ = λ f g → ∀ {x y} → x ≈₁ y → f ⟨$⟩ x ≈₂ g ⟨$⟩ y
+ ; isEquivalence = record
+ { refl = λ {f} → cong f
+ ; sym = λ f∼g x∼y → To.sym (f∼g (From.sym x∼y))
+ ; trans = λ f∼g g∼h x∼y → To.trans (f∼g From.refl) (g∼h x∼y)
+ }
+ }
+ where
+ open module From = B.Setoid From using () renaming (_≈_ to _≈₁_)
+ open module To = I.Setoid To using () renaming (_≈_ to _≈₂_)
+
+-- Non-dependent.
+
+infixr 0 _⇨_
+
+_⇨_ : ∀ {f₁ f₂ t₁ t₂} → B.Setoid f₁ f₂ → B.Setoid t₁ t₂ → B.Setoid _ _
+From ⇨ To = setoid From (B.Setoid.indexedSetoid To)
+
+-- A variant of setoid which uses the propositional equality setoid
+-- for the domain, and a more convenient definition of _≈_.
+
+≡-setoid : ∀ {f t₁ t₂} (From : Set f) → I.Setoid From t₁ t₂ → B.Setoid _ _
+≡-setoid From To = record
+ { Carrier = (x : From) → Carrier x
+ ; _≈_ = λ f g → ∀ x → f x ≈ g x
+ ; isEquivalence = record
+ { refl = λ {f} x → refl
+ ; sym = λ f∼g x → sym (f∼g x)
+ ; trans = λ f∼g g∼h x → trans (f∼g x) (g∼h x)
+ }
+ } where open I.Setoid To
diff --git a/src/Function/Equivalence.agda b/src/Function/Equivalence.agda
new file mode 100644
index 0000000..39f563c
--- /dev/null
+++ b/src/Function/Equivalence.agda
@@ -0,0 +1,105 @@
+------------------------------------------------------------------------
+-- Equivalence (coinhabitance)
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Function.Equivalence where
+
+open import Function using (flip)
+open import Function.Equality as F
+ using (_⟶_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_)
+open import Level
+open import Relation.Binary
+import Relation.Binary.PropositionalEquality as P
+
+-- Setoid equivalence.
+
+record Equivalent {f₁ f₂ t₁ t₂}
+ (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) :
+ Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
+ field
+ to : From ⟶ To
+ from : To ⟶ From
+
+-- Set equivalence.
+
+infix 3 _⇔_
+
+_⇔_ : ∀ {f t} → Set f → Set t → Set _
+From ⇔ To = Equivalent (P.setoid From) (P.setoid To)
+
+equivalent : ∀ {f t} {From : Set f} {To : Set t} →
+ (From → To) → (To → From) → From ⇔ To
+equivalent to from = record { to = P.→-to-⟶ to; from = P.→-to-⟶ from }
+
+------------------------------------------------------------------------
+-- Map and zip
+
+map : ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
+ {f₁′ f₂′ t₁′ t₂′}
+ {From′ : Setoid f₁′ f₂′} {To′ : Setoid t₁′ t₂′} →
+ ((From ⟶ To) → (From′ ⟶ To′)) →
+ ((To ⟶ From) → (To′ ⟶ From′)) →
+ Equivalent From To → Equivalent From′ To′
+map t f eq = record { to = t to; from = f from }
+ where open Equivalent eq
+
+zip : ∀ {f₁₁ f₂₁ t₁₁ t₂₁}
+ {From₁ : Setoid f₁₁ f₂₁} {To₁ : Setoid t₁₁ t₂₁}
+ {f₁₂ f₂₂ t₁₂ t₂₂}
+ {From₂ : Setoid f₁₂ f₂₂} {To₂ : Setoid t₁₂ t₂₂}
+ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} →
+ ((From₁ ⟶ To₁) → (From₂ ⟶ To₂) → (From ⟶ To)) →
+ ((To₁ ⟶ From₁) → (To₂ ⟶ From₂) → (To ⟶ From)) →
+ Equivalent From₁ To₁ → Equivalent From₂ To₂ → Equivalent From To
+zip t f eq₁ eq₂ =
+ record { to = t (to eq₁) (to eq₂); from = f (from eq₁) (from eq₂) }
+ where open Equivalent
+
+------------------------------------------------------------------------
+-- Equivalent is an equivalence relation
+
+-- Identity and composition (reflexivity and transitivity).
+
+id : ∀ {s₁ s₂} → Reflexive (Equivalent {s₁} {s₂})
+id {x = S} = record
+ { to = F.id
+ ; from = F.id
+ }
+
+infixr 9 _∘_
+
+_∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂} →
+ TransFlip (Equivalent {f₁} {f₂} {m₁} {m₂})
+ (Equivalent {m₁} {m₂} {t₁} {t₂})
+ (Equivalent {f₁} {f₂} {t₁} {t₂})
+f ∘ g = record
+ { to = to f ⟪∘⟫ to g
+ ; from = from g ⟪∘⟫ from f
+ } where open Equivalent
+
+-- Symmetry.
+
+sym : ∀ {f₁ f₂ t₁ t₂} →
+ Sym (Equivalent {f₁} {f₂} {t₁} {t₂}) (Equivalent {t₁} {t₂} {f₁} {f₂})
+sym eq = record
+ { from = to
+ ; to = from
+ } where open Equivalent eq
+
+-- For fixed universe levels we can construct setoids.
+
+setoid : (s₁ s₂ : Level) → Setoid (suc (s₁ ⊔ s₂)) (s₁ ⊔ s₂)
+setoid s₁ s₂ = record
+ { Carrier = Setoid s₁ s₂
+ ; _≈_ = Equivalent
+ ; isEquivalence = record {refl = id; sym = sym; trans = flip _∘_}
+ }
+
+⇔-setoid : (ℓ : Level) → Setoid (suc ℓ) ℓ
+⇔-setoid ℓ = record
+ { Carrier = Set ℓ
+ ; _≈_ = _⇔_
+ ; isEquivalence = record {refl = id; sym = sym; trans = flip _∘_}
+ }
diff --git a/src/Data/Function/Injection.agda b/src/Function/Injection.agda
index ddbae58..dc2e4b2 100644
--- a/src/Data/Function/Injection.agda
+++ b/src/Function/Injection.agda
@@ -4,14 +4,16 @@
{-# OPTIONS --universe-polymorphism #-}
-module Data.Function.Injection where
+module Function.Injection where
-open import Data.Function as Fun using () renaming (_∘_ to _⟨∘⟩_)
+open import Function as Fun using () renaming (_∘_ to _⟨∘⟩_)
open import Level
open import Relation.Binary
-open import Data.Function.Equality as F
+open import Function.Equality as F
using (_⟶_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_)
+-- Injective functions.
+
Injective : ∀ {a₁ a₂ b₁ b₂} {A : Setoid a₁ a₂} {B : Setoid b₁ b₂} →
A ⟶ B → Set _
Injective {A = A} {B} f = ∀ {x y} → f ⟨$⟩ x ≈₂ f ⟨$⟩ y → x ≈₁ y
@@ -19,6 +21,8 @@ Injective {A = A} {B} f = ∀ {x y} → f ⟨$⟩ x ≈₂ f ⟨$⟩ y → x ≈
open Setoid A renaming (_≈_ to _≈₁_)
open Setoid B renaming (_≈_ to _≈₂_)
+-- The set of all injections between two setoids.
+
record Injection {f₁ f₂ t₁ t₂}
(From : Setoid f₁ f₂) (To : Setoid t₁ t₂) :
Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
@@ -26,6 +30,8 @@ record Injection {f₁ f₂ t₁ t₂}
to : From ⟶ To
injective : Injective to
+-- Identity and composition.
+
infixr 9 _∘_
id : ∀ {s₁ s₂} {S : Setoid s₁ s₂} → Injection S S
diff --git a/src/Function/Inverse.agda b/src/Function/Inverse.agda
new file mode 100644
index 0000000..d8a8e76
--- /dev/null
+++ b/src/Function/Inverse.agda
@@ -0,0 +1,273 @@
+------------------------------------------------------------------------
+-- Inverses
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Function.Inverse where
+
+open import Level
+open import Function as Fun using (flip)
+open import Function.Bijection hiding (id; _∘_)
+open import Function.Equality as F
+ using (_⟶_) renaming (_∘_ to _⟪∘⟫_)
+open import Function.Equivalence as Eq using (Equivalent; ⇔-setoid)
+open import Function.LeftInverse as Left hiding (id; _∘_)
+open import Function.Surjection as Surj hiding (id; _∘_)
+open import Relation.Binary
+import Relation.Binary.PropositionalEquality as P
+
+-- Inverses.
+
+record _InverseOf_ {f₁ f₂ t₁ t₂}
+ {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
+ (from : To ⟶ From) (to : From ⟶ To) :
+ Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
+ field
+ left-inverse-of : from LeftInverseOf to
+ right-inverse-of : from RightInverseOf to
+
+-- The set of all inverses between two setoids.
+
+record Inverse {f₁ f₂ t₁ t₂}
+ (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) :
+ Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
+ field
+ to : From ⟶ To
+ from : To ⟶ From
+ inverse-of : from InverseOf to
+
+ open _InverseOf_ inverse-of public
+
+ left-inverse : LeftInverse From To
+ left-inverse = record
+ { to = to
+ ; from = from
+ ; left-inverse-of = left-inverse-of
+ }
+
+ open LeftInverse left-inverse public
+ using (injective; injection)
+
+ bijection : Bijection From To
+ bijection = record
+ { to = to
+ ; bijective = record
+ { injective = injective
+ ; surjective = record
+ { from = from
+ ; right-inverse-of = right-inverse-of
+ }
+ }
+ }
+
+ open Bijection bijection public
+ using (equivalent; surjective; surjection; right-inverse)
+
+-- The set of all inverses between two sets.
+
+infix 3 _⇿_
+
+_⇿_ : ∀ {f t} → Set f → Set t → Set _
+From ⇿ To = Inverse (P.setoid From) (P.setoid To)
+
+-- If two setoids are in bijective correspondence, then there is an
+-- inverse between them.
+
+fromBijection :
+ ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} →
+ Bijection From To → Inverse From To
+fromBijection b = record
+ { to = Bijection.to b
+ ; from = Bijection.from b
+ ; inverse-of = record
+ { left-inverse-of = Bijection.left-inverse-of b
+ ; right-inverse-of = Bijection.right-inverse-of b
+ }
+ }
+
+------------------------------------------------------------------------
+-- Map and zip
+
+map : ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
+ {f₁′ f₂′ t₁′ t₂′}
+ {From′ : Setoid f₁′ f₂′} {To′ : Setoid t₁′ t₂′} →
+ (t : (From ⟶ To) → (From′ ⟶ To′)) →
+ (f : (To ⟶ From) → (To′ ⟶ From′)) →
+ (∀ {to from} → from InverseOf to → f from InverseOf t to) →
+ Inverse From To → Inverse From′ To′
+map t f pres eq = record
+ { to = t to
+ ; from = f from
+ ; inverse-of = pres inverse-of
+ } where open Inverse eq
+
+zip : ∀ {f₁₁ f₂₁ t₁₁ t₂₁}
+ {From₁ : Setoid f₁₁ f₂₁} {To₁ : Setoid t₁₁ t₂₁}
+ {f₁₂ f₂₂ t₁₂ t₂₂}
+ {From₂ : Setoid f₁₂ f₂₂} {To₂ : Setoid t₁₂ t₂₂}
+ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} →
+ (t : (From₁ ⟶ To₁) → (From₂ ⟶ To₂) → (From ⟶ To)) →
+ (f : (To₁ ⟶ From₁) → (To₂ ⟶ From₂) → (To ⟶ From)) →
+ (∀ {to₁ from₁ to₂ from₂} →
+ from₁ InverseOf to₁ → from₂ InverseOf to₂ →
+ f from₁ from₂ InverseOf t to₁ to₂) →
+ Inverse From₁ To₁ → Inverse From₂ To₂ → Inverse From To
+zip t f pres eq₁ eq₂ = record
+ { to = t (to eq₁) (to eq₂)
+ ; from = f (from eq₁) (from eq₂)
+ ; inverse-of = pres (inverse-of eq₁) (inverse-of eq₂)
+ } where open Inverse
+
+------------------------------------------------------------------------
+-- Inverse is an equivalence relation
+
+-- Identity and composition (reflexivity and transitivity).
+
+id : ∀ {s₁ s₂} → Reflexive (Inverse {s₁} {s₂})
+id {x = S} = record
+ { to = F.id
+ ; from = F.id
+ ; inverse-of = record
+ { left-inverse-of = LeftInverse.left-inverse-of id′
+ ; right-inverse-of = LeftInverse.left-inverse-of id′
+ }
+ } where id′ = Left.id {S = S}
+
+infixr 9 _∘_
+
+_∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂} →
+ TransFlip (Inverse {f₁} {f₂} {m₁} {m₂})
+ (Inverse {m₁} {m₂} {t₁} {t₂})
+ (Inverse {f₁} {f₂} {t₁} {t₂})
+f ∘ g = record
+ { to = to f ⟪∘⟫ to g
+ ; from = from g ⟪∘⟫ from f
+ ; inverse-of = record
+ { left-inverse-of = LeftInverse.left-inverse-of (Left._∘_ (left-inverse f) (left-inverse g))
+ ; right-inverse-of = LeftInverse.left-inverse-of (Left._∘_ (right-inverse g) (right-inverse f))
+ }
+ } where open Inverse
+
+-- Symmetry.
+
+private
+ module Dummy where
+
+ sym : ∀ {f₁ f₂ t₁ t₂} →
+ Sym (Inverse {f₁} {f₂} {t₁} {t₂}) (Inverse {t₁} {t₂} {f₁} {f₂})
+ sym inv = record
+ { from = to
+ ; to = from
+ ; inverse-of = record
+ { left-inverse-of = right-inverse-of
+ ; right-inverse-of = left-inverse-of
+ }
+ } where open Inverse inv
+
+-- For fixed universe levels we can construct a setoid.
+
+setoid : (s₁ s₂ : Level) → Setoid (suc (s₁ ⊔ s₂)) (s₁ ⊔ s₂)
+setoid s₁ s₂ = record
+ { Carrier = Setoid s₁ s₂
+ ; _≈_ = Inverse
+ ; isEquivalence =
+ record {refl = id; sym = Dummy.sym; trans = flip _∘_}
+ }
+
+------------------------------------------------------------------------
+-- A generalisation which includes both _⇔_ and _⇿_
+
+-- There are two kinds of "isomorphisms": equivalences and inverses.
+
+data Kind : Set where
+ equivalent inverse : Kind
+
+Isomorphism : Kind → ∀ {ℓ₁ ℓ₂} → Set ℓ₁ → Set ℓ₂ → Set _
+Isomorphism inverse A B = Inverse (P.setoid A) (P.setoid B)
+Isomorphism equivalent A B = Equivalent (P.setoid A) (P.setoid B)
+
+-- Inverses are stronger, equivalences weaker.
+
+⇿⇒ : ∀ {k x y} {X : Set x} {Y : Set y} →
+ Isomorphism inverse X Y → Isomorphism k X Y
+⇿⇒ {inverse} = Fun.id
+⇿⇒ {equivalent} = Inverse.equivalent
+
+⇒⇔ : ∀ {k x y} {X : Set x} {Y : Set y} →
+ Isomorphism k X Y → Isomorphism equivalent X Y
+⇒⇔ {inverse} = Inverse.equivalent
+⇒⇔ {equivalent} = Fun.id
+
+-- Equational reasoning for isomorphisms.
+
+module EquationalReasoning where
+
+ private
+
+ refl : ∀ {k ℓ} → Reflexive (Isomorphism k {ℓ})
+ refl {inverse} = id
+ refl {equivalent} = Eq.id
+
+ trans : ∀ {k ℓ₁ ℓ₂ ℓ₃} →
+ Trans (Isomorphism k {ℓ₁} {ℓ₂})
+ (Isomorphism k {ℓ₂} {ℓ₃})
+ (Isomorphism k {ℓ₁} {ℓ₃})
+ trans {inverse} = flip _∘_
+ trans {equivalent} = flip Eq._∘_
+
+ sym : ∀ {k ℓ₁ ℓ₂} →
+ Sym (Isomorphism k {ℓ₁} {ℓ₂})
+ (Isomorphism k {ℓ₂} {ℓ₁})
+ sym {inverse} = Dummy.sym
+ sym {equivalent} = Eq.sym
+
+ infix 2 _∎
+ infixr 2 _≈⟨_⟩_ _⇿⟨_⟩_
+
+ _≈⟨_⟩_ : ∀ {k x y z} (X : Set x) {Y : Set y} {Z : Set z} →
+ Isomorphism k X Y → Isomorphism k Y Z → Isomorphism k X Z
+ _ ≈⟨ X≈Y ⟩ Y≈Z = trans X≈Y Y≈Z
+
+ _⇿⟨_⟩_ : ∀ {k x y z} (X : Set x) {Y : Set y} {Z : Set z} →
+ X ⇿ Y → Isomorphism k Y Z → Isomorphism k X Z
+ X ⇿⟨ X⇿Y ⟩ Y⇔Z = X ≈⟨ ⇿⇒ X⇿Y ⟩ Y⇔Z
+
+ _∎ : ∀ {k x} (X : Set x) → Isomorphism k X X
+ X ∎ = refl
+
+-- For fixed universe levels we can construct a setoid.
+
+Isomorphism-setoid : Kind → (ℓ : Level) → Setoid _ _
+Isomorphism-setoid k ℓ = record
+ { Carrier = Set ℓ
+ ; _≈_ = Isomorphism k
+ ; isEquivalence =
+ record {refl = _ ∎; sym = sym; trans = _≈⟨_⟩_ _}
+ } where open EquationalReasoning
+
+-- Every unary relation induces an equivalence relation. (No claim is
+-- made that this relation is unique.)
+
+InducedEquivalence₁ : Kind → ∀ {a s} {A : Set a}
+ (S : A → Set s) → Setoid _ _
+InducedEquivalence₁ k S = record
+ { _≈_ = λ x y → Isomorphism k (S x) (S y)
+ ; isEquivalence = record {refl = refl; sym = sym; trans = trans}
+ } where open Setoid (Isomorphism-setoid _ _)
+
+-- Every binary relation induces an equivalence relation. (No claim is
+-- made that this relation is unique.)
+
+InducedEquivalence₂ : Kind → ∀ {a b s} {A : Set a} {B : Set b}
+ (_S_ : A → B → Set s) → Setoid _ _
+InducedEquivalence₂ k _S_ = record
+ { _≈_ = λ x y → ∀ {z} → Isomorphism k (z S x) (z S y)
+ ; isEquivalence = record
+ { refl = refl
+ ; sym = λ i≈j → sym i≈j
+ ; trans = λ i≈j j≈k → trans i≈j j≈k
+ }
+ } where open Setoid (Isomorphism-setoid _ _)
+
+open Dummy public
diff --git a/src/Function/Inverse/TypeIsomorphisms.agda b/src/Function/Inverse/TypeIsomorphisms.agda
new file mode 100644
index 0000000..4c58303
--- /dev/null
+++ b/src/Function/Inverse/TypeIsomorphisms.agda
@@ -0,0 +1,270 @@
+------------------------------------------------------------------------
+-- Various basic type isomorphisms
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Function.Inverse.TypeIsomorphisms where
+
+open import Algebra
+import Algebra.FunctionProperties as FP
+open import Data.Empty
+open import Data.Product as Prod hiding (swap)
+open import Data.Sum as Sum
+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.Inverse as Inv
+ using (Kind; Isomorphism; _⇿_; module Inverse)
+open import Relation.Binary
+open import Relation.Binary.Product.Pointwise
+open import Relation.Binary.PropositionalEquality as P using (_≡_; _≗_)
+open import Relation.Binary.Sum
+open import Relation.Nullary
+
+------------------------------------------------------------------------
+-- ⊥, ⊤, _×_ and _⊎_ form a commutative semiring
+
+×-CommutativeMonoid : Kind → (ℓ : Level) → CommutativeMonoid _ _
+×-CommutativeMonoid k ℓ = record
+ { Carrier = Set ℓ
+ ; _≈_ = Isomorphism k
+ ; _∙_ = _×_
+ ; ε = Lift ⊤
+ ; isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = Setoid.isEquivalence $
+ Inv.Isomorphism-setoid k ℓ
+ ; assoc = λ A B C → Inv.⇿⇒ $ assoc A B C
+ ; ∙-cong = ×-cong k
+ }
+ ; identityˡ = λ A → Inv.⇿⇒ $ left-identity A
+ ; comm = λ A B → Inv.⇿⇒ $ comm A B
+ }
+ }
+ where
+ open FP _⇿_
+
+ ×-cong : ∀ k {A B C D : Set ℓ} →
+ Isomorphism k A B → Isomorphism k C D →
+ Isomorphism k (A × C) (B × D)
+ ×-cong Inv.equivalent = _×-⇔_
+ ×-cong Inv.inverse = _×-⇿_
+
+ left-identity : LeftIdentity (Lift {ℓ = ℓ} ⊤) _×_
+ left-identity _ = record
+ { to = P.→-to-⟶ $ proj₂ {a = ℓ} {b = ℓ}
+ ; from = P.→-to-⟶ {b₁ = ℓ} λ y → _ , y
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.refl
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+
+ assoc : Associative _×_
+ assoc _ _ _ = record
+ { to = P.→-to-⟶ {a = ℓ} {b₁ = ℓ} λ t → (proj₁ (proj₁ t) , (proj₂ (proj₁ t) , proj₂ t))
+ ; from = P.→-to-⟶ {a = ℓ} {b₁ = ℓ} λ t → ((proj₁ t , proj₁ (proj₂ t)) , proj₂ (proj₂ t))
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.refl
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+
+ comm : Commutative _×_
+ comm _ _ = record
+ { to = P.→-to-⟶ $ Prod.swap {a = ℓ} {b = ℓ}
+ ; from = P.→-to-⟶ $ Prod.swap {a = ℓ} {b = ℓ}
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.refl
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+
+⊎-CommutativeMonoid : Kind → (ℓ : Level) → CommutativeMonoid _ _
+⊎-CommutativeMonoid k ℓ = record
+ { Carrier = Set ℓ
+ ; _≈_ = Isomorphism k
+ ; _∙_ = _⊎_
+ ; ε = Lift ⊥
+ ; isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = Setoid.isEquivalence $
+ Inv.Isomorphism-setoid k ℓ
+ ; assoc = λ A B C → Inv.⇿⇒ $ assoc A B C
+ ; ∙-cong = ⊎-cong k
+ }
+ ; identityˡ = λ A → Inv.⇿⇒ $ left-identity A
+ ; comm = λ A B → Inv.⇿⇒ $ comm A B
+ }
+ }
+ where
+ open FP _⇿_
+
+ ⊎-cong : ∀ k {A B C D : Set ℓ} →
+ Isomorphism k A B → Isomorphism k C D →
+ Isomorphism k (A ⊎ C) (B ⊎ D)
+ ⊎-cong Inv.equivalent = _⊎-⇔_
+ ⊎-cong Inv.inverse = _⊎-⇿_
+
+ left-identity : LeftIdentity (Lift ⊥) (_⊎_ {a = ℓ} {b = ℓ})
+ left-identity A = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { right-inverse-of = λ _ → P.refl
+ ; left-inverse-of =
+ [ ⊥-elim {Whatever = _ ≡ _} ∘ lower , (λ _ → P.refl) ]
+ }
+ }
+ where
+ to : Lift {ℓ = ℓ} ⊥ ⊎ A → A
+ to = [ (λ ()) ∘′ lower , id ]
+
+ from : A → Lift {ℓ = ℓ} ⊥ ⊎ A
+ from = inj₂
+
+ assoc : Associative _⊎_
+ assoc A B C = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = [ [ (λ _ → P.refl) , (λ _ → P.refl) ] , (λ _ → P.refl) ]
+ ; right-inverse-of = [ (λ _ → P.refl) , [ (λ _ → P.refl) , (λ _ → P.refl) ] ]
+ }
+ }
+ where
+ to : (A ⊎ B) ⊎ C → A ⊎ (B ⊎ C)
+ to = [ [ inj₁ , inj₂ ∘ inj₁ ] , inj₂ ∘ inj₂ ]
+
+ from : A ⊎ (B ⊎ C) → (A ⊎ B) ⊎ C
+ from = [ inj₁ ∘ inj₁ , [ inj₁ ∘ inj₂ , inj₂ ] ]
+
+ comm : Commutative _⊎_
+ comm _ _ = record
+ { to = P.→-to-⟶ swap
+ ; from = P.→-to-⟶ swap
+ ; inverse-of = record
+ { left-inverse-of = inv
+ ; right-inverse-of = inv
+ }
+ }
+ where
+ swap : {A B : Set ℓ} → A ⊎ B → B ⊎ A
+ swap = [ inj₂ , inj₁ ]
+
+ inv : ∀ {A B} → swap ∘ swap {A} {B} ≗ id
+ inv = [ (λ _ → P.refl) , (λ _ → P.refl) ]
+
+×⊎-CommutativeSemiring : Kind → (ℓ : Level) → CommutativeSemiring _ _
+×⊎-CommutativeSemiring k ℓ = record
+ { Carrier = Set ℓ
+ ; _≈_ = Isomorphism k
+ ; _+_ = _⊎_
+ ; _*_ = _×_
+ ; 0# = Lift ⊥
+ ; 1# = Lift ⊤
+ ; isCommutativeSemiring = record
+ { +-isCommutativeMonoid = isCommutativeMonoid $
+ ⊎-CommutativeMonoid k ℓ
+ ; *-isCommutativeMonoid = isCommutativeMonoid $
+ ×-CommutativeMonoid k ℓ
+ ; distribʳ = λ A B C → Inv.⇿⇒ $ right-distrib A B C
+ ; zeroˡ = λ A → Inv.⇿⇒ $ left-zero A
+ }
+ }
+ where
+ open CommutativeMonoid
+ open FP _⇿_
+
+ left-zero : LeftZero (Lift ⊥) (_×_ {a = ℓ} {b = ℓ})
+ left-zero A = record
+ { to = P.→-to-⟶ $ proj₁ {a = ℓ} {b = ℓ}
+ ; from = P.→-to-⟶ (⊥-elim ∘′ lower)
+ ; inverse-of = record
+ { left-inverse-of = λ p → ⊥-elim (lower $ proj₁ p)
+ ; right-inverse-of = λ x → ⊥-elim (lower x)
+ }
+ }
+
+ right-distrib : _×_ DistributesOverʳ _⊎_
+ right-distrib A B C = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { right-inverse-of = [ (λ _ → P.refl) , (λ _ → P.refl) ]
+ ; left-inverse-of =
+ uncurry [ (λ _ _ → P.refl) , (λ _ _ → P.refl) ]
+ }
+ }
+ where
+ to : (B ⊎ C) × A → B × A ⊎ C × A
+ to = uncurry [ curry inj₁ , curry inj₂ ]
+
+ from : B × A ⊎ C × A → (B ⊎ C) × A
+ from = [ Prod.map inj₁ id , Prod.map inj₂ id ]
+
+------------------------------------------------------------------------
+-- Some reordering lemmas
+
+ΠΠ⇿ΠΠ : ∀ {a b p} {A : Set a} {B : Set b} (P : A → B → Set p) →
+ ((x : A) (y : B) → P x y) ⇿ ((y : B) (x : A) → P x y)
+ΠΠ⇿ΠΠ _ = record
+ { to = P.→-to-⟶ λ f x y → f y x
+ ; from = P.→-to-⟶ λ f y x → f x y
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.refl
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+
+∃∃⇿∃∃ : ∀ {a b p} {A : Set a} {B : Set b} (P : A → B → Set p) →
+ (∃₂ λ x y → P x y) ⇿ (∃₂ λ y x → P x y)
+∃∃⇿∃∃ {a} {b} {p} _ = record
+ { to = P.→-to-⟶ {a = ℓ} λ p → (proj₁ (proj₂ p) , proj₁ p , proj₂ (proj₂ p))
+ ; from = P.→-to-⟶ {a = ℓ} λ p → (proj₁ (proj₂ p) , proj₁ p , proj₂ (proj₂ p))
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.refl
+ ; right-inverse-of = λ _ → P.refl
+ }
+ } where ℓ = p ⊔ (b ⊔ a)
+
+------------------------------------------------------------------------
+-- Implicit and explicit function spaces are isomorphic
+
+Π⇿Π : ∀ {a b} {A : Set a} {B : A → Set b} →
+ ((x : A) → B x) ⇿ ({x : A} → B x)
+Π⇿Π = record
+ { to = P.→-to-⟶ λ f {x} → f x
+ ; from = P.→-to-⟶ λ f x → f {x}
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.refl
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+
+------------------------------------------------------------------------
+-- ¬_ 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} →
+ 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)
diff --git a/src/Function/LeftInverse.agda b/src/Function/LeftInverse.agda
new file mode 100644
index 0000000..36dae63
--- /dev/null
+++ b/src/Function/LeftInverse.agda
@@ -0,0 +1,90 @@
+------------------------------------------------------------------------
+-- Left inverses
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Function.LeftInverse where
+
+open import Data.Product
+open import Level
+import Relation.Binary.EqReasoning as EqReasoning
+open import Relation.Binary
+open import Function.Equality as F
+ using (_⟶_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_)
+open import Function.Equivalence using (Equivalent)
+open import Function.Injection using (Injective; Injection)
+
+-- Left and right inverses.
+
+_LeftInverseOf_ :
+ ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} →
+ To ⟶ From → From ⟶ To → Set _
+_LeftInverseOf_ {From = From} f g = ∀ x → f ⟨$⟩ (g ⟨$⟩ x) ≈ x
+ where open Setoid From
+
+_RightInverseOf_ :
+ ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} →
+ To ⟶ From → From ⟶ To → Set _
+f RightInverseOf g = g LeftInverseOf f
+
+-- The set of all left inverses between two setoids.
+
+record LeftInverse {f₁ f₂ t₁ t₂}
+ (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) :
+ Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
+ field
+ to : From ⟶ To
+ from : To ⟶ From
+ left-inverse-of : from LeftInverseOf to
+
+ open Setoid From
+ open EqReasoning From
+
+ injective : Injective to
+ injective {x} {y} eq = begin
+ x ≈⟨ sym (left-inverse-of x) ⟩
+ from ⟨$⟩ (to ⟨$⟩ x) ≈⟨ F.cong from eq ⟩
+ from ⟨$⟩ (to ⟨$⟩ y) ≈⟨ left-inverse-of y ⟩
+ y ∎
+
+ injection : Injection From To
+ injection = record { to = to; injective = injective }
+
+ equivalent : Equivalent From To
+ equivalent = record
+ { to = to
+ ; from = from
+ }
+
+-- The set of all right inverses between two setoids.
+
+RightInverse : ∀ {f₁ f₂ t₁ t₂}
+ (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) → Set _
+RightInverse From To = LeftInverse To From
+
+-- Identity and composition.
+
+id : ∀ {s₁ s₂} {S : Setoid s₁ s₂} → LeftInverse S S
+id {S = S} = record
+ { to = F.id
+ ; from = F.id
+ ; left-inverse-of = λ _ → Setoid.refl S
+ }
+
+infixr 9 _∘_
+
+_∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂}
+ {F : Setoid f₁ f₂} {M : Setoid m₁ m₂} {T : Setoid t₁ t₂} →
+ LeftInverse M T → LeftInverse F M → LeftInverse F T
+_∘_ {F = F} f g = record
+ { to = to f ⟪∘⟫ to g
+ ; from = from g ⟪∘⟫ from f
+ ; left-inverse-of = λ x → begin
+ from g ⟨$⟩ (from f ⟨$⟩ (to f ⟨$⟩ (to g ⟨$⟩ x))) ≈⟨ F.cong (from g) (left-inverse-of f (to g ⟨$⟩ x)) ⟩
+ from g ⟨$⟩ (to g ⟨$⟩ x) ≈⟨ left-inverse-of g x ⟩
+ x ∎
+ }
+ where
+ open LeftInverse
+ open EqReasoning F
diff --git a/src/Function/Surjection.agda b/src/Function/Surjection.agda
new file mode 100644
index 0000000..a0269cc
--- /dev/null
+++ b/src/Function/Surjection.agda
@@ -0,0 +1,83 @@
+------------------------------------------------------------------------
+-- Surjections
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Function.Surjection where
+
+open import Level
+open import Function.Equality as F
+ using (_⟶_) renaming (_∘_ to _⟪∘⟫_)
+open import Function.Equivalence using (Equivalent)
+open import Function.Injection hiding (id; _∘_)
+open import Function.LeftInverse as Left hiding (id; _∘_)
+open import Data.Product
+open import Relation.Binary
+
+-- Surjective functions.
+
+record Surjective {f₁ f₂ t₁ t₂}
+ {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
+ (to : From ⟶ To) :
+ Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
+ field
+ from : To ⟶ From
+ right-inverse-of : from RightInverseOf to
+
+-- The set of all surjections between two setoids.
+
+record Surjection {f₁ f₂ t₁ t₂}
+ (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) :
+ Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
+ field
+ to : From ⟶ To
+ surjective : Surjective to
+
+ open Surjective surjective public
+
+ right-inverse : RightInverse From To
+ right-inverse = record
+ { to = from
+ ; from = to
+ ; left-inverse-of = right-inverse-of
+ }
+
+ injective : Injective from
+ injective = LeftInverse.injective right-inverse
+
+ injection : Injection To From
+ injection = LeftInverse.injection right-inverse
+
+ equivalent : Equivalent From To
+ equivalent = record
+ { to = to
+ ; from = from
+ }
+
+-- Identity and composition.
+
+id : ∀ {s₁ s₂} {S : Setoid s₁ s₂} → Surjection S S
+id {S = S} = record
+ { to = F.id
+ ; surjective = record
+ { from = LeftInverse.to id′
+ ; right-inverse-of = LeftInverse.left-inverse-of id′
+ }
+ } where id′ = Left.id {S = S}
+
+infixr 9 _∘_
+
+_∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂}
+ {F : Setoid f₁ f₂} {M : Setoid m₁ m₂} {T : Setoid t₁ t₂} →
+ Surjection M T → Surjection F M → Surjection F T
+f ∘ g = record
+ { to = to f ⟪∘⟫ to g
+ ; surjective = record
+ { from = LeftInverse.to g∘f
+ ; right-inverse-of = LeftInverse.left-inverse-of g∘f
+ }
+ }
+ where
+ open Surjection
+ g∘f = Left._∘_ (right-inverse g) (right-inverse f)
diff --git a/src/IO.agda b/src/IO.agda
index cfc3c32..94fbbaf 100644
--- a/src/IO.agda
+++ b/src/IO.agda
@@ -88,6 +88,14 @@ writeFile∞ f s =
writeFile : String → String → IO ⊤
writeFile f s = writeFile∞ f (toCostring s)
+appendFile∞ : String → Costring → IO ⊤
+appendFile∞ f s =
+ ♯ lift (Prim.appendFile f (Haskell.fromColist s)) >>
+ ♯ return _
+
+appendFile : String → String → IO ⊤
+appendFile f s = appendFile∞ f (toCostring s)
+
putStr∞ : Costring → IO ⊤
putStr∞ s =
♯ lift (Prim.putStr (Haskell.fromColist s)) >>
diff --git a/src/IO/Primitive.agda b/src/IO/Primitive.agda
index a529c6f..7764519 100644
--- a/src/IO/Primitive.agda
+++ b/src/IO/Primitive.agda
@@ -41,11 +41,13 @@ postulate
getContents : IO Costring
readFile : String → IO Costring
writeFile : String → Costring → IO Unit
+ appendFile : String → Costring → IO Unit
putStr : Costring → IO Unit
putStrLn : Costring → IO Unit
{-# COMPILED getContents getContents #-}
{-# COMPILED readFile readFile #-}
{-# COMPILED writeFile writeFile #-}
+{-# COMPILED appendFile appendFile #-}
{-# COMPILED putStr putStr #-}
{-# COMPILED putStrLn putStrLn #-}
diff --git a/src/Induction.agda b/src/Induction.agda
index 26d1459..3902c65 100644
--- a/src/Induction.agda
+++ b/src/Induction.agda
@@ -4,6 +4,9 @@
{-# OPTIONS --universe-polymorphism #-}
+-- The idea underlying Induction.* comes from Epigram 1, see Section 4
+-- of "The view from the left" by McBride and McKinna.
+
-- Note: The types in this module can perhaps be easier to understand
-- if they are normalised. Note also that Agda can do the
-- normalisation for you.
diff --git a/src/Induction/Lexicographic.agda b/src/Induction/Lexicographic.agda
index 9181d75..6cf5a7c 100644
--- a/src/Induction/Lexicographic.agda
+++ b/src/Induction/Lexicographic.agda
@@ -11,28 +11,35 @@ open import Data.Product
-- The structure of lexicographic induction.
-_⊗_ : ∀ {ℓ} {A B : Set ℓ} →
- RecStruct A → RecStruct B → RecStruct (A × B)
-_⊗_ RecA RecB P (x , y) =
+Σ-Rec : ∀ {ℓ} {A : Set ℓ} {B : A → Set ℓ} →
+ RecStruct A → (∀ x → RecStruct (B x)) → RecStruct (Σ A B)
+Σ-Rec RecA RecB P (x , y) =
-- Either x is constant and y is "smaller", ...
- RecB (λ y' → P (x , y')) y
+ RecB x (λ y' → P (x , y')) y
×
-- ...or x is "smaller" and y is arbitrary.
RecA (λ x' → ∀ y' → P (x' , y')) x
+_⊗_ : ∀ {ℓ} {A B : Set ℓ} →
+ RecStruct A → RecStruct B → RecStruct (A × B)
+RecA ⊗ RecB = Σ-Rec RecA (λ _ → RecB)
+
-- Constructs a recursor builder for lexicographic induction.
-[_⊗_] : ∀ {ℓ} {A B : Set ℓ}
- {RecA : RecStruct A} → RecursorBuilder RecA →
- {RecB : RecStruct B} → RecursorBuilder RecB →
- RecursorBuilder (RecA ⊗ RecB)
-[_⊗_] {RecA = RecA} recA {RecB = RecB} recB P f (x , y) =
+Σ-rec-builder :
+ ∀ {ℓ} {A : Set ℓ} {B : A → Set ℓ}
+ {RecA : RecStruct A}
+ {RecB : ∀ x → RecStruct (B x)} →
+ RecursorBuilder RecA → (∀ x → RecursorBuilder (RecB x)) →
+ RecursorBuilder (Σ-Rec RecA RecB)
+Σ-rec-builder {RecA = RecA} {RecB = RecB} recA recB P f (x , y) =
(p₁ x y p₂x , p₂x)
where
p₁ : ∀ x y →
RecA (λ x' → ∀ y' → P (x' , y')) x →
- RecB (λ y' → P (x , y')) y
- p₁ x y x-rec = recB (λ y' → P (x , y'))
+ RecB x (λ y' → P (x , y')) y
+ p₁ x y x-rec = recB x
+ (λ y' → P (x , y'))
(λ y y-rec → f (x , y) (y-rec , x-rec))
y
@@ -42,6 +49,11 @@ _⊗_ RecA RecB P (x , y) =
p₂x = p₂ x
+[_⊗_] : ∀ {ℓ} {A B : Set ℓ} {RecA : RecStruct A} {RecB : RecStruct B} →
+ RecursorBuilder RecA → RecursorBuilder RecB →
+ RecursorBuilder (RecA ⊗ RecB)
+[ recA ⊗ recB ] = Σ-rec-builder recA (λ _ → recB)
+
------------------------------------------------------------------------
-- Example
diff --git a/src/Induction/Nat.agda b/src/Induction/Nat.agda
index c3cdf5d..78b6b81 100644
--- a/src/Induction/Nat.agda
+++ b/src/Induction/Nat.agda
@@ -4,15 +4,16 @@
module Induction.Nat where
-open import Data.Function
+open import Function
open import Data.Nat
open import Data.Fin
open import Data.Fin.Props
open import Data.Product
open import Data.Unit
open import Induction
-import Induction.WellFounded as WF
+open import Induction.WellFounded as WF
open import Relation.Binary.PropositionalEquality
+open import Relation.Unary
------------------------------------------------------------------------
-- Ordinary induction
@@ -46,18 +47,20 @@ cRec = build cRec-builder
------------------------------------------------------------------------
-- Complete induction based on _<′_
-open WF _<′_ using (acc) renaming (Acc to <-Acc)
+<-Rec : RecStruct ℕ
+<-Rec = WfRec _<′_
-<-allAcc : ∀ n → <-Acc n
-<-allAcc n = acc (helper n)
- where
- helper : ∀ n m → m <′ n → <-Acc m
- helper zero _ ()
- helper (suc n) .n ≤′-refl = acc (helper n)
- helper (suc n) m (≤′-step m<n) = helper n m m<n
+mutual
-open WF _<′_ public using () renaming (WfRec to <-Rec)
-open WF.All _<′_ <-allAcc public
+ <-well-founded : Well-founded _<′_
+ <-well-founded n = acc (<-well-founded′ n)
+
+ <-well-founded′ : ∀ n → <-Rec (Acc _<′_) n
+ <-well-founded′ zero _ ()
+ <-well-founded′ (suc n) .n ≤′-refl = <-well-founded n
+ <-well-founded′ (suc n) m (≤′-step m<n) = <-well-founded′ n m m<n
+
+open WF.All <-well-founded public
renaming ( wfRec-builder to <-rec-builder
; wfRec to <-rec
)
@@ -65,17 +68,13 @@ open WF.All _<′_ <-allAcc public
------------------------------------------------------------------------
-- Complete induction based on _≺_
-open WF _≺_ renaming (Acc to ≺-Acc)
-
-<-Acc⇒≺-Acc : ∀ {n} → <-Acc n → ≺-Acc n
-<-Acc⇒≺-Acc (acc rs) =
- acc (λ m m≺n → <-Acc⇒≺-Acc (rs m (≺⇒<′ m≺n)))
+≺-Rec : RecStruct ℕ
+≺-Rec = WfRec _≺_
-≺-allAcc : ∀ n → ≺-Acc n
-≺-allAcc n = <-Acc⇒≺-Acc (<-allAcc n)
+≺-well-founded : Well-founded _≺_
+≺-well-founded = Subrelation.well-founded ≺⇒<′ <-well-founded
-open WF _≺_ public using () renaming (WfRec to ≺-Rec)
-open WF.All _≺_ ≺-allAcc public
+open WF.All ≺-well-founded public
renaming ( wfRec-builder to ≺-rec-builder
; wfRec to ≺-rec
)
diff --git a/src/Induction/WellFounded.agda b/src/Induction/WellFounded.agda
index 58750c3..9ed9e52 100644
--- a/src/Induction/WellFounded.agda
+++ b/src/Induction/WellFounded.agda
@@ -6,42 +6,126 @@
open import Relation.Binary
-module Induction.WellFounded {a} {A : Set a} (_<_ : Rel A a) where
+module Induction.WellFounded where
+open import Data.Product
+open import Function
open import Induction
+open import Relation.Unary
-- When using well-founded recursion you can recurse arbitrarily, as
-- long as the arguments become smaller, and "smaller" is
-- well-founded.
-WfRec : RecStruct A
-WfRec P x = ∀ y → y < x → P y
+WfRec : ∀ {a} {A : Set a} → Rel A a → RecStruct A
+WfRec _<_ P x = ∀ y → y < x → P y
+
+-- The accessibility predicate: x is accessible if everything which is
+-- smaller than x is also accessible (inductively).
+
+data Acc {a} {A : Set a} (_<_ : Rel A a) (x : A) : Set a where
+ acc : (rs : WfRec _<_ (Acc _<_) x) → Acc _<_ x
-- The accessibility predicate encodes what it means to be
-- well-founded; if all elements are accessible, then _<_ is
-- well-founded.
-data Acc (x : A) : Set a where
- acc : (rs : WfRec Acc x) → Acc x
+Well-founded : ∀ {a} {A : Set a} → Rel A a → Set a
+Well-founded _<_ = ∀ x → Acc _<_ x
-- Well-founded induction for the subset of accessible elements:
-module Some where
+module Some {a} {A : Set a} {_<_ : Rel A a} where
- wfRec-builder : SubsetRecursorBuilder Acc WfRec
+ wfRec-builder : SubsetRecursorBuilder (Acc _<_) (WfRec _<_)
wfRec-builder P f x (acc rs) = λ y y<x →
f y (wfRec-builder P f y (rs y y<x))
- wfRec : SubsetRecursor Acc WfRec
+ wfRec : SubsetRecursor (Acc _<_) (WfRec _<_)
wfRec = subsetBuild wfRec-builder
-- Well-founded induction for all elements, assuming they are all
-- accessible:
-module All (allAcc : ∀ x → Acc x) where
+module All {a} {A : Set a} {_<_ : Rel A a}
+ (wf : Well-founded _<_) where
- wfRec-builder : RecursorBuilder WfRec
- wfRec-builder P f x = Some.wfRec-builder P f x (allAcc x)
+ wfRec-builder : RecursorBuilder (WfRec _<_)
+ wfRec-builder P f x = Some.wfRec-builder P f x (wf x)
- wfRec : Recursor WfRec
+ wfRec : Recursor (WfRec _<_)
wfRec = build wfRec-builder
+
+-- It might be useful to establish proofs of Acc or Well-founded using
+-- combinators such as the ones below (see, for instance,
+-- "Constructing Recursion Operators in Intuitionistic Type Theory" by
+-- Lawrence C Paulson).
+
+module Subrelation {a} {A : Set a} {_<₁_ _<₂_ : Rel A a}
+ (<₁⇒<₂ : ∀ {x y} → x <₁ y → x <₂ y) where
+
+ accessible : Acc _<₂_ ⊆ Acc _<₁_
+ accessible (acc rs) = acc (λ y y<x → accessible (rs y (<₁⇒<₂ y<x)))
+
+ well-founded : Well-founded _<₂_ → Well-founded _<₁_
+ well-founded wf = λ x → accessible (wf x)
+
+module Inverse-image {ℓ} {A B : Set ℓ} {_<_ : Rel B ℓ}
+ (f : A → B) where
+
+ accessible : ∀ {x} → Acc _<_ (f x) → Acc (_<_ on f) x
+ accessible (acc rs) = acc (λ y fy<fx → accessible (rs (f y) fy<fx))
+
+ well-founded : Well-founded _<_ → Well-founded (_<_ on f)
+ well-founded wf = λ x → accessible (wf (f x))
+
+module Transitive-closure {a} {A : Set a} (_<_ : Rel A a) where
+
+ infix 4 _<⁺_
+
+ data _<⁺_ : Rel A a where
+ [_] : ∀ {x y} (x<y : x < y) → x <⁺ y
+ trans : ∀ {x y z} (x<y : x <⁺ y) (y<z : y <⁺ z) → x <⁺ z
+
+ downwards-closed : ∀ {x y} → Acc _<⁺_ y → x <⁺ y → Acc _<⁺_ x
+ downwards-closed (acc rs) x<y = acc (λ z z<x → rs z (trans z<x x<y))
+
+ mutual
+
+ accessible : ∀ {x} → Acc _<_ x → Acc _<⁺_ x
+ accessible acc-x = acc (accessible′ acc-x)
+
+ accessible′ : ∀ {x} → Acc _<_ x → WfRec _<⁺_ (Acc _<⁺_) x
+ accessible′ (acc rs) y [ y<x ] = accessible (rs y y<x)
+ accessible′ acc-x y (trans y<z z<x) =
+ downwards-closed (accessible′ acc-x _ z<x) y<z
+
+ well-founded : Well-founded _<_ → Well-founded _<⁺_
+ well-founded wf = λ x → accessible (wf x)
+
+module Lexicographic {ℓ} {A : Set ℓ} {B : A → Set ℓ}
+ (RelA : Rel A ℓ)
+ (RelB : ∀ x → Rel (B x) ℓ) where
+
+ data _<_ : Rel (Σ A B) ℓ where
+ left : ∀ {x₁ y₁ x₂ y₂} (x₁<x₂ : RelA x₁ x₂) → (x₁ , y₁) < (x₂ , y₂)
+ right : ∀ {x y₁ y₂} (y₁<y₂ : RelB x y₁ y₂) → (x , y₁) < (x , y₂)
+
+ mutual
+
+ accessible : ∀ {x y} →
+ Acc RelA x → (∀ {x} → Well-founded (RelB x)) →
+ Acc _<_ (x , y)
+ accessible accA wfB = acc (accessible′ accA (wfB _) wfB)
+
+ accessible′ :
+ ∀ {x y} →
+ Acc RelA x → Acc (RelB x) y → (∀ {x} → Well-founded (RelB x)) →
+ WfRec _<_ (Acc _<_) (x , y)
+ accessible′ (acc rsA) _ wfB ._ (left x′<x) = accessible (rsA _ x′<x) wfB
+ accessible′ accA (acc rsB) wfB ._ (right y′<y) =
+ acc (accessible′ accA (rsB _ y′<y) wfB)
+
+ well-founded : Well-founded RelA → (∀ {x} → Well-founded (RelB x)) →
+ Well-founded _<_
+ well-founded wfA wfB p = accessible (wfA (proj₁ p)) wfB
diff --git a/src/Level.agda b/src/Level.agda
index e1dc8e9..2aefb6a 100644
--- a/src/Level.agda
+++ b/src/Level.agda
@@ -25,9 +25,12 @@ zero ⊔ j = j
suc i ⊔ zero = suc i
suc i ⊔ suc j = suc (i ⊔ j)
+{-# BUILTIN LEVELMAX _⊔_ #-}
+
-- Lifting.
-{-# BUILTIN LEVELMAX _⊔_ #-}
+record Lift {a ℓ} (A : Set a) : Set (a ⊔ ℓ) where
+ constructor lift
+ field lower : A
-data Lift {a ℓ} (A : Set a) : Set (a ⊔ ℓ) where
- lift : (x : A) → Lift A
+open Lift public
diff --git a/src/Reflection.agda b/src/Reflection.agda
new file mode 100644
index 0000000..974c9d8
--- /dev/null
+++ b/src/Reflection.agda
@@ -0,0 +1,81 @@
+------------------------------------------------------------------------
+-- Support for reflection
+------------------------------------------------------------------------
+
+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 Relation.Binary
+open import Relation.Binary.PropositionalEquality
+open import Relation.Binary.PropositionalEquality.TrustMe
+open import Relation.Nullary
+
+------------------------------------------------------------------------
+-- Names
+
+-- Names.
+
+postulate Name : Set
+
+{-# BUILTIN QNAME Name #-}
+
+private
+ primitive
+ primQNameEquality : Name → Name → Bool
+
+-- Equality of names is decidable.
+
+infix 4 _==_ _≟_
+
+_==_ : Name → Name → Bool
+_==_ = primQNameEquality
+
+_≟_ : Decidable {A = Name} _≡_
+s₁ ≟ s₂ with s₁ == s₂
+... | true = yes trustMe
+... | false = no whatever
+ where postulate whatever : _
+
+------------------------------------------------------------------------
+-- Terms
+
+-- Is the argument explicit?
+
+Explicit? = Bool
+
+-- Arguments.
+
+data Arg A : Set where
+ arg : Explicit? → A → Arg A
+
+{-# BUILTIN ARG Arg #-}
+{-# BUILTIN ARGARG arg #-}
+
+-- Terms.
+
+data Term : Set where
+ -- Variable applied to arguments.
+ var : ℕ → List (Arg Term) → Term
+ -- Constructor applied to arguments.
+ con : Name → List (Arg Term) → Term
+ -- Identifier applied to arguments.
+ def : Name → List (Arg Term) → Term
+ -- Explicit or implicit λ abstraction.
+ lam : Explicit? → Term → Term
+ -- Pi-type.
+ pi : Arg Term → Term → Term
+ -- An arbitrary sort (Set, for instance).
+ sort : Term
+ -- Anything.
+ unknown : Term
+
+{-# BUILTIN AGDATERM Term #-}
+{-# BUILTIN AGDATERMVAR var #-}
+{-# BUILTIN AGDATERMCON con #-}
+{-# BUILTIN AGDATERMDEF def #-}
+{-# BUILTIN AGDATERMLAM lam #-}
+{-# BUILTIN AGDATERMPI pi #-}
+{-# BUILTIN AGDATERMSORT sort #-}
+{-# BUILTIN AGDATERMUNSUPPORTED unknown #-}
diff --git a/src/Relation/Binary.agda b/src/Relation/Binary.agda
index 1940876..c5aaed4 100644
--- a/src/Relation/Binary.agda
+++ b/src/Relation/Binary.agda
@@ -8,11 +8,12 @@ module Relation.Binary where
open import Data.Product
open import Data.Sum
-open import Data.Function
+open import Function
open import Level
import Relation.Binary.PropositionalEquality.Core as PropEq
open import Relation.Binary.Consequences
open import Relation.Binary.Core as Core using (_≡_)
+import Relation.Binary.Indexed.Core as I
------------------------------------------------------------------------
-- Simple properties and equivalence relations
@@ -31,13 +32,16 @@ record IsPreorder {a ℓ₁ ℓ₂} {A : Set a}
-- Reflexivity is expressed in terms of an underlying equality:
reflexive : _≈_ ⇒ _∼_
trans : Transitive _∼_
- ∼-resp-≈ : _∼_ Respects₂ _≈_
module Eq = IsEquivalence isEquivalence
refl : Reflexive _∼_
refl = reflexive Eq.refl
+ ∼-resp-≈ : _∼_ Respects₂ _≈_
+ ∼-resp-≈ = (λ x≈y z∼x → trans z∼x (reflexive x≈y))
+ , (λ x≈y x∼z → trans (reflexive $ Eq.sym x≈y) x∼z)
+
record Preorder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where
infix 4 _≈_ _∼_
field
@@ -67,12 +71,24 @@ record Setoid c ℓ : Set (suc (c ⊔ ℓ)) where
{ isEquivalence = PropEq.isEquivalence
; reflexive = reflexive
; trans = trans
- ; ∼-resp-≈ = PropEq.resp₂ _≈_
}
- preorder : Preorder c zero ℓ
+ preorder : Preorder c c ℓ
preorder = record { isPreorder = isPreorder }
+ -- A trivially indexed setoid.
+
+ indexedSetoid : ∀ {i} {I : Set i} → I.Setoid I c _
+ indexedSetoid = record
+ { Carrier = λ _ → Carrier
+ ; _≈_ = _≈_
+ ; isEquivalence = record
+ { refl = refl
+ ; sym = sym
+ ; trans = trans
+ }
+ }
+
------------------------------------------------------------------------
-- Decidable equivalence relations
@@ -126,6 +142,56 @@ record Poset c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where
preorder = record { isPreorder = isPreorder }
------------------------------------------------------------------------
+-- Decidable partial orders
+
+record IsDecPartialOrder {a ℓ₁ ℓ₂} {A : Set a}
+ (_≈_ : Rel A ℓ₁) (_≤_ : Rel A ℓ₂) :
+ Set (a ⊔ ℓ₁ ⊔ ℓ₂) where
+ infix 4 _≟_ _≤?_
+ field
+ isPartialOrder : IsPartialOrder _≈_ _≤_
+ _≟_ : Decidable _≈_
+ _≤?_ : Decidable _≤_
+
+ private
+ module PO = IsPartialOrder isPartialOrder
+ open PO public hiding (module Eq)
+
+ module Eq where
+
+ isDecEquivalence : IsDecEquivalence _≈_
+ isDecEquivalence = record
+ { isEquivalence = PO.isEquivalence
+ ; _≟_ = _≟_
+ }
+
+ open IsDecEquivalence isDecEquivalence public
+
+record DecPoset c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where
+ infix 4 _≈_ _≤_
+ field
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ₁
+ _≤_ : Rel Carrier ℓ₂
+ isDecPartialOrder : IsDecPartialOrder _≈_ _≤_
+
+ private
+ module DPO = IsDecPartialOrder isDecPartialOrder
+ open DPO public hiding (module Eq)
+
+ poset : Poset c ℓ₁ ℓ₂
+ poset = record { isPartialOrder = isPartialOrder }
+
+ open Poset poset public using (preorder)
+
+ module Eq where
+
+ decSetoid : DecSetoid c ℓ₁
+ decSetoid = record { isDecEquivalence = DPO.Eq.isDecEquivalence }
+
+ open DecSetoid decSetoid public
+
+------------------------------------------------------------------------
-- Strict partial orders
record IsStrictPartialOrder {a ℓ₁ ℓ₂} {A : Set a}
diff --git a/src/Relation/Binary/Consequences.agda b/src/Relation/Binary/Consequences.agda
index 7972cea..d07c2a2 100644
--- a/src/Relation/Binary/Consequences.agda
+++ b/src/Relation/Binary/Consequences.agda
@@ -7,9 +7,9 @@
module Relation.Binary.Consequences where
open import Relation.Binary.Core hiding (refl)
-open import Relation.Nullary.Core
+open import Relation.Nullary
open import Relation.Binary.PropositionalEquality.Core
-open import Data.Function
+open import Function
open import Data.Sum
open import Data.Product
open import Data.Empty
@@ -19,28 +19,28 @@ open import Data.Empty
open import Relation.Binary.Consequences.Core public
trans∧irr⟶asym :
- ∀ {a ℓ₁ ℓ₂} {A : Set a} → {≈ : Rel A ℓ₁} {< : Rel A ℓ₂} →
- Reflexive ≈ →
- Transitive < → Irreflexive ≈ < → Asymmetric <
+ ∀ {a ℓ₁ ℓ₂} {A : Set a} → {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} →
+ Reflexive _≈_ →
+ Transitive _<_ → Irreflexive _≈_ _<_ → Asymmetric _<_
trans∧irr⟶asym refl trans irrefl = λ x<y y<x →
irrefl refl (trans x<y y<x)
irr∧antisym⟶asym :
- ∀ {a ℓ₁ ℓ₂} {A : Set a} {≈ : Rel A ℓ₁} {< : Rel A ℓ₂} →
- Irreflexive ≈ < → Antisymmetric ≈ < → Asymmetric <
+ ∀ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} →
+ Irreflexive _≈_ _<_ → Antisymmetric _≈_ _<_ → Asymmetric _<_
irr∧antisym⟶asym irrefl antisym = λ x<y y<x →
irrefl (antisym x<y y<x) x<y
asym⟶antisym :
- ∀ {a ℓ₁ ℓ₂} {A : Set a} {≈ : Rel A ℓ₁} {< : Rel A ℓ₂} →
- Asymmetric < → Antisymmetric ≈ <
+ ∀ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} →
+ Asymmetric _<_ → Antisymmetric _≈_ _<_
asym⟶antisym asym x<y y<x = ⊥-elim (asym x<y y<x)
asym⟶irr :
- ∀ {a ℓ₁ ℓ₂} {A : Set a} {≈ : Rel A ℓ₁} {< : Rel A ℓ₂} →
- < Respects₂ ≈ → Symmetric ≈ →
- Asymmetric < → Irreflexive ≈ <
-asym⟶irr {< = _<_} resp sym asym {x} {y} x≈y x<y = asym x<y y<x
+ ∀ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} →
+ _<_ Respects₂ _≈_ → Symmetric _≈_ →
+ Asymmetric _<_ → Irreflexive _≈_ _<_
+asym⟶irr {_<_ = _<_} resp sym asym {x} {y} x≈y x<y = asym x<y y<x
where
y<y : y < y
y<y = proj₂ resp x≈y x<y
@@ -48,24 +48,24 @@ asym⟶irr {< = _<_} resp sym asym {x} {y} x≈y x<y = asym x<y y<x
y<x = proj₁ resp (sym x≈y) y<y
total⟶refl :
- ∀ {a ℓ₁ ℓ₂} {A : Set a} {≈ : Rel A ℓ₁} {∼ : Rel A ℓ₂} →
- ∼ Respects₂ ≈ → Symmetric ≈ →
- Total ∼ → ≈ ⇒ ∼
-total⟶refl {≈ = ≈} {∼ = ∼} resp sym total = refl
+ ∀ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_∼_ : Rel A ℓ₂} →
+ _∼_ Respects₂ _≈_ → Symmetric _≈_ →
+ Total _∼_ → _≈_ ⇒ _∼_
+total⟶refl {_≈_ = _≈_} {_∼_ = _∼_} resp sym total = refl
where
- refl : ≈ ⇒ ∼
+ refl : _≈_ ⇒ _∼_
refl {x} {y} x≈y with total x y
... | inj₁ x∼y = x∼y
... | inj₂ y∼x =
proj₁ resp x≈y (proj₂ resp (sym x≈y) y∼x)
total+dec⟶dec :
- ∀ {a ℓ₁ ℓ₂} {A : Set a} {≈ : Rel A ℓ₁} {≤ : Rel A ℓ₂} →
- ≈ ⇒ ≤ → Antisymmetric ≈ ≤ →
- Total ≤ → Decidable ≈ → Decidable ≤
-total+dec⟶dec {≈ = ≈} {≤ = ≤} refl antisym total _≟_ = dec
+ ∀ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_≤_ : Rel A ℓ₂} →
+ _≈_ ⇒ _≤_ → Antisymmetric _≈_ _≤_ →
+ Total _≤_ → Decidable _≈_ → Decidable _≤_
+total+dec⟶dec {_≈_ = _≈_} {_≤_ = _≤_} refl antisym total _≟_ = dec
where
- dec : Decidable ≤
+ dec : Decidable _≤_
dec x y with total x y
... | inj₁ x≤y = yes x≤y
... | inj₂ y≤x with x ≟ y
@@ -73,30 +73,30 @@ total+dec⟶dec {≈ = ≈} {≤ = ≤} refl antisym total _≟_ = dec
... | no ¬x≈y = no (λ x≤y → ¬x≈y (antisym x≤y y≤x))
tri⟶asym :
- ∀ {a ℓ₁ ℓ₂} {A : Set a} {≈ : Rel A ℓ₁} {< : Rel A ℓ₂} →
- Trichotomous ≈ < → Asymmetric <
+ ∀ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} →
+ Trichotomous _≈_ _<_ → Asymmetric _<_
tri⟶asym tri {x} {y} x<y x>y with tri x y
... | tri< _ _ x≯y = x≯y x>y
... | tri≈ _ _ x≯y = x≯y x>y
... | tri> x≮y _ _ = x≮y x<y
tri⟶irr :
- ∀ {a ℓ₁ ℓ₂} {A : Set a} {≈ : Rel A ℓ₁} {< : Rel A ℓ₂} →
- < Respects₂ ≈ → Symmetric ≈ →
- Trichotomous ≈ < → Irreflexive ≈ <
+ ∀ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} →
+ _<_ Respects₂ _≈_ → Symmetric _≈_ →
+ Trichotomous _≈_ _<_ → Irreflexive _≈_ _<_
tri⟶irr resp sym tri = asym⟶irr resp sym (tri⟶asym tri)
tri⟶dec≈ :
- ∀ {a ℓ₁ ℓ₂} {A : Set a} {≈ : Rel A ℓ₁} {< : Rel A ℓ₂} →
- Trichotomous ≈ < → Decidable ≈
+ ∀ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} →
+ Trichotomous _≈_ _<_ → Decidable _≈_
tri⟶dec≈ compare x y with compare x y
... | tri< _ x≉y _ = no x≉y
... | tri≈ _ x≈y _ = yes x≈y
... | tri> _ x≉y _ = no x≉y
tri⟶dec< :
- ∀ {a ℓ₁ ℓ₂} {A : Set a} {≈ : Rel A ℓ₁} {< : Rel A ℓ₂} →
- Trichotomous ≈ < → Decidable <
+ ∀ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} →
+ Trichotomous _≈_ _<_ → Decidable _<_
tri⟶dec< compare x y with compare x y
... | tri< x<y _ _ = yes x<y
... | tri≈ x≮y _ _ = no x≮y
diff --git a/src/Relation/Binary/Core.agda b/src/Relation/Binary/Core.agda
index bb3bc50..472ca58 100644
--- a/src/Relation/Binary/Core.agda
+++ b/src/Relation/Binary/Core.agda
@@ -11,12 +11,12 @@ module Relation.Binary.Core where
open import Data.Product
open import Data.Sum
-open import Data.Function
+open import Function
open import Level
-open import Relation.Nullary.Core
+open import Relation.Nullary
------------------------------------------------------------------------
--- Binary relations.
+-- Binary relations
-- Heterogeneous binary relations
@@ -86,6 +86,12 @@ Trans : ∀ {a b c ℓ₁ ℓ₂ ℓ₃} {A : Set a} {B : Set b} {C : Set c} →
REL A B ℓ₁ → REL B C ℓ₂ → REL A C ℓ₃ → Set _
Trans P Q R = ∀ {i j k} → P i j → Q j k → R i k
+-- A variant of Trans.
+
+TransFlip : ∀ {a b c ℓ₁ ℓ₂ ℓ₃} {A : Set a} {B : Set b} {C : Set c} →
+ REL A B ℓ₁ → REL B C ℓ₂ → REL A C ℓ₃ → Set _
+TransFlip P Q R = ∀ {i j k} → Q j k → P i j → R i k
+
Transitive : ∀ {a ℓ} {A : Set a} → Rel A ℓ → Set _
Transitive _∼_ = Trans _∼_ _∼_ _∼_
@@ -142,7 +148,7 @@ private
infix 4 _≡_ _≢_
- data _≡_ {a} {A : Set a} (x : A) : A → Set where
+ data _≡_ {a} {A : Set a} (x : A) : A → Set a where
refl : x ≡ x
{-# BUILTIN EQUALITY _≡_ #-}
@@ -150,7 +156,7 @@ private
-- Nonequality.
- _≢_ : ∀ {a} {A : Set a} → A → A → Set
+ _≢_ : ∀ {a} {A : Set a} → A → A → Set a
x ≢ y = ¬ x ≡ y
------------------------------------------------------------------------
diff --git a/src/Relation/Binary/Flip.agda b/src/Relation/Binary/Flip.agda
index 72d8d53..74b1300 100644
--- a/src/Relation/Binary/Flip.agda
+++ b/src/Relation/Binary/Flip.agda
@@ -8,7 +8,7 @@ open import Relation.Binary
module Relation.Binary.Flip where
-open import Data.Function
+open import Function
open import Data.Product
implies : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b}
@@ -83,7 +83,6 @@ isPreorder {≈ = ≈} {∼} pre = record
{ isEquivalence = isEquivalence Pre.isEquivalence
; reflexive = implies ≈ ∼ Pre.reflexive
; trans = transitive ∼ Pre.trans
- ; ∼-resp-≈ = respects₂ ∼ ≈ Pre.Eq.sym Pre.∼-resp-≈
}
where module Pre = IsPreorder pre
diff --git a/src/Relation/Binary/HeterogeneousEquality.agda b/src/Relation/Binary/HeterogeneousEquality.agda
index 3c3420f..b82e3ce 100644
--- a/src/Relation/Binary/HeterogeneousEquality.agda
+++ b/src/Relation/Binary/HeterogeneousEquality.agda
@@ -6,12 +6,14 @@
module Relation.Binary.HeterogeneousEquality where
-open import Data.Function
open import Data.Product
+open import Function
+open import Function.Inverse using (Inverse)
open import Level
open import Relation.Nullary
open import Relation.Binary
open import Relation.Binary.Consequences
+open import Relation.Binary.Indexed as I using (_at_)
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; refl)
@@ -99,6 +101,28 @@ setoid A = record
; isEquivalence = isEquivalence
}
+indexedSetoid : ∀ {a b} {A : Set a} → (A → Set b) → I.Setoid A _ _
+indexedSetoid B = record
+ { Carrier = B
+ ; _≈_ = λ x y → x ≅ y
+ ; isEquivalence = record
+ { refl = refl
+ ; sym = sym
+ ; trans = trans
+ }
+ }
+
+≡⇿≅ : ∀ {a b} {A : Set a} (B : A → Set b) {x : A} →
+ Inverse (PropEq.setoid (B x)) (indexedSetoid B at x)
+≡⇿≅ B = record
+ { to = record { _⟨$⟩_ = id; cong = ≡-to-≅ }
+ ; from = record { _⟨$⟩_ = id; cong = ≅-to-≡ }
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+
decSetoid : ∀ {a} {A : Set a} →
Decidable {A = A} {B = A} (λ x y → x ≅ y) →
DecSetoid _ _
@@ -116,7 +140,6 @@ isPreorder = record
{ isEquivalence = isEquivalence
; reflexive = id
; trans = trans
- ; ∼-resp-≈ = resp₂ (λ x y → x ≅ y)
}
isPreorder-≡ : ∀ {a} {A : Set a} →
@@ -125,7 +148,6 @@ isPreorder-≡ = record
{ isEquivalence = PropEq.isEquivalence
; reflexive = reflexive
; trans = trans
- ; ∼-resp-≈ = PropEq.resp₂ (λ x y → x ≅ y)
}
preorder : ∀ {a} → Set a → Preorder _ _ _
diff --git a/src/Relation/Binary/Indexed.agda b/src/Relation/Binary/Indexed.agda
new file mode 100644
index 0000000..fdb750a
--- /dev/null
+++ b/src/Relation/Binary/Indexed.agda
@@ -0,0 +1,35 @@
+------------------------------------------------------------------------
+-- Indexed binary relations
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+module Relation.Binary.Indexed where
+
+import Relation.Binary as B
+
+open import Relation.Binary.Indexed.Core public
+
+-- By "instantiating" an indexed setoid one gets an ordinary setoid.
+
+_at_ : ∀ {i s₁ s₂} {I : Set i} → Setoid I s₁ s₂ → I → B.Setoid _ _
+S at i = record
+ { Carrier = S.Carrier i
+ ; _≈_ = S._≈_
+ ; isEquivalence = record
+ { refl = S.refl
+ ; sym = S.sym
+ ; trans = S.trans
+ }
+ } where module S = Setoid S
+
+------------------------------------------------------------------------
+-- Simple properties of indexed binary relations
+
+-- Generalised implication.
+
+infixr 4 _=[_]⇒_
+
+_=[_]⇒_ : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : A → Set b} →
+ B.Rel A ℓ₁ → ((x : A) → B x) → Rel B ℓ₂ → Set _
+P =[ f ]⇒ Q = ∀ {i j} → P i j → Q (f i) (f j)
diff --git a/src/Relation/Binary/Indexed/Core.agda b/src/Relation/Binary/Indexed/Core.agda
new file mode 100644
index 0000000..ecc3b03
--- /dev/null
+++ b/src/Relation/Binary/Indexed/Core.agda
@@ -0,0 +1,69 @@
+------------------------------------------------------------------------
+-- Indexed binary relations
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+-- This file contains some core definitions which are reexported by
+-- Relation.Binary.Indexed.
+
+module Relation.Binary.Indexed.Core where
+
+open import Function
+open import Level
+import Relation.Binary.Core as B
+import Relation.Binary.Core as P
+
+------------------------------------------------------------------------
+-- Indexed binary relations
+
+-- Heterogeneous.
+
+REL : ∀ {i₁ i₂ a₁ a₂} {I₁ : Set i₁} {I₂ : Set i₂} →
+ (I₁ → Set a₁) → (I₂ → Set a₂) → (ℓ : Level) → Set _
+REL A₁ A₂ ℓ = ∀ {i₁ i₂} → A₁ i₁ → A₂ i₂ → Set ℓ
+
+-- Homogeneous.
+
+Rel : ∀ {i a} {I : Set i} → (I → Set a) → (ℓ : Level) → Set _
+Rel A ℓ = REL A A ℓ
+
+------------------------------------------------------------------------
+-- Simple properties of indexed binary relations
+
+-- Reflexivity.
+
+Reflexive : ∀ {i a ℓ} {I : Set i} (A : I → Set a) → Rel A ℓ → Set _
+Reflexive _ _∼_ = ∀ {i} → B.Reflexive (_∼_ {i})
+
+-- Symmetry.
+
+Symmetric : ∀ {i a ℓ} {I : Set i} (A : I → Set a) → Rel A ℓ → Set _
+Symmetric _ _∼_ = ∀ {i j} → B.Sym (_∼_ {i} {j}) _∼_
+
+-- Transitivity.
+
+Transitive : ∀ {i a ℓ} {I : Set i} (A : I → Set a) → Rel A ℓ → Set _
+Transitive _ _∼_ = ∀ {i j k} → B.Trans _∼_ (_∼_ {j}) (_∼_ {i} {k})
+
+------------------------------------------------------------------------
+-- Setoids
+
+record IsEquivalence {i a ℓ} {I : Set i} (A : I → Set a)
+ (_≈_ : Rel A ℓ) : Set (i ⊔ a ⊔ ℓ) where
+ field
+ refl : Reflexive A _≈_
+ sym : Symmetric A _≈_
+ trans : Transitive A _≈_
+
+ reflexive : ∀ {i} → P._≡_ ⟨ B._⇒_ ⟩ _≈_ {i}
+ reflexive P.refl = refl
+
+record Setoid {i} (I : Set i) c ℓ : Set (suc (i ⊔ c ⊔ ℓ)) where
+ infix 4 _≈_
+ field
+ Carrier : I → Set c
+ _≈_ : Rel Carrier ℓ
+ isEquivalence : IsEquivalence Carrier _≈_
+
+ open IsEquivalence isEquivalence public
diff --git a/src/Relation/Binary/InducedPreorders.agda b/src/Relation/Binary/InducedPreorders.agda
index 454250b..057421e 100644
--- a/src/Relation/Binary/InducedPreorders.agda
+++ b/src/Relation/Binary/InducedPreorders.agda
@@ -11,7 +11,7 @@ module Relation.Binary.InducedPreorders
(S : Setoid s₁ s₂) -- The underlying equality.
where
-open import Data.Function
+open import Function
open import Data.Product
open Setoid S
@@ -28,8 +28,6 @@ InducedPreorder₁ P resp = record
{ isEquivalence = isEquivalence
; reflexive = resp
; trans = flip _∘′_
- ; ∼-resp-≈ = (λ c₁≈c₂ c₃∼c₁ → resp c₁≈c₂ ∘ c₃∼c₁)
- , (λ c₁≈c₂ c₁∼c₃ → c₁∼c₃ ∘ resp (sym c₁≈c₂))
}
}
@@ -46,7 +44,5 @@ InducedPreorder₂ _R_ resp = record
{ isEquivalence = isEquivalence
; reflexive = λ c₁≈c₂ → resp c₁≈c₂
; trans = λ c₁∼c₂ c₂∼c₃ → c₂∼c₃ ∘ c₁∼c₂
- ; ∼-resp-≈ = (λ c₁≈c₂ c₃∼c₁ → resp c₁≈c₂ ∘ c₃∼c₁)
- , (λ c₁≈c₂ c₁∼c₃ → c₁∼c₃ ∘ resp (sym c₁≈c₂))
}
}
diff --git a/src/Relation/Binary/List/NonStrictLex.agda b/src/Relation/Binary/List/NonStrictLex.agda
index 933d88f..4ba9842 100644
--- a/src/Relation/Binary/List/NonStrictLex.agda
+++ b/src/Relation/Binary/List/NonStrictLex.agda
@@ -10,7 +10,7 @@
module Relation.Binary.List.NonStrictLex where
open import Data.Empty
-open import Data.Function
+open import Function
open import Data.Unit using (⊤; tt)
open import Data.Product
open import Data.List
@@ -26,31 +26,32 @@ open import Relation.Binary.List.Pointwise as Pointwise using ([])
private
module Dummy {A : Set} where
- Lex : (P : Set) → (≈ ≤ : Rel A zero) → Rel (List A) zero
- Lex P ≈ ≤ = Strict.Lex P ≈ (Conv._<_ ≈ ≤)
+ Lex : (P : Set) → (_≈_ _≤_ : Rel A zero) → Rel (List A) zero
+ Lex P _≈_ _≤_ = Strict.Lex P _≈_ (Conv._<_ _≈_ _≤_)
-- Strict lexicographic ordering.
- Lex-< : (≈ ≤ : Rel A zero) → Rel (List A) zero
+ Lex-< : (_≈_ _≤_ : Rel A zero) → Rel (List A) zero
Lex-< = Lex ⊥
-- Non-strict lexicographic ordering.
- Lex-≤ : (≈ ≤ : Rel A zero) → Rel (List A) zero
+ Lex-≤ : (_≈_ _≤_ : Rel A zero) → Rel (List A) zero
Lex-≤ = Lex ⊤
------------------------------------------------------------------------
-- Properties
- ≤-reflexive : ∀ ≈ ≤ → Pointwise.Rel ≈ ⇒ Lex-≤ ≈ ≤
- ≤-reflexive ≈ ≤ = Strict.≤-reflexive ≈ (Conv._<_ ≈ ≤)
+ ≤-reflexive : ∀ _≈_ _≤_ → Pointwise.Rel _≈_ ⇒ Lex-≤ _≈_ _≤_
+ ≤-reflexive _≈_ _≤_ = Strict.≤-reflexive _≈_ (Conv._<_ _≈_ _≤_)
- <-irreflexive : ∀ {≈ ≤} → Irreflexive (Pointwise.Rel ≈) (Lex-< ≈ ≤)
- <-irreflexive {≈} {≤} =
- Strict.<-irreflexive {< = Conv._<_ ≈ ≤} (Conv.irrefl ≈ ≤)
+ <-irreflexive : ∀ {_≈_ _≤_} →
+ Irreflexive (Pointwise.Rel _≈_) (Lex-< _≈_ _≤_)
+ <-irreflexive {_≈_} {_≤_} =
+ Strict.<-irreflexive {_<_ = Conv._<_ _≈_ _≤_} (Conv.irrefl _≈_ _≤_)
- transitive : ∀ {P ≈ ≤} →
- IsPartialOrder ≈ ≤ → Transitive (Lex P ≈ ≤)
+ transitive : ∀ {P _≈_ _≤_} →
+ IsPartialOrder _≈_ _≤_ → Transitive (Lex P _≈_ _≤_)
transitive po =
Strict.transitive
isEquivalence
@@ -58,95 +59,98 @@ private
(Conv.trans _ _ po)
where open IsPartialOrder po
- antisymmetric : ∀ {P ≈ ≤} →
- Symmetric ≈ → Antisymmetric ≈ ≤ →
- Antisymmetric (Pointwise.Rel ≈) (Lex P ≈ ≤)
- antisymmetric {≈ = ≈} {≤} sym antisym =
- Strict.antisymmetric sym (Conv.irrefl ≈ ≤)
- (Conv.antisym⟶asym ≈ _ antisym)
-
- <-asymmetric : ∀ {≈ ≤} →
- IsEquivalence ≈ → ≤ Respects₂ ≈ →
- Antisymmetric ≈ ≤ → Asymmetric (Lex-< ≈ ≤)
- <-asymmetric {≈} eq resp antisym =
+ antisymmetric : ∀ {P _≈_ _≤_} →
+ Symmetric _≈_ → Antisymmetric _≈_ _≤_ →
+ Antisymmetric (Pointwise.Rel _≈_) (Lex P _≈_ _≤_)
+ antisymmetric {_≈_ = _≈_} {_≤_} sym antisym =
+ Strict.antisymmetric sym (Conv.irrefl _≈_ _≤_)
+ (Conv.antisym⟶asym _≈_ _ antisym)
+
+ <-asymmetric : ∀ {_≈_ _≤_} →
+ IsEquivalence _≈_ → _≤_ Respects₂ _≈_ →
+ Antisymmetric _≈_ _≤_ → Asymmetric (Lex-< _≈_ _≤_)
+ <-asymmetric {_≈_} eq resp antisym =
Strict.<-asymmetric sym (Conv.<-resp-≈ _ _ eq resp)
- (Conv.antisym⟶asym ≈ _ antisym)
+ (Conv.antisym⟶asym _≈_ _ antisym)
where open IsEquivalence eq
- respects₂ : ∀ {P ≈ ≤} →
- IsEquivalence ≈ → ≤ Respects₂ ≈ →
- Lex P ≈ ≤ Respects₂ Pointwise.Rel ≈
+ respects₂ : ∀ {P _≈_ _≤_} →
+ IsEquivalence _≈_ → _≤_ Respects₂ _≈_ →
+ Lex P _≈_ _≤_ Respects₂ Pointwise.Rel _≈_
respects₂ eq resp = Strict.respects₂ eq (Conv.<-resp-≈ _ _ eq resp)
- decidable : ∀ {P ≈ ≤} →
- Dec P → Decidable ≈ → Decidable ≤ →
- Decidable (Lex P ≈ ≤)
+ decidable : ∀ {P _≈_ _≤_} →
+ Dec P → Decidable _≈_ → Decidable _≤_ →
+ Decidable (Lex P _≈_ _≤_)
decidable dec-P dec-≈ dec-≤ =
Strict.decidable dec-P dec-≈ (Conv.decidable _ _ dec-≈ dec-≤)
- <-decidable : ∀ {≈ ≤} →
- Decidable ≈ → Decidable ≤ → Decidable (Lex-< ≈ ≤)
+ <-decidable :
+ ∀ {_≈_ _≤_} →
+ Decidable _≈_ → Decidable _≤_ → Decidable (Lex-< _≈_ _≤_)
<-decidable = decidable (no id)
- ≤-decidable : ∀ {≈ ≤} →
- Decidable ≈ → Decidable ≤ → Decidable (Lex-≤ ≈ ≤)
+ ≤-decidable :
+ ∀ {_≈_ _≤_} →
+ Decidable _≈_ → Decidable _≤_ → Decidable (Lex-≤ _≈_ _≤_)
≤-decidable = decidable (yes tt)
- ≤-total : ∀ {≈ ≤} → Symmetric ≈ → Decidable ≈ →
- Antisymmetric ≈ ≤ → Total ≤ →
- Total (Lex-≤ ≈ ≤)
+ ≤-total : ∀ {_≈_ _≤_} → Symmetric _≈_ → Decidable _≈_ →
+ Antisymmetric _≈_ _≤_ → Total _≤_ →
+ Total (Lex-≤ _≈_ _≤_)
≤-total sym dec-≈ antisym tot =
Strict.≤-total sym (Conv.trichotomous _ _ sym dec-≈ antisym tot)
- <-compare : ∀ {≈ ≤} → Symmetric ≈ → Decidable ≈ →
- Antisymmetric ≈ ≤ → Total ≤ →
- Trichotomous (Pointwise.Rel ≈) (Lex-< ≈ ≤)
+ <-compare : ∀ {_≈_ _≤_} → Symmetric _≈_ → Decidable _≈_ →
+ Antisymmetric _≈_ _≤_ → Total _≤_ →
+ Trichotomous (Pointwise.Rel _≈_) (Lex-< _≈_ _≤_)
<-compare sym dec-≈ antisym tot =
Strict.<-compare sym (Conv.trichotomous _ _ sym dec-≈ antisym tot)
-- Some collections of properties which are preserved by Lex-≤ or
-- Lex-<.
- ≤-isPreorder : ∀ {≈ ≤} →
- IsPartialOrder ≈ ≤ →
- IsPreorder (Pointwise.Rel ≈) (Lex-≤ ≈ ≤)
+ ≤-isPreorder : ∀ {_≈_ _≤_} →
+ IsPartialOrder _≈_ _≤_ →
+ IsPreorder (Pointwise.Rel _≈_) (Lex-≤ _≈_ _≤_)
≤-isPreorder po =
Strict.≤-isPreorder
isEquivalence (Conv.trans _ _ po)
(Conv.<-resp-≈ _ _ isEquivalence ≤-resp-≈)
where open IsPartialOrder po
- ≤-isPartialOrder : ∀ {≈ ≤} →
- IsPartialOrder ≈ ≤ →
- IsPartialOrder (Pointwise.Rel ≈) (Lex-≤ ≈ ≤)
+ ≤-isPartialOrder : ∀ {_≈_ _≤_} →
+ IsPartialOrder _≈_ _≤_ →
+ IsPartialOrder (Pointwise.Rel _≈_) (Lex-≤ _≈_ _≤_)
≤-isPartialOrder po =
Strict.≤-isPartialOrder
(Conv.isPartialOrder⟶isStrictPartialOrder _ _ po)
- ≤-isTotalOrder : ∀ {≈ ≤} →
- Decidable ≈ → IsTotalOrder ≈ ≤ →
- IsTotalOrder (Pointwise.Rel ≈) (Lex-≤ ≈ ≤)
+ ≤-isTotalOrder : ∀ {_≈_ _≤_} →
+ Decidable _≈_ → IsTotalOrder _≈_ _≤_ →
+ IsTotalOrder (Pointwise.Rel _≈_) (Lex-≤ _≈_ _≤_)
≤-isTotalOrder dec tot =
Strict.≤-isTotalOrder
(Conv.isTotalOrder⟶isStrictTotalOrder _ _ dec tot)
- ≤-isDecTotalOrder : ∀ {≈ ≤} →
- IsDecTotalOrder ≈ ≤ →
- IsDecTotalOrder (Pointwise.Rel ≈) (Lex-≤ ≈ ≤)
+ ≤-isDecTotalOrder :
+ ∀ {_≈_ _≤_} →
+ IsDecTotalOrder _≈_ _≤_ →
+ IsDecTotalOrder (Pointwise.Rel _≈_) (Lex-≤ _≈_ _≤_)
≤-isDecTotalOrder dtot =
Strict.≤-isDecTotalOrder
(Conv.isDecTotalOrder⟶isStrictTotalOrder _ _ dtot)
- <-isStrictPartialOrder : ∀ {≈ ≤} →
- IsPartialOrder ≈ ≤ →
- IsStrictPartialOrder (Pointwise.Rel ≈) (Lex-< ≈ ≤)
+ <-isStrictPartialOrder : ∀ {_≈_ _≤_} →
+ IsPartialOrder _≈_ _≤_ →
+ IsStrictPartialOrder (Pointwise.Rel _≈_) (Lex-< _≈_ _≤_)
<-isStrictPartialOrder po =
Strict.<-isStrictPartialOrder
(Conv.isPartialOrder⟶isStrictPartialOrder _ _ po)
- <-isStrictTotalOrder : ∀ {≈ ≤} →
- Decidable ≈ → IsTotalOrder ≈ ≤ →
- IsStrictTotalOrder (Pointwise.Rel ≈) (Lex-< ≈ ≤)
+ <-isStrictTotalOrder : ∀ {_≈_ _≤_} →
+ Decidable _≈_ → IsTotalOrder _≈_ _≤_ →
+ IsStrictTotalOrder (Pointwise.Rel _≈_) (Lex-< _≈_ _≤_)
<-isStrictTotalOrder dec tot =
Strict.<-isStrictTotalOrder
(Conv.isTotalOrder⟶isStrictTotalOrder _ _ dec tot)
diff --git a/src/Relation/Binary/List/Pointwise.agda b/src/Relation/Binary/List/Pointwise.agda
index 4fdeadd..d0314f0 100644
--- a/src/Relation/Binary/List/Pointwise.agda
+++ b/src/Relation/Binary/List/Pointwise.agda
@@ -4,7 +4,7 @@
module Relation.Binary.List.Pointwise where
-open import Data.Function
+open import Function
open import Data.Product
open import Data.List
open import Level
@@ -28,45 +28,46 @@ private
tail : ∀ {_∼_ x y xs ys} → Rel _∼_ (x ∷ xs) (y ∷ ys) → Rel _∼_ xs ys
tail (x∼y ∷ xs∼ys) = xs∼ys
- reflexive : ∀ {≈ ∼} → ≈ ⇒ ∼ → (Rel ≈) ⇒ (Rel ∼)
+ reflexive : ∀ {_≈_ _∼_} → _≈_ ⇒ _∼_ → Rel _≈_ ⇒ Rel _∼_
reflexive ≈⇒∼ [] = []
reflexive ≈⇒∼ (x≈y ∷ xs≈ys) = ≈⇒∼ x≈y ∷ reflexive ≈⇒∼ xs≈ys
- refl : ∀ {∼} → Reflexive ∼ → Reflexive (Rel ∼)
+ refl : ∀ {_∼_} → Reflexive _∼_ → Reflexive (Rel _∼_)
refl rfl {[]} = []
refl rfl {x ∷ xs} = rfl ∷ refl rfl
- symmetric : ∀ {∼} → Symmetric ∼ → Symmetric (Rel ∼)
+ symmetric : ∀ {_∼_} → Symmetric _∼_ → Symmetric (Rel _∼_)
symmetric sym [] = []
symmetric sym (x∼y ∷ xs∼ys) = sym x∼y ∷ symmetric sym xs∼ys
- transitive : ∀ {∼} → Transitive ∼ → Transitive (Rel ∼)
+ transitive : ∀ {_∼_} → 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 ≈) (Rel ≤)
+ antisymmetric : ∀ {_≈_ _≤_} → 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 ∼) Respects₂ (Rel ≈)
- respects₂ {≈} {∼} resp =
+ respects₂ : ∀ {_≈_ _∼_} →
+ _∼_ Respects₂ _≈_ → (Rel _∼_) Respects₂ (Rel _≈_)
+ respects₂ {_≈_} {_∼_} resp =
(λ {xs} {ys} {zs} → resp¹ {xs} {ys} {zs}) ,
(λ {xs} {ys} {zs} → resp² {xs} {ys} {zs})
where
- resp¹ : ∀ {xs} → (Rel ∼ xs) Respects (Rel ≈)
+ resp¹ : ∀ {xs} → (Rel _∼_ xs) Respects (Rel _≈_)
resp¹ [] [] = []
resp¹ (x≈y ∷ xs≈ys) (z∼x ∷ zs∼xs) =
proj₁ resp x≈y z∼x ∷ resp¹ xs≈ys zs∼xs
- resp² : ∀ {ys} → (flip (Rel ∼) ys) Respects (Rel ≈)
+ resp² : ∀ {ys} → (flip (Rel _∼_) ys) Respects (Rel _≈_)
resp² [] [] = []
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 : ∀ {_∼_} → Decidable _∼_ → Decidable (Rel _∼_)
decidable dec [] [] = yes []
decidable dec [] (y ∷ ys) = no (λ ())
decidable dec (x ∷ xs) [] = no (λ ())
@@ -76,30 +77,30 @@ private
... | no ¬xs∼ys = no (¬xs∼ys ∘ tail)
... | yes xs∼ys = yes (x∼y ∷ xs∼ys)
- isEquivalence : ∀ {≈} → IsEquivalence ≈ → IsEquivalence (Rel ≈)
+ isEquivalence : ∀ {_≈_} → 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 ≈ ∼ → IsPreorder (Rel ≈) (Rel ∼)
+ isPreorder : ∀ {_≈_ _∼_} →
+ IsPreorder _≈_ _∼_ → IsPreorder (Rel _≈_) (Rel _∼_)
isPreorder pre = record
{ isEquivalence = isEquivalence Pre.isEquivalence
; reflexive = reflexive Pre.reflexive
; trans = transitive Pre.trans
- ; ∼-resp-≈ = respects₂ Pre.∼-resp-≈
} where module Pre = IsPreorder pre
- isDecEquivalence : ∀ {≈} → IsDecEquivalence ≈ →
- IsDecEquivalence (Rel ≈)
+ isDecEquivalence : ∀ {_≈_} → IsDecEquivalence _≈_ →
+ IsDecEquivalence (Rel _≈_)
isDecEquivalence eq = record
{ isEquivalence = isEquivalence Dec.isEquivalence
; _≟_ = decidable Dec._≟_
} where module Dec = IsDecEquivalence eq
- isPartialOrder : ∀ {≈ ≤} → IsPartialOrder ≈ ≤ →
- IsPartialOrder (Rel ≈) (Rel ≤)
+ isPartialOrder : ∀ {_≈_ _≤_} → IsPartialOrder _≈_ _≤_ →
+ IsPartialOrder (Rel _≈_) (Rel _≤_)
isPartialOrder po = record
{ isPreorder = isPreorder PO.isPreorder
; antisym = antisymmetric PO.antisym
diff --git a/src/Relation/Binary/List/StrictLex.agda b/src/Relation/Binary/List/StrictLex.agda
index 422a5d4..7530895 100644
--- a/src/Relation/Binary/List/StrictLex.agda
+++ b/src/Relation/Binary/List/StrictLex.agda
@@ -11,7 +11,7 @@ module Relation.Binary.List.StrictLex where
open import Data.Empty
open import Data.Unit using (⊤; tt)
-open import Data.Function
+open import Function
open import Data.Product
open import Data.Sum
open import Data.List
@@ -25,58 +25,60 @@ open import Relation.Binary.List.Pointwise as Pointwise
private
module Dummy {A : Set} where
- data Lex (P : Set) (≈ < : Rel A zero) : Rel (List A) zero where
- base : P → Lex P ≈ < [] []
- halt : ∀ {y ys} → Lex P ≈ < [] (y ∷ ys)
- this : ∀ {x xs y ys} (x<y : < x y) → Lex P ≈ < (x ∷ xs) (y ∷ ys)
- next : ∀ {x xs y ys} (x≈y : ≈ x y)
- (xs⊴ys : Lex P ≈ < xs ys) → Lex P ≈ < (x ∷ xs) (y ∷ ys)
+ data Lex (P : Set) (_≈_ _<_ : Rel A zero) : Rel (List A) zero where
+ base : P → Lex P _≈_ _<_ [] []
+ halt : ∀ {y ys} → Lex P _≈_ _<_ [] (y ∷ ys)
+ this : ∀ {x xs y ys} (x<y : x < y) → Lex P _≈_ _<_ (x ∷ xs) (y ∷ ys)
+ next : ∀ {x xs y ys} (x≈y : x ≈ y)
+ (xs⊴ys : Lex P _≈_ _<_ xs ys) → Lex P _≈_ _<_ (x ∷ xs) (y ∷ ys)
-- Strict lexicographic ordering.
- Lex-< : (≈ < : Rel A zero) → Rel (List A) zero
+ Lex-< : (_≈_ _<_ : Rel A zero) → Rel (List A) zero
Lex-< = Lex ⊥
- ¬[]<[] : ∀ {≈ <} → ¬ Lex-< ≈ < [] []
+ ¬[]<[] : ∀ {_≈_ _<_} → ¬ Lex-< _≈_ _<_ [] []
¬[]<[] (base ())
-- Non-strict lexicographic ordering.
- Lex-≤ : (≈ < : Rel A zero) → Rel (List A) zero
+ Lex-≤ : (_≈_ _<_ : Rel A zero) → Rel (List A) zero
Lex-≤ = Lex ⊤
-- Utilities.
- ¬≤-this : ∀ {P ≈ < x y xs ys} → ¬ ≈ x y → ¬ < x y →
- ¬ Lex P ≈ < (x ∷ xs) (y ∷ ys)
+ ¬≤-this : ∀ {P _≈_ _<_ x y xs ys} → ¬ (x ≈ y) → ¬ (x < y) →
+ ¬ Lex P _≈_ _<_ (x ∷ xs) (y ∷ ys)
¬≤-this ¬x≈y ¬x<y (this x<y) = ¬x<y x<y
¬≤-this ¬x≈y ¬x<y (next x≈y xs⊴ys) = ¬x≈y x≈y
- ¬≤-next : ∀ {P ≈ < x y xs ys} → ¬ < x y → ¬ Lex P ≈ < xs ys →
- ¬ Lex P ≈ < (x ∷ xs) (y ∷ ys)
+ ¬≤-next : ∀ {P _≈_ _<_ x y xs ys} →
+ ¬ (x < y) → ¬ Lex P _≈_ _<_ xs ys →
+ ¬ Lex P _≈_ _<_ (x ∷ xs) (y ∷ ys)
¬≤-next ¬x<y ¬xs⊴ys (this x<y) = ¬x<y x<y
¬≤-next ¬x<y ¬xs⊴ys (next x≈y xs⊴ys) = ¬xs⊴ys xs⊴ys
----------------------------------------------------------------------
-- Properties
- ≤-reflexive : ∀ ≈ < → Pointwise.Rel ≈ ⇒ Lex-≤ ≈ <
- ≤-reflexive ≈ < [] = base tt
- ≤-reflexive ≈ < (x≈y ∷ xs≈ys) = next x≈y (≤-reflexive ≈ < xs≈ys)
+ ≤-reflexive : ∀ _≈_ _<_ → Pointwise.Rel _≈_ ⇒ Lex-≤ _≈_ _<_
+ ≤-reflexive _≈_ _<_ [] = base tt
+ ≤-reflexive _≈_ _<_ (x≈y ∷ xs≈ys) =
+ next x≈y (≤-reflexive _≈_ _<_ xs≈ys)
- <-irreflexive : ∀ {≈ <} → Irreflexive ≈ < →
- Irreflexive (Pointwise.Rel ≈) (Lex-< ≈ <)
+ <-irreflexive : ∀ {_≈_ _<_} → Irreflexive _≈_ _<_ →
+ Irreflexive (Pointwise.Rel _≈_) (Lex-< _≈_ _<_)
<-irreflexive irr [] (base ())
<-irreflexive irr (x≈y ∷ xs≈ys) (this x<y) = irr x≈y x<y
<-irreflexive irr (x≈y ∷ xs≈ys) (next x≊y xs⊴ys) =
<-irreflexive irr xs≈ys xs⊴ys
- transitive : ∀ {P ≈ <} →
- IsEquivalence ≈ → < Respects₂ ≈ → Transitive < →
- Transitive (Lex P ≈ <)
- transitive {P} {≈} {<} eq resp tr = trans
+ transitive : ∀ {P _≈_ _<_} →
+ IsEquivalence _≈_ → _<_ Respects₂ _≈_ → Transitive _<_ →
+ Transitive (Lex P _≈_ _<_)
+ transitive {P} {_≈_} {_<_} eq resp tr = trans
where
- trans : Transitive (Lex P ≈ <)
+ trans : Transitive (Lex P _≈_ _<_)
trans (base p) (base _) = base p
trans (base y) halt = halt
trans halt (this y<z) = halt
@@ -88,12 +90,13 @@ private
trans (next x≈y xs⊴ys) (next y≈z ys⊴zs) =
next (IsEquivalence.trans eq x≈y y≈z) (trans xs⊴ys ys⊴zs)
- antisymmetric : ∀ {P ≈ <} →
- Symmetric ≈ → Irreflexive ≈ < → Asymmetric < →
- Antisymmetric (Pointwise.Rel ≈) (Lex P ≈ <)
- antisymmetric {P} {≈} {<} sym ir asym = as
+ antisymmetric :
+ ∀ {P _≈_ _<_} →
+ Symmetric _≈_ → Irreflexive _≈_ _<_ → Asymmetric _<_ →
+ Antisymmetric (Pointwise.Rel _≈_) (Lex P _≈_ _<_)
+ antisymmetric {P} {_≈_} {_<_} sym ir asym = as
where
- as : Antisymmetric (Pointwise.Rel ≈) (Lex P ≈ <)
+ as : Antisymmetric (Pointwise.Rel _≈_) (Lex P _≈_ _<_)
as (base _) (base _) = []
as halt ()
as (this x<y) (this y<x) = ⊥-elim (asym x<y y<x)
@@ -101,15 +104,15 @@ private
as (next x≈y xs⊴ys) (this y<x) = ⊥-elim (ir (sym x≈y) y<x)
as (next x≈y xs⊴ys) (next y≈x ys⊴xs) = x≈y ∷ as xs⊴ys ys⊴xs
- <-asymmetric : ∀ {≈ <} →
- Symmetric ≈ → < Respects₂ ≈ → Asymmetric < →
- Asymmetric (Lex-< ≈ <)
- <-asymmetric {≈} {<} sym resp as = asym
+ <-asymmetric : ∀ {_≈_ _<_} →
+ Symmetric _≈_ → _<_ Respects₂ _≈_ → Asymmetric _<_ →
+ Asymmetric (Lex-< _≈_ _<_)
+ <-asymmetric {_≈_} {_<_} sym resp as = asym
where
- irrefl : Irreflexive ≈ <
+ irrefl : Irreflexive _≈_ _<_
irrefl = asym⟶irr resp sym as
- asym : Asymmetric (Lex-< ≈ <)
+ asym : Asymmetric (Lex-< _≈_ _<_)
asym (base bot) _ = bot
asym halt ()
asym (this x<y) (this y<x) = as x<y y<x
@@ -117,14 +120,14 @@ private
asym (next x≈y xs⊴ys) (this y<x) = irrefl (sym x≈y) y<x
asym (next x≈y xs⊴ys) (next y≈x ys⊴xs) = asym xs⊴ys ys⊴xs
- respects₂ : ∀ {P ≈ <} →
- IsEquivalence ≈ → < Respects₂ ≈ →
- Lex P ≈ < Respects₂ Pointwise.Rel ≈
- respects₂ {P} {≈} {<} eq resp =
+ respects₂ : ∀ {P _≈_ _<_} →
+ IsEquivalence _≈_ → _<_ Respects₂ _≈_ →
+ Lex P _≈_ _<_ Respects₂ Pointwise.Rel _≈_
+ respects₂ {P} {_≈_} {_<_} eq resp =
(λ {xs} {ys} {zs} → resp¹ {xs} {ys} {zs}) ,
(λ {xs} {ys} {zs} → resp² {xs} {ys} {zs})
where
- resp¹ : ∀ {xs} → Lex P ≈ < xs Respects Pointwise.Rel ≈
+ resp¹ : ∀ {xs} → Lex P _≈_ _<_ xs Respects Pointwise.Rel _≈_
resp¹ [] xs⊴[] = xs⊴[]
resp¹ (x≈y ∷ xs≈ys) halt = halt
resp¹ (x≈y ∷ xs≈ys) (this z<x) = this (proj₁ resp x≈y z<x)
@@ -132,24 +135,24 @@ private
next (Eq.trans z≈x x≈y) (resp¹ xs≈ys zs⊴xs)
where module Eq = IsEquivalence eq
- resp² : ∀ {ys} → flip (Lex P ≈ <) ys Respects Pointwise.Rel ≈
+ resp² : ∀ {ys} → flip (Lex P _≈_ _<_) ys Respects Pointwise.Rel _≈_
resp² [] []⊴ys = []⊴ys
resp² (x≈z ∷ xs≈zs) (this x<y) = this (proj₂ resp x≈z x<y)
resp² (x≈z ∷ xs≈zs) (next x≈y xs⊴ys) =
next (Eq.trans (Eq.sym x≈z) x≈y) (resp² xs≈zs xs⊴ys)
where module Eq = IsEquivalence eq
- decidable : ∀ {P ≈ <} →
- Dec P → Decidable ≈ → Decidable < →
- Decidable (Lex P ≈ <)
- decidable {P} {≈} {<} dec-P dec-≈ dec-< = dec
+ decidable : ∀ {P _≈_ _<_} →
+ Dec P → Decidable _≈_ → Decidable _<_ →
+ Decidable (Lex P _≈_ _<_)
+ decidable {P} {_≈_} {_<_} dec-P dec-≈ dec-< = dec
where
- dec : Decidable (Lex P ≈ <)
+ dec : Decidable (Lex P _≈_ _<_)
dec [] [] with dec-P
... | yes p = yes (base p)
... | no ¬p = no helper
where
- helper : ¬ Lex P ≈ < [] []
+ helper : ¬ Lex P _≈_ _<_ [] []
helper (base p) = ¬p p
dec [] (y ∷ ys) = yes halt
dec (x ∷ xs) [] = no (λ ())
@@ -161,22 +164,25 @@ private
... | yes xs⊴ys = yes (next x≈y xs⊴ys)
... | no ¬xs⊴ys = no (¬≤-next ¬x<y ¬xs⊴ys)
- <-decidable : ∀ {≈ <} →
- Decidable ≈ → Decidable < → Decidable (Lex-< ≈ <)
+ <-decidable :
+ ∀ {_≈_ _<_} →
+ Decidable _≈_ → Decidable _<_ → Decidable (Lex-< _≈_ _<_)
<-decidable = decidable (no id)
- ≤-decidable : ∀ {≈ <} →
- Decidable ≈ → Decidable < → Decidable (Lex-≤ ≈ <)
+ ≤-decidable :
+ ∀ {_≈_ _<_} →
+ Decidable _≈_ → Decidable _<_ → Decidable (Lex-≤ _≈_ _<_)
≤-decidable = decidable (yes tt)
-- Note that trichotomy is an unnecessarily strong precondition for
-- the following lemma.
- ≤-total : ∀ {≈ <} →
- Symmetric ≈ → Trichotomous ≈ < → Total (Lex-≤ ≈ <)
- ≤-total {≈} {<} sym tri = total
+ ≤-total :
+ ∀ {_≈_ _<_} →
+ Symmetric _≈_ → Trichotomous _≈_ _<_ → Total (Lex-≤ _≈_ _<_)
+ ≤-total {_≈_} {_<_} sym tri = total
where
- total : Total (Lex-≤ ≈ <)
+ total : Total (Lex-≤ _≈_ _<_)
total [] [] = inj₁ (base tt)
total [] (x ∷ xs) = inj₁ halt
total (x ∷ xs) [] = inj₂ halt
@@ -187,12 +193,12 @@ private
... | inj₁ xs≲ys = inj₁ (next x≈y xs≲ys)
... | inj₂ ys≲xs = inj₂ (next (sym x≈y) ys≲xs)
- <-compare : ∀ {≈ <} →
- Symmetric ≈ → Trichotomous ≈ < →
- Trichotomous (Pointwise.Rel ≈) (Lex-< ≈ <)
- <-compare {≈} {<} sym tri = cmp
+ <-compare : ∀ {_≈_ _<_} →
+ Symmetric _≈_ → Trichotomous _≈_ _<_ →
+ Trichotomous (Pointwise.Rel _≈_) (Lex-< _≈_ _<_)
+ <-compare {_≈_} {_<_} sym tri = cmp
where
- cmp : Trichotomous (Pointwise.Rel ≈) (Lex-< ≈ <)
+ cmp : Trichotomous (Pointwise.Rel _≈_) (Lex-< _≈_ _<_)
cmp [] [] = tri≈ ¬[]<[] [] ¬[]<[]
cmp [] (y ∷ ys) = tri< halt (λ ()) (λ ())
cmp (x ∷ xs) [] = tri> (λ ()) (λ ()) halt
@@ -212,28 +218,28 @@ private
-- Some collections of properties which are preserved by Lex-≤ or
-- Lex-<.
- ≤-isPreorder : ∀ {≈ <} →
- IsEquivalence ≈ → Transitive < → < Respects₂ ≈ →
- IsPreorder (Pointwise.Rel ≈) (Lex-≤ ≈ <)
- ≤-isPreorder {≈} {<} eq tr resp = record
+ ≤-isPreorder : ∀ {_≈_ _<_} →
+ IsEquivalence _≈_ → Transitive _<_ → _<_ Respects₂ _≈_ →
+ IsPreorder (Pointwise.Rel _≈_) (Lex-≤ _≈_ _<_)
+ ≤-isPreorder {_≈_} {_<_} eq tr resp = record
{ isEquivalence = Pointwise.isEquivalence eq
- ; reflexive = ≤-reflexive ≈ <
+ ; reflexive = ≤-reflexive _≈_ _<_
; trans = transitive eq resp tr
- ; ∼-resp-≈ = respects₂ eq resp
}
- ≤-isPartialOrder : ∀ {≈ <} →
- IsStrictPartialOrder ≈ < →
- IsPartialOrder (Pointwise.Rel ≈) (Lex-≤ ≈ <)
- ≤-isPartialOrder {≈} {<} spo = record
+ ≤-isPartialOrder : ∀ {_≈_ _<_} →
+ IsStrictPartialOrder _≈_ _<_ →
+ IsPartialOrder (Pointwise.Rel _≈_) (Lex-≤ _≈_ _<_)
+ ≤-isPartialOrder {_≈_} {_<_} spo = record
{ isPreorder = ≤-isPreorder isEquivalence trans <-resp-≈
; antisym = antisymmetric Eq.sym irrefl
- (trans∧irr⟶asym {≈ = ≈} {<} Eq.refl trans irrefl)
+ (trans∧irr⟶asym {_≈_ = _≈_} {_<_ = _<_}
+ Eq.refl trans irrefl)
} where open IsStrictPartialOrder spo
- ≤-isTotalOrder : ∀ {≈ <} →
- IsStrictTotalOrder ≈ < →
- IsTotalOrder (Pointwise.Rel ≈) (Lex-≤ ≈ <)
+ ≤-isTotalOrder : ∀ {_≈_ _<_} →
+ IsStrictTotalOrder _≈_ _<_ →
+ IsTotalOrder (Pointwise.Rel _≈_) (Lex-≤ _≈_ _<_)
≤-isTotalOrder sto = record
{ isPartialOrder =
≤-isPartialOrder (record
@@ -245,9 +251,9 @@ private
; total = ≤-total Eq.sym compare
} where open IsStrictTotalOrder sto
- ≤-isDecTotalOrder : ∀ {≈ <} →
- IsStrictTotalOrder ≈ < →
- IsDecTotalOrder (Pointwise.Rel ≈) (Lex-≤ ≈ <)
+ ≤-isDecTotalOrder : ∀ {_≈_ _<_} →
+ IsStrictTotalOrder _≈_ _<_ →
+ IsDecTotalOrder (Pointwise.Rel _≈_) (Lex-≤ _≈_ _<_)
≤-isDecTotalOrder sto = record
{ isTotalOrder = ≤-isTotalOrder sto
; _≟_ = Pointwise.decidable _≟_
@@ -255,8 +261,8 @@ private
} where open IsStrictTotalOrder sto
<-isStrictPartialOrder
- : ∀ {≈ <} → IsStrictPartialOrder ≈ < →
- IsStrictPartialOrder (Pointwise.Rel ≈) (Lex-< ≈ <)
+ : ∀ {_≈_ _<_} → IsStrictPartialOrder _≈_ _<_ →
+ IsStrictPartialOrder (Pointwise.Rel _≈_) (Lex-< _≈_ _<_)
<-isStrictPartialOrder spo = record
{ isEquivalence = Pointwise.isEquivalence isEquivalence
; irrefl = <-irreflexive irrefl
@@ -265,8 +271,8 @@ private
} where open IsStrictPartialOrder spo
<-isStrictTotalOrder
- : ∀ {≈ <} → IsStrictTotalOrder ≈ < →
- IsStrictTotalOrder (Pointwise.Rel ≈) (Lex-< ≈ <)
+ : ∀ {_≈_ _<_} → IsStrictTotalOrder _≈_ _<_ →
+ IsStrictTotalOrder (Pointwise.Rel _≈_) (Lex-< _≈_ _<_)
<-isStrictTotalOrder sto = record
{ isEquivalence = Pointwise.isEquivalence isEquivalence
; trans = transitive isEquivalence <-resp-≈ trans
diff --git a/src/Relation/Binary/NonStrictToStrict.agda b/src/Relation/Binary/NonStrictToStrict.agda
index 37183d4..2ec6806 100644
--- a/src/Relation/Binary/NonStrictToStrict.agda
+++ b/src/Relation/Binary/NonStrictToStrict.agda
@@ -17,7 +17,7 @@ module Relation.Binary.NonStrictToStrict
open import Relation.Nullary
open import Relation.Binary.Consequences
-open import Data.Function
+open import Function
open import Data.Product
open import Data.Sum
diff --git a/src/Relation/Binary/On.agda b/src/Relation/Binary/On.agda
index 48a9a8b..6e5f7d9 100644
--- a/src/Relation/Binary/On.agda
+++ b/src/Relation/Binary/On.agda
@@ -9,7 +9,7 @@ open import Relation.Binary
module Relation.Binary.On {a b} {A : Set a} {B : Set b}
(f : B → A) where
-open import Data.Function
+open import Function
open import Data.Product
implies : ∀ {ℓ₁ ℓ₂} (≈ : Rel A ℓ₁) (∼ : Rel A ℓ₂) →
@@ -70,7 +70,6 @@ isPreorder {≈ = ≈} {∼} pre = record
{ isEquivalence = isEquivalence Pre.isEquivalence
; reflexive = implies ≈ ∼ Pre.reflexive
; trans = transitive ∼ Pre.trans
- ; ∼-resp-≈ = respects₂ ∼ ≈ Pre.∼-resp-≈
}
where module Pre = IsPreorder pre
diff --git a/src/Relation/Binary/OrderMorphism.agda b/src/Relation/Binary/OrderMorphism.agda
index ac35d76..0c2f981 100644
--- a/src/Relation/Binary/OrderMorphism.agda
+++ b/src/Relation/Binary/OrderMorphism.agda
@@ -8,7 +8,7 @@ module Relation.Binary.OrderMorphism where
open import Relation.Binary
open Poset
-import Data.Function as F
+import Function as F
open import Level
record _⇒-Poset_ {p₁ p₂ p₃ p₄ p₅ p₆}
diff --git a/src/Relation/Binary/Product/NonStrictLex.agda b/src/Relation/Binary/Product/NonStrictLex.agda
index 8f2c75c..6467a24 100644
--- a/src/Relation/Binary/Product/NonStrictLex.agda
+++ b/src/Relation/Binary/Product/NonStrictLex.agda
@@ -2,6 +2,8 @@
-- Lexicographic products of binary relations
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- The definition of lexicographic product used here is suitable if
-- the left-hand relation is a (non-strict) partial order.
@@ -18,74 +20,77 @@ open import Relation.Binary.Product.Pointwise as Pointwise
import Relation.Binary.Product.StrictLex as Strict
private
- module Dummy {a₁ a₂ : Set} where
+ module Dummy {a₁ a₂ ℓ₁ ℓ₂} {A₁ : Set a₁} {A₂ : Set a₂} where
- ×-Lex : (≈₁ ≤₁ : Rel a₁ zero) → (≤₂ : Rel a₂ zero) →
- Rel (a₁ × a₂) zero
- ×-Lex ≈₁ ≤₁ ≤₂ = Strict.×-Lex ≈₁ (Conv._<_ ≈₁ ≤₁) ≤₂
+ ×-Lex : (_≈₁_ _≤₁_ : Rel A₁ ℓ₁) → (_≤₂_ : Rel A₂ ℓ₂) →
+ Rel (A₁ × A₂) _
+ ×-Lex _≈₁_ _≤₁_ _≤₂_ = Strict.×-Lex _≈₁_ (Conv._<_ _≈₁_ _≤₁_) _≤₂_
-- Some properties which are preserved by ×-Lex (under certain
-- assumptions).
- ×-reflexive : ∀ ≈₁ ≤₁ {≈₂} ≤₂ →
- ≈₂ ⇒ ≤₂ → (≈₁ ×-Rel ≈₂) ⇒ (×-Lex ≈₁ ≤₁ ≤₂)
- ×-reflexive ≈₁ ≤₁ ≤₂ refl₂ {x} {y} =
- Strict.×-reflexive ≈₁ (Conv._<_ ≈₁ ≤₁) ≤₂ refl₂ {x} {y}
+ ×-reflexive : ∀ _≈₁_ _≤₁_ {_≈₂_} _≤₂_ →
+ _≈₂_ ⇒ _≤₂_ → (_≈₁_ ×-Rel _≈₂_) ⇒ (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-reflexive _≈₁_ _≤₁_ _≤₂_ refl₂ {x} {y} =
+ Strict.×-reflexive _≈₁_ (Conv._<_ _≈₁_ _≤₁_) _≤₂_ refl₂ {x} {y}
- ×-transitive : ∀ {≈₁ ≤₁} → IsPartialOrder ≈₁ ≤₁ →
- ∀ {≤₂} → Transitive ≤₂ →
- Transitive (×-Lex ≈₁ ≤₁ ≤₂)
- ×-transitive {≈₁ = ≈₁} {≤₁ = ≤₁} po₁ {≤₂ = ≤₂} trans₂
+ ×-transitive : ∀ {_≈₁_ _≤₁_} → IsPartialOrder _≈₁_ _≤₁_ →
+ ∀ {_≤₂_} → Transitive _≤₂_ →
+ Transitive (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-transitive {_≈₁_ = _≈₁_} {_≤₁_ = _≤₁_} po₁ {_≤₂_ = _≤₂_} trans₂
{x} {y} {z} =
Strict.×-transitive
- {<₁ = Conv._<_ ≈₁ ≤₁}
+ {_<₁_ = Conv._<_ _≈₁_ _≤₁_}
isEquivalence (Conv.<-resp-≈ _ _ isEquivalence ≤-resp-≈)
(Conv.trans _ _ po₁)
- {≤₂ = ≤₂} trans₂ {x} {y} {z}
+ {_≤₂_ = _≤₂_} trans₂ {x} {y} {z}
where open IsPartialOrder po₁
- ×-antisymmetric : ∀ {≈₁ ≤₁} → IsPartialOrder ≈₁ ≤₁ →
- ∀ {≈₂ ≤₂} → Antisymmetric ≈₂ ≤₂ →
- Antisymmetric (≈₁ ×-Rel ≈₂) (×-Lex ≈₁ ≤₁ ≤₂)
- ×-antisymmetric {≈₁ = ≈₁} {≤₁ = ≤₁} po₁ {≤₂ = ≤₂} antisym₂
- {x} {y} =
- Strict.×-antisymmetric {<₁ = Conv._<_ ≈₁ ≤₁} ≈-sym₁ irrefl₁ asym₁
- {≤₂ = ≤₂} antisym₂ {x} {y}
+ ×-antisymmetric :
+ ∀ {_≈₁_ _≤₁_} → IsPartialOrder _≈₁_ _≤₁_ →
+ ∀ {_≈₂_ _≤₂_} → Antisymmetric _≈₂_ _≤₂_ →
+ Antisymmetric (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-antisymmetric {_≈₁_ = _≈₁_} {_≤₁_ = _≤₁_}
+ po₁ {_≤₂_ = _≤₂_} antisym₂ {x} {y} =
+ Strict.×-antisymmetric {_<₁_ = Conv._<_ _≈₁_ _≤₁_}
+ ≈-sym₁ irrefl₁ asym₁
+ {_≤₂_ = _≤₂_} antisym₂ {x} {y}
where
open IsPartialOrder po₁
open Eq renaming (refl to ≈-refl₁; sym to ≈-sym₁)
- irrefl₁ : Irreflexive ≈₁ (Conv._<_ ≈₁ ≤₁)
- irrefl₁ = Conv.irrefl ≈₁ ≤₁
+ irrefl₁ : Irreflexive _≈₁_ (Conv._<_ _≈₁_ _≤₁_)
+ irrefl₁ = Conv.irrefl _≈₁_ _≤₁_
- asym₁ : Asymmetric (Conv._<_ ≈₁ ≤₁)
- asym₁ = trans∧irr⟶asym {≈ = ≈₁}
+ asym₁ : Asymmetric (Conv._<_ _≈₁_ _≤₁_)
+ asym₁ = trans∧irr⟶asym {_≈_ = _≈₁_}
≈-refl₁ (Conv.trans _ _ po₁) irrefl₁
- ×-≈-respects₂ : ∀ {≈₁ ≤₁} → IsEquivalence ≈₁ → ≤₁ Respects₂ ≈₁ →
- ∀ {≈₂ ≤₂} → ≤₂ Respects₂ ≈₂ →
- (×-Lex ≈₁ ≤₁ ≤₂) Respects₂ (≈₁ ×-Rel ≈₂)
+ ×-≈-respects₂ :
+ ∀ {_≈₁_ _≤₁_} → IsEquivalence _≈₁_ → _≤₁_ Respects₂ _≈₁_ →
+ ∀ {_≈₂_ _≤₂_ : Rel A₂ ℓ₂} → _≤₂_ Respects₂ _≈₂_ →
+ (×-Lex _≈₁_ _≤₁_ _≤₂_) Respects₂ (_≈₁_ ×-Rel _≈₂_)
×-≈-respects₂ eq₁ resp₁ resp₂ =
Strict.×-≈-respects₂ eq₁ (Conv.<-resp-≈ _ _ eq₁ resp₁) resp₂
- ×-decidable : ∀ {≈₁ ≤₁} → Decidable ≈₁ → Decidable ≤₁ →
- ∀ {≤₂} → Decidable ≤₂ →
- Decidable (×-Lex ≈₁ ≤₁ ≤₂)
+ ×-decidable : ∀ {_≈₁_ _≤₁_} → Decidable _≈₁_ → Decidable _≤₁_ →
+ ∀ {_≤₂_} → Decidable _≤₂_ →
+ Decidable (×-Lex _≈₁_ _≤₁_ _≤₂_)
×-decidable dec-≈₁ dec-≤₁ dec-≤₂ =
Strict.×-decidable dec-≈₁ (Conv.decidable _ _ dec-≈₁ dec-≤₁)
dec-≤₂
- ×-total : ∀ {≈₁ ≤₁} → Symmetric ≈₁ → Decidable ≈₁ →
- Antisymmetric ≈₁ ≤₁ → Total ≤₁ →
- ∀ {≤₂} → Total ≤₂ →
- Total (×-Lex ≈₁ ≤₁ ≤₂)
- ×-total {≈₁ = ≈₁} {≤₁ = ≤₁} sym₁ dec₁ antisym₁ total₁
- {≤₂ = ≤₂} total₂ = total
+ ×-total : ∀ {_≈₁_ _≤₁_} → Symmetric _≈₁_ → Decidable _≈₁_ →
+ Antisymmetric _≈₁_ _≤₁_ → Total _≤₁_ →
+ ∀ {_≤₂_} → Total _≤₂_ →
+ Total (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-total {_≈₁_ = _≈₁_} {_≤₁_ = _≤₁_} sym₁ dec₁ antisym₁ total₁
+ {_≤₂_ = _≤₂_} total₂ = total
where
- tri₁ : Trichotomous ≈₁ (Conv._<_ ≈₁ ≤₁)
+ tri₁ : Trichotomous _≈₁_ (Conv._<_ _≈₁_ _≤₁_)
tri₁ = Conv.trichotomous _ _ sym₁ dec₁ antisym₁ total₁
- total : Total (×-Lex ≈₁ ≤₁ ≤₂)
+ total : Total (×-Lex _≈₁_ _≤₁_ _≤₂_)
total x y with tri₁ (proj₁ x) (proj₁ y)
... | tri< x₁<y₁ x₁≉y₁ x₁≯y₁ = inj₁ (inj₁ x₁<y₁)
... | tri> x₁≮y₁ x₁≉y₁ x₁>y₁ = inj₂ (inj₁ x₁>y₁)
@@ -96,45 +101,47 @@ private
-- Some collections of properties which are preserved by ×-Lex
-- (under certain assumptions).
- _×-isPartialOrder_ : ∀ {≈₁ ≤₁} → IsPartialOrder ≈₁ ≤₁ →
- ∀ {≈₂ ≤₂} → IsPartialOrder ≈₂ ≤₂ →
- IsPartialOrder (≈₁ ×-Rel ≈₂) (×-Lex ≈₁ ≤₁ ≤₂)
- _×-isPartialOrder_ {≈₁ = ≈₁} {≤₁ = ≤₁} po₁ {≤₂ = ≤₂} po₂ = record
+ _×-isPartialOrder_ :
+ ∀ {_≈₁_ _≤₁_} → IsPartialOrder _≈₁_ _≤₁_ →
+ ∀ {_≈₂_ _≤₂_} → IsPartialOrder _≈₂_ _≤₂_ →
+ IsPartialOrder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ _×-isPartialOrder_ {_≈₁_ = _≈₁_} {_≤₁_ = _≤₁_} po₁
+ {_≤₂_ = _≤₂_} po₂ = record
{ isPreorder = record
{ isEquivalence = Pointwise._×-isEquivalence_
(isEquivalence po₁)
(isEquivalence po₂)
; reflexive = λ {x y} →
- ×-reflexive ≈₁ ≤₁ ≤₂ (reflexive po₂) {x} {y}
+ ×-reflexive _≈₁_ _≤₁_ _≤₂_ (reflexive po₂)
+ {x} {y}
; trans = λ {x y z} →
- ×-transitive po₁ {≤₂ = ≤₂} (trans po₂)
+ ×-transitive po₁ {_≤₂_ = _≤₂_} (trans po₂)
{x} {y} {z}
- ; ∼-resp-≈ = ×-≈-respects₂ (isEquivalence po₁)
- (≤-resp-≈ po₁)
- (≤-resp-≈ po₂)
}
; antisym = λ {x y} →
- ×-antisymmetric {≤₁ = ≤₁} po₁
- {≤₂ = ≤₂} (antisym po₂) {x} {y}
+ ×-antisymmetric {_≤₁_ = _≤₁_} po₁
+ {_≤₂_ = _≤₂_} (antisym po₂) {x} {y}
}
where open IsPartialOrder
- ×-isTotalOrder : ∀ {≈₁ ≤₁} → Decidable ≈₁ → IsTotalOrder ≈₁ ≤₁ →
- ∀ {≈₂ ≤₂} → IsTotalOrder ≈₂ ≤₂ →
- IsTotalOrder (≈₁ ×-Rel ≈₂) (×-Lex ≈₁ ≤₁ ≤₂)
- ×-isTotalOrder {≤₁ = ≤₁} ≈₁-dec to₁ {≤₂ = ≤₂} to₂ = record
+ ×-isTotalOrder :
+ ∀ {_≈₁_ _≤₁_} → Decidable _≈₁_ → IsTotalOrder _≈₁_ _≤₁_ →
+ ∀ {_≈₂_ _≤₂_} → IsTotalOrder _≈₂_ _≤₂_ →
+ IsTotalOrder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-isTotalOrder {_≤₁_ = _≤₁_} ≈₁-dec to₁ {_≤₂_ = _≤₂_} to₂ = record
{ isPartialOrder = isPartialOrder to₁ ×-isPartialOrder
isPartialOrder to₂
- ; total = ×-total {≤₁ = ≤₁} (Eq.sym to₁) ≈₁-dec
- (antisym to₁) (total to₁)
- {≤₂ = ≤₂} (total to₂)
+ ; total = ×-total {_≤₁_ = _≤₁_} (Eq.sym to₁) ≈₁-dec
+ (antisym to₁) (total to₁)
+ {_≤₂_ = _≤₂_} (total to₂)
}
where open IsTotalOrder
- _×-isDecTotalOrder_ : ∀ {≈₁ ≤₁} → IsDecTotalOrder ≈₁ ≤₁ →
- ∀ {≈₂ ≤₂} → IsDecTotalOrder ≈₂ ≤₂ →
- IsDecTotalOrder (≈₁ ×-Rel ≈₂) (×-Lex ≈₁ ≤₁ ≤₂)
- _×-isDecTotalOrder_ {≤₁ = ≤₁} to₁ {≤₂ = ≤₂} to₂ = record
+ _×-isDecTotalOrder_ :
+ ∀ {_≈₁_ _≤₁_} → IsDecTotalOrder _≈₁_ _≤₁_ →
+ ∀ {_≈₂_ _≤₂_} → IsDecTotalOrder _≈₂_ _≤₂_ →
+ IsDecTotalOrder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ _×-isDecTotalOrder_ {_≤₁_ = _≤₁_} to₁ {_≤₂_ = _≤₂_} to₂ = record
{ isTotalOrder = ×-isTotalOrder (_≟_ to₁)
(isTotalOrder to₁)
(isTotalOrder to₂)
@@ -147,14 +154,16 @@ open Dummy public
-- "Packages" (e.g. posets) can also be combined.
-_×-poset_ : Poset _ _ _ → Poset _ _ _ → Poset _ _ _
+_×-poset_ :
+ ∀ {p₁ p₂ p₃ p₄} → Poset p₁ p₂ _ → Poset p₃ p₄ _ → Poset _ _ _
p₁ ×-poset p₂ = record
{ isPartialOrder = isPartialOrder p₁ ×-isPartialOrder
isPartialOrder p₂
} where open Poset
_×-totalOrder_ :
- DecTotalOrder _ _ _ → TotalOrder _ _ _ → TotalOrder _ _ _
+ ∀ {d₁ d₂ t₃ t₄} →
+ DecTotalOrder d₁ d₂ _ → TotalOrder t₃ t₄ _ → TotalOrder _ _ _
t₁ ×-totalOrder t₂ = record
{ isTotalOrder = ×-isTotalOrder T₁._≟_ T₁.isTotalOrder T₂.isTotalOrder
}
@@ -163,7 +172,8 @@ t₁ ×-totalOrder t₂ = record
module T₂ = TotalOrder t₂
_×-decTotalOrder_ :
- DecTotalOrder _ _ _ → DecTotalOrder _ _ _ → DecTotalOrder _ _ _
+ ∀ {d₁ d₂ d₃ d₄} →
+ DecTotalOrder d₁ d₂ _ → DecTotalOrder d₃ d₄ _ → DecTotalOrder _ _ _
t₁ ×-decTotalOrder t₂ = record
{ isDecTotalOrder = isDecTotalOrder t₁ ×-isDecTotalOrder
isDecTotalOrder t₂
diff --git a/src/Relation/Binary/Product/Pointwise.agda b/src/Relation/Binary/Product/Pointwise.agda
index e9582b6..38d167d 100644
--- a/src/Relation/Binary/Product/Pointwise.agda
+++ b/src/Relation/Binary/Product/Pointwise.agda
@@ -2,92 +2,115 @@
-- Pointwise products of binary relations
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Relation.Binary.Product.Pointwise where
-open import Data.Function
-open import Data.Product
+open import Data.Product as Prod
open import Data.Sum
+open import Data.Unit using (⊤)
+open import Function
+open import Function.Equality as F using (_⟨$⟩_)
+open import Function.Equivalence as Eq
+ using (Equivalent; _⇔_; module Equivalent)
+ renaming (_∘_ to _⟨∘⟩_)
+open import Function.Inverse as Inv
+ using (Inverse; _⇿_; module Inverse)
+ renaming (_∘_ to _⟪∘⟫_)
+open import Function.LeftInverse
+ using (_LeftInverseOf_; _RightInverseOf_)
open import Level
open import Relation.Nullary.Product
open import Relation.Binary
+import Relation.Binary.PropositionalEquality as P
private
- module Dummy {a₁ a₂ : Set} where
+ module Dummy {a₁ a₂ ℓ₁ ℓ₂} {A₁ : Set a₁} {A₂ : Set a₂} where
infixr 2 _×-Rel_
- _×-Rel_ : Rel a₁ zero → Rel a₂ zero → Rel (a₁ × a₂) zero
- ∼₁ ×-Rel ∼₂ = (∼₁ on proj₁) -×- (∼₂ on proj₂)
+ _×-Rel_ : Rel A₁ ℓ₁ → Rel A₂ ℓ₂ → Rel (A₁ × A₂) _
+ _∼₁_ ×-Rel _∼₂_ = (_∼₁_ on proj₁) -×- (_∼₂_ on proj₂)
-- Some properties which are preserved by ×-Rel (under certain
-- assumptions).
- _×-reflexive_ : ∀ {≈₁ ∼₁ ≈₂ ∼₂} →
- ≈₁ ⇒ ∼₁ → ≈₂ ⇒ ∼₂ → (≈₁ ×-Rel ≈₂) ⇒ (∼₁ ×-Rel ∼₂)
+ _×-reflexive_ :
+ ∀ {_≈₁_ _∼₁_ _≈₂_ _∼₂_} →
+ _≈₁_ ⇒ _∼₁_ → _≈₂_ ⇒ _∼₂_ → (_≈₁_ ×-Rel _≈₂_) ⇒ (_∼₁_ ×-Rel _∼₂_)
refl₁ ×-reflexive refl₂ = λ x≈y →
(refl₁ (proj₁ x≈y) , refl₂ (proj₂ x≈y))
- _×-refl_ : ∀ {∼₁ ∼₂} →
- Reflexive ∼₁ → Reflexive ∼₂ → Reflexive (∼₁ ×-Rel ∼₂)
+ _×-refl_ :
+ ∀ {_∼₁_ _∼₂_} →
+ Reflexive _∼₁_ → Reflexive _∼₂_ → Reflexive (_∼₁_ ×-Rel _∼₂_)
refl₁ ×-refl refl₂ = (refl₁ , refl₂)
×-irreflexive₁ :
- ∀ {≈₁ <₁ ≈₂ <₂} →
- Irreflexive ≈₁ <₁ → Irreflexive (≈₁ ×-Rel ≈₂) (<₁ ×-Rel <₂)
+ ∀ {_≈₁_ _<₁_ _≈₂_ _<₂_} →
+ Irreflexive _≈₁_ _<₁_ →
+ Irreflexive (_≈₁_ ×-Rel _≈₂_) (_<₁_ ×-Rel _<₂_)
×-irreflexive₁ ir = λ x≈y x<y → ir (proj₁ x≈y) (proj₁ x<y)
×-irreflexive₂ :
- ∀ {≈₁ <₁ ≈₂ <₂} →
- Irreflexive ≈₂ <₂ → Irreflexive (≈₁ ×-Rel ≈₂) (<₁ ×-Rel <₂)
+ ∀ {_≈₁_ _<₁_ _≈₂_ _<₂_} →
+ Irreflexive _≈₂_ _<₂_ →
+ Irreflexive (_≈₁_ ×-Rel _≈₂_) (_<₁_ ×-Rel _<₂_)
×-irreflexive₂ ir = λ x≈y x<y → ir (proj₂ x≈y) (proj₂ x<y)
- _×-symmetric_ : ∀ {∼₁ ∼₂} →
- Symmetric ∼₁ → Symmetric ∼₂ → Symmetric (∼₁ ×-Rel ∼₂)
+ _×-symmetric_ :
+ ∀ {_∼₁_ _∼₂_} →
+ Symmetric _∼₁_ → Symmetric _∼₂_ → Symmetric (_∼₁_ ×-Rel _∼₂_)
sym₁ ×-symmetric sym₂ = λ x∼y → sym₁ (proj₁ x∼y) , sym₂ (proj₂ x∼y)
- _×-transitive_ : ∀ {∼₁ ∼₂} →
- Transitive ∼₁ → Transitive ∼₂ →
- Transitive (∼₁ ×-Rel ∼₂)
+ _×-transitive_ : ∀ {_∼₁_ _∼₂_} →
+ Transitive _∼₁_ → Transitive _∼₂_ →
+ Transitive (_∼₁_ ×-Rel _∼₂_)
trans₁ ×-transitive trans₂ = λ x∼y y∼z →
trans₁ (proj₁ x∼y) (proj₁ y∼z) ,
trans₂ (proj₂ x∼y) (proj₂ y∼z)
- _×-antisymmetric_ : ∀ {≈₁ ≤₁ ≈₂ ≤₂} →
- Antisymmetric ≈₁ ≤₁ → Antisymmetric ≈₂ ≤₂ →
- Antisymmetric (≈₁ ×-Rel ≈₂) (≤₁ ×-Rel ≤₂)
+ _×-antisymmetric_ :
+ ∀ {_≈₁_ _≤₁_ _≈₂_ _≤₂_} →
+ Antisymmetric _≈₁_ _≤₁_ → Antisymmetric _≈₂_ _≤₂_ →
+ Antisymmetric (_≈₁_ ×-Rel _≈₂_) (_≤₁_ ×-Rel _≤₂_)
antisym₁ ×-antisymmetric antisym₂ = λ x≤y y≤x →
( antisym₁ (proj₁ x≤y) (proj₁ y≤x)
, antisym₂ (proj₂ x≤y) (proj₂ y≤x) )
- ×-asymmetric₁ : ∀ {<₁ ∼₂} → Asymmetric <₁ → Asymmetric (<₁ ×-Rel ∼₂)
+ ×-asymmetric₁ :
+ ∀ {_<₁_ _∼₂_} → Asymmetric _<₁_ → Asymmetric (_<₁_ ×-Rel _∼₂_)
×-asymmetric₁ asym₁ = λ x<y y<x → asym₁ (proj₁ x<y) (proj₁ y<x)
- ×-asymmetric₂ : ∀ {∼₁ <₂} → Asymmetric <₂ → Asymmetric (∼₁ ×-Rel <₂)
+ ×-asymmetric₂ :
+ ∀ {_∼₁_ _<₂_} → Asymmetric _<₂_ → Asymmetric (_∼₁_ ×-Rel _<₂_)
×-asymmetric₂ asym₂ = λ x<y y<x → asym₂ (proj₂ x<y) (proj₂ y<x)
- _×-≈-respects₂_ : ∀ {≈₁ ∼₁ ≈₂ ∼₂} →
- ∼₁ Respects₂ ≈₁ → ∼₂ Respects₂ ≈₂ →
- (∼₁ ×-Rel ∼₂) Respects₂ (≈₁ ×-Rel ≈₂)
- _×-≈-respects₂_ {≈₁ = ≈₁} {∼₁ = ∼₁} {≈₂ = ≈₂} {∼₂ = ∼₂}
- resp₁ resp₂ =
+ _×-≈-respects₂_ : ∀ {_≈₁_ _∼₁_ _≈₂_ _∼₂_} →
+ _∼₁_ Respects₂ _≈₁_ → _∼₂_ Respects₂ _≈₂_ →
+ (_∼₁_ ×-Rel _∼₂_) Respects₂ (_≈₁_ ×-Rel _≈₂_)
+ _×-≈-respects₂_
+ {_≈₁_ = _≈₁_} {_∼₁_ = _∼₁_} {_≈₂_ = _≈₂_} {_∼₂_ = _∼₂_}
+ resp₁ resp₂ =
(λ {x y z} → resp¹ {x} {y} {z}) ,
(λ {x y z} → resp² {x} {y} {z})
where
- ∼ = ∼₁ ×-Rel ∼₂
+ _∼_ = _∼₁_ ×-Rel _∼₂_
- resp¹ : ∀ {x} → (∼ x) Respects (≈₁ ×-Rel ≈₂)
+ resp¹ : ∀ {x} → (_∼_ x) Respects (_≈₁_ ×-Rel _≈₂_)
resp¹ y≈y' x∼y = proj₁ resp₁ (proj₁ y≈y') (proj₁ x∼y) ,
proj₁ resp₂ (proj₂ y≈y') (proj₂ x∼y)
- resp² : ∀ {y} → (flip ∼ y) Respects (≈₁ ×-Rel ≈₂)
+ resp² : ∀ {y} → (flip _∼_ y) Respects (_≈₁_ ×-Rel _≈₂_)
resp² x≈x' x∼y = proj₂ resp₁ (proj₁ x≈x') (proj₁ x∼y) ,
proj₂ resp₂ (proj₂ x≈x') (proj₂ x∼y)
- ×-total : ∀ {∼₁ ∼₂} →
- Symmetric ∼₁ → Total ∼₁ → Total ∼₂ → Total (∼₁ ×-Rel ∼₂)
- ×-total {∼₁ = ∼₁} {∼₂ = ∼₂} sym₁ total₁ total₂ = total
+ ×-total :
+ ∀ {_∼₁_ _∼₂_} →
+ Symmetric _∼₁_ → Total _∼₁_ → Total _∼₂_ → Total (_∼₁_ ×-Rel _∼₂_)
+ ×-total {_∼₁_ = _∼₁_} {_∼₂_ = _∼₂_} sym₁ total₁ total₂ = total
where
- total : Total (∼₁ ×-Rel ∼₂)
+ total : Total (_∼₁_ ×-Rel _∼₂_)
total x y with total₁ (proj₁ x) (proj₁ y)
| total₂ (proj₂ x) (proj₂ y)
... | inj₁ x₁∼y₁ | inj₁ x₂∼y₂ = inj₁ ( x₁∼y₁ , x₂∼y₂)
@@ -95,8 +118,9 @@ private
... | inj₂ y₁∼x₁ | inj₂ y₂∼x₂ = inj₂ ( y₁∼x₁ , y₂∼x₂)
... | inj₂ y₁∼x₁ | inj₁ x₂∼y₂ = inj₁ (sym₁ y₁∼x₁ , x₂∼y₂)
- _×-decidable_ : ∀ {∼₁ ∼₂} →
- Decidable ∼₁ → Decidable ∼₂ → Decidable (∼₁ ×-Rel ∼₂)
+ _×-decidable_ :
+ ∀ {_∼₁_ _∼₂_} →
+ Decidable _∼₁_ → Decidable _∼₂_ → Decidable (_∼₁_ ×-Rel _∼₂_)
dec₁ ×-decidable dec₂ = λ x y →
dec₁ (proj₁ x) (proj₁ y)
×-dec
@@ -104,43 +128,43 @@ private
-- Some collections of properties which are preserved by ×-Rel.
- _×-isEquivalence_ : ∀ {≈₁ ≈₂} →
- IsEquivalence ≈₁ → IsEquivalence ≈₂ →
- IsEquivalence (≈₁ ×-Rel ≈₂)
- _×-isEquivalence_ {≈₁ = ≈₁} {≈₂ = ≈₂} eq₁ eq₂ = record
+ _×-isEquivalence_ : ∀ {_≈₁_ _≈₂_} →
+ IsEquivalence _≈₁_ → IsEquivalence _≈₂_ →
+ IsEquivalence (_≈₁_ ×-Rel _≈₂_)
+ _×-isEquivalence_ {_≈₁_ = _≈₁_} {_≈₂_ = _≈₂_} eq₁ eq₂ = record
{ refl = λ {x} →
- _×-refl_ {∼₁ = ≈₁} {∼₂ = ≈₂}
+ _×-refl_ {_∼₁_ = _≈₁_} {_∼₂_ = _≈₂_}
(refl eq₁) (refl eq₂) {x}
; sym = λ {x y} →
- _×-symmetric_ {∼₁ = ≈₁} {∼₂ = ≈₂}
+ _×-symmetric_ {_∼₁_ = _≈₁_} {_∼₂_ = _≈₂_}
(sym eq₁) (sym eq₂) {x} {y}
; trans = λ {x y z} →
- _×-transitive_ {∼₁ = ≈₁} {∼₂ = ≈₂}
+ _×-transitive_ {_∼₁_ = _≈₁_} {_∼₂_ = _≈₂_}
(trans eq₁) (trans eq₂) {x} {y} {z}
}
where open IsEquivalence
- _×-isPreorder_ : ∀ {≈₁ ∼₁ ≈₂ ∼₂} →
- IsPreorder ≈₁ ∼₁ → IsPreorder ≈₂ ∼₂ →
- IsPreorder (≈₁ ×-Rel ≈₂) (∼₁ ×-Rel ∼₂)
- _×-isPreorder_ {∼₁ = ∼₁} {∼₂ = ∼₂} pre₁ pre₂ = record
+ _×-isPreorder_ : ∀ {_≈₁_ _∼₁_ _≈₂_ _∼₂_} →
+ IsPreorder _≈₁_ _∼₁_ → IsPreorder _≈₂_ _∼₂_ →
+ IsPreorder (_≈₁_ ×-Rel _≈₂_) (_∼₁_ ×-Rel _∼₂_)
+ _×-isPreorder_ {_∼₁_ = _∼₁_} {_∼₂_ = _∼₂_} pre₁ pre₂ = record
{ isEquivalence = isEquivalence pre₁ ×-isEquivalence
isEquivalence pre₂
; reflexive = λ {x y} →
- _×-reflexive_ {∼₁ = ∼₁} {∼₂ = ∼₂}
+ _×-reflexive_ {_∼₁_ = _∼₁_} {_∼₂_ = _∼₂_}
(reflexive pre₁) (reflexive pre₂)
{x} {y}
; trans = λ {x y z} →
- _×-transitive_ {∼₁ = ∼₁} {∼₂ = ∼₂}
+ _×-transitive_ {_∼₁_ = _∼₁_} {_∼₂_ = _∼₂_}
(trans pre₁) (trans pre₂)
{x} {y} {z}
- ; ∼-resp-≈ = ∼-resp-≈ pre₁ ×-≈-respects₂ ∼-resp-≈ pre₂
}
where open IsPreorder
- _×-isDecEquivalence_ : ∀ {≈₁ ≈₂} →
- IsDecEquivalence ≈₁ → IsDecEquivalence ≈₂ →
- IsDecEquivalence (≈₁ ×-Rel ≈₂)
+ _×-isDecEquivalence_ :
+ ∀ {_≈₁_ _≈₂_} →
+ IsDecEquivalence _≈₁_ → IsDecEquivalence _≈₂_ →
+ IsDecEquivalence (_≈₁_ ×-Rel _≈₂_)
eq₁ ×-isDecEquivalence eq₂ = record
{ isEquivalence = isEquivalence eq₁ ×-isEquivalence
isEquivalence eq₂
@@ -148,31 +172,34 @@ private
}
where open IsDecEquivalence
- _×-isPartialOrder_ : ∀ {≈₁ ≤₁ ≈₂ ≤₂} →
- IsPartialOrder ≈₁ ≤₁ → IsPartialOrder ≈₂ ≤₂ →
- IsPartialOrder (≈₁ ×-Rel ≈₂) (≤₁ ×-Rel ≤₂)
- _×-isPartialOrder_ {≤₁ = ≤₁} {≤₂ = ≤₂} po₁ po₂ = record
+ _×-isPartialOrder_ :
+ ∀ {_≈₁_ _≤₁_ _≈₂_ _≤₂_} →
+ IsPartialOrder _≈₁_ _≤₁_ → IsPartialOrder _≈₂_ _≤₂_ →
+ IsPartialOrder (_≈₁_ ×-Rel _≈₂_) (_≤₁_ ×-Rel _≤₂_)
+ _×-isPartialOrder_ {_≤₁_ = _≤₁_} {_≤₂_ = _≤₂_} po₁ po₂ = record
{ isPreorder = isPreorder po₁ ×-isPreorder isPreorder po₂
; antisym = λ {x y} →
- _×-antisymmetric_ {≤₁ = ≤₁} {≤₂ = ≤₂}
+ _×-antisymmetric_ {_≤₁_ = _≤₁_} {_≤₂_ = _≤₂_}
(antisym po₁) (antisym po₂)
{x} {y}
}
where open IsPartialOrder
_×-isStrictPartialOrder_ :
- ∀ {≈₁ <₁ ≈₂ <₂} →
- IsStrictPartialOrder ≈₁ <₁ → IsStrictPartialOrder ≈₂ <₂ →
- IsStrictPartialOrder (≈₁ ×-Rel ≈₂) (<₁ ×-Rel <₂)
- _×-isStrictPartialOrder_ {<₁ = <₁} {≈₂ = ≈₂} {<₂ = <₂} spo₁ spo₂ =
+ ∀ {_≈₁_ _<₁_ _≈₂_ _<₂_} →
+ IsStrictPartialOrder _≈₁_ _<₁_ → IsStrictPartialOrder _≈₂_ _<₂_ →
+ IsStrictPartialOrder (_≈₁_ ×-Rel _≈₂_) (_<₁_ ×-Rel _<₂_)
+ _×-isStrictPartialOrder_ {_<₁_ = _<₁_} {_≈₂_ = _≈₂_} {_<₂_ = _<₂_}
+ spo₁ spo₂ =
record
{ isEquivalence = isEquivalence spo₁ ×-isEquivalence
isEquivalence spo₂
; irrefl = λ {x y} →
- ×-irreflexive₁ {<₁ = <₁} {≈₂ = ≈₂} {<₂ = <₂}
+ ×-irreflexive₁ {_<₁_ = _<₁_}
+ {_≈₂_ = _≈₂_} {_<₂_ = _<₂_}
(irrefl spo₁) {x} {y}
; trans = λ {x y z} →
- _×-transitive_ {∼₁ = <₁} {∼₂ = <₂}
+ _×-transitive_ {_∼₁_ = _<₁_} {_∼₂_ = _<₂_}
(trans spo₁) (trans spo₂)
{x} {y} {z}
; <-resp-≈ = <-resp-≈ spo₁ ×-≈-respects₂ <-resp-≈ spo₂
@@ -183,33 +210,134 @@ open Dummy public
-- "Packages" (e.g. setoids) can also be combined.
-_×-preorder_ : Preorder _ _ _ → Preorder _ _ _ → Preorder _ _ _
+_×-preorder_ :
+ ∀ {p₁ p₂ p₃ p₄} →
+ Preorder p₁ p₂ _ → Preorder p₃ p₄ _ → Preorder _ _ _
p₁ ×-preorder p₂ = record
{ isPreorder = isPreorder p₁ ×-isPreorder isPreorder p₂
} where open Preorder
-_×-setoid_ : Setoid _ _ → Setoid _ _ → Setoid _ _
+_×-setoid_ :
+ ∀ {s₁ s₂ s₃ s₄} → Setoid s₁ s₂ → Setoid s₃ s₄ → Setoid _ _
s₁ ×-setoid s₂ = record
{ isEquivalence = isEquivalence s₁ ×-isEquivalence isEquivalence s₂
} where open Setoid
-_×-decSetoid_ : DecSetoid _ _ → DecSetoid _ _ → DecSetoid _ _
+_×-decSetoid_ :
+ ∀ {d₁ d₂ d₃ d₄} → DecSetoid d₁ d₂ → DecSetoid d₃ d₄ → DecSetoid _ _
s₁ ×-decSetoid s₂ = record
{ isDecEquivalence = isDecEquivalence s₁ ×-isDecEquivalence
isDecEquivalence s₂
} where open DecSetoid
-_×-poset_ : Poset _ _ _ → Poset _ _ _ → Poset _ _ _
+_×-poset_ :
+ ∀ {p₁ p₂ p₃ p₄} → Poset p₁ p₂ _ → Poset p₃ p₄ _ → Poset _ _ _
s₁ ×-poset s₂ = record
{ isPartialOrder = isPartialOrder s₁ ×-isPartialOrder
isPartialOrder s₂
} where open Poset
_×-strictPartialOrder_ :
- StrictPartialOrder _ _ _ → StrictPartialOrder _ _ _ →
+ ∀ {s₁ s₂ s₃ s₄} →
+ StrictPartialOrder s₁ s₂ _ → StrictPartialOrder s₃ s₄ _ →
StrictPartialOrder _ _ _
s₁ ×-strictPartialOrder s₂ = record
{ isStrictPartialOrder = isStrictPartialOrder s₁
×-isStrictPartialOrder
isStrictPartialOrder s₂
} where open StrictPartialOrder
+
+------------------------------------------------------------------------
+-- Some properties related to equivalences and inverses
+
+×-Rel⇿≡ : ∀ {a b} {A : Set a} {B : Set b} →
+ Inverse (P.setoid A ×-setoid P.setoid B) (P.setoid (A × B))
+×-Rel⇿≡ {A = A} {B} = record
+ { to = record { _⟨$⟩_ = id; cong = to-cong }
+ ; from = record { _⟨$⟩_ = id; cong = from-cong }
+ ; inverse-of = record
+ { left-inverse-of = λ _ → (P.refl , P.refl)
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+ where
+ to-cong : (P._≡_ ×-Rel P._≡_) ⇒ P._≡_
+ to-cong (eq₁ , eq₂) = helper eq₁ eq₂
+ where
+ open P using (_≡_)
+
+ helper : {x x′ : A} {y y′ : B} →
+ x ≡ x′ → y ≡ y′ → _≡_ {A = A × B} (x , y) (x′ , y′)
+ helper P.refl P.refl = P.refl
+
+ from-cong : P._≡_ ⇒ (P._≡_ ×-Rel P._≡_)
+ from-cong P.refl = (P.refl , P.refl)
+
+_×-equivalent_ :
+ ∀ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
+ {A : Setoid a₁ a₂} {B : Setoid b₁ b₂}
+ {C : Setoid c₁ c₂} {D : Setoid d₁ d₂} →
+ Equivalent A B → Equivalent C D →
+ Equivalent (A ×-setoid C) (B ×-setoid D)
+_×-equivalent_ {A = A} {B} {C} {D} A⇔B C⇔D = record
+ { to = record { _⟨$⟩_ = to; cong = λ {x y} → to-cong {x} {y} }
+ ; from = record { _⟨$⟩_ = from; cong = λ {x y} → from-cong {x} {y} }
+ }
+ where
+ open Setoid (A ×-setoid C) using () renaming (_≈_ to _≈AC_)
+ open Setoid (B ×-setoid D) using () renaming (_≈_ to _≈BD_)
+
+ to = Prod.map (_⟨$⟩_ (Equivalent.to A⇔B))
+ (_⟨$⟩_ (Equivalent.to C⇔D))
+
+ to-cong : _≈AC_ =[ to ]⇒ _≈BD_
+ to-cong (_∼₁_ , _∼₂_) =
+ (F.cong (Equivalent.to A⇔B) _∼₁_ , F.cong (Equivalent.to C⇔D) _∼₂_)
+
+ from = Prod.map (_⟨$⟩_ (Equivalent.from A⇔B))
+ (_⟨$⟩_ (Equivalent.from C⇔D))
+
+ from-cong : _≈BD_ =[ from ]⇒ _≈AC_
+ from-cong (_∼₁_ , _∼₂_) =
+ (F.cong (Equivalent.from A⇔B) _∼₁_ ,
+ F.cong (Equivalent.from C⇔D) _∼₂_)
+
+_×-⇔_ : ∀ {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 = A} {B} {C} {D} A⇔B C⇔D =
+ Inverse.equivalent (×-Rel⇿≡ {A = B} {B = D}) ⟨∘⟩
+ A⇔B ×-equivalent C⇔D ⟨∘⟩
+ Eq.sym (Inverse.equivalent (×-Rel⇿≡ {A = A} {B = C}))
+
+_×-inverse_ :
+ ∀ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
+ {A : Setoid a₁ a₂} {B : Setoid b₁ b₂}
+ {C : Setoid c₁ c₂} {D : Setoid d₁ d₂} →
+ Inverse A B → Inverse C D → Inverse (A ×-setoid C) (B ×-setoid D)
+A⇿B ×-inverse C⇿D = record
+ { to = Equivalent.to eq
+ ; from = Equivalent.from eq
+ ; inverse-of = record
+ { left-inverse-of = left
+ ; right-inverse-of = right
+ }
+ }
+ where
+ eq = Inverse.equivalent A⇿B ×-equivalent Inverse.equivalent C⇿D
+
+ left : Equivalent.from eq LeftInverseOf Equivalent.to eq
+ left (x , y) = ( Inverse.left-inverse-of A⇿B x
+ , Inverse.left-inverse-of C⇿D y
+ )
+
+ right : Equivalent.from eq RightInverseOf Equivalent.to eq
+ right (x , y) = ( Inverse.right-inverse-of A⇿B x
+ , Inverse.right-inverse-of C⇿D y
+ )
+
+_×-⇿_ : ∀ {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 = A} {B} {C} {D} A⇿B C⇿D =
+ ×-Rel⇿≡ {A = B} {B = D} ⟪∘⟫
+ A⇿B ×-inverse C⇿D ⟪∘⟫
+ Inv.sym (×-Rel⇿≡ {A = A} {B = C})
diff --git a/src/Relation/Binary/Product/StrictLex.agda b/src/Relation/Binary/Product/StrictLex.agda
index 59b4a4b..3a3542b 100644
--- a/src/Relation/Binary/Product/StrictLex.agda
+++ b/src/Relation/Binary/Product/StrictLex.agda
@@ -2,12 +2,14 @@
-- Lexicographic products of binary relations
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
-- The definition of lexicographic product used here is suitable if
-- the left-hand relation is a strict partial order.
module Relation.Binary.Product.StrictLex where
-open import Data.Function
+open import Function
open import Data.Product
open import Data.Sum
open import Data.Empty
@@ -18,39 +20,43 @@ open import Relation.Binary
open import Relation.Binary.Consequences
open import Relation.Binary.Product.Pointwise as Pointwise
using (_×-Rel_)
+open import Relation.Nullary
private
- module Dummy {a₁ a₂ : Set} where
+ module Dummy {a₁ a₂ ℓ₁ ℓ₂} {A₁ : Set a₁} {A₂ : Set a₂} where
- ×-Lex : (≈₁ <₁ : Rel a₁ zero) → (≤₂ : Rel a₂ zero) →
- Rel (a₁ × a₂) zero
- ×-Lex ≈₁ <₁ ≤₂ = (<₁ on proj₁) -⊎- (≈₁ on proj₁) -×- (≤₂ on proj₂)
+ ×-Lex : (_≈₁_ _<₁_ : Rel A₁ ℓ₁) → (_≤₂_ : Rel A₂ ℓ₂) →
+ Rel (A₁ × A₂) _
+ ×-Lex _≈₁_ _<₁_ _≤₂_ =
+ (_<₁_ on proj₁) -⊎- (_≈₁_ on proj₁) -×- (_≤₂_ on proj₂)
-- Some properties which are preserved by ×-Lex (under certain
-- assumptions).
- ×-reflexive : ∀ ≈₁ ∼₁ {≈₂} ≤₂ →
- ≈₂ ⇒ ≤₂ → (≈₁ ×-Rel ≈₂) ⇒ (×-Lex ≈₁ ∼₁ ≤₂)
+ ×-reflexive : ∀ _≈₁_ _∼₁_ {_≈₂_ : Rel A₂ ℓ₂} _≤₂_ →
+ _≈₂_ ⇒ _≤₂_ → (_≈₁_ ×-Rel _≈₂_) ⇒ (×-Lex _≈₁_ _∼₁_ _≤₂_)
×-reflexive _ _ _ refl₂ = λ x≈y →
inj₂ (proj₁ x≈y , refl₂ (proj₂ x≈y))
- _×-irreflexive_ : ∀ {≈₁ <₁} → Irreflexive ≈₁ <₁ →
- ∀ {≈₂ <₂} → Irreflexive ≈₂ <₂ →
- Irreflexive (≈₁ ×-Rel ≈₂) (×-Lex ≈₁ <₁ <₂)
+ _×-irreflexive_ : ∀ {_≈₁_ _<₁_} → Irreflexive _≈₁_ _<₁_ →
+ ∀ {_≈₂_ _<₂_ : Rel A₂ ℓ₂} → Irreflexive _≈₂_ _<₂_ →
+ Irreflexive (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_)
(ir₁ ×-irreflexive ir₂) x≈y (inj₁ x₁<y₁) = ir₁ (proj₁ x≈y) x₁<y₁
(ir₁ ×-irreflexive ir₂) x≈y (inj₂ x≈<y) =
ir₂ (proj₂ x≈y) (proj₂ x≈<y)
×-transitive :
- ∀ {≈₁ <₁} → IsEquivalence ≈₁ → <₁ Respects₂ ≈₁ → Transitive <₁ →
- ∀ {≤₂} → Transitive ≤₂ →
- Transitive (×-Lex ≈₁ <₁ ≤₂)
- ×-transitive {≈₁ = ≈₁} {<₁ = <₁} eq₁ resp₁ trans₁
- {≤₂ = ≤₂} trans₂ {x} {y} {z} = trans {x} {y} {z}
+ ∀ {_≈₁_ _<₁_} →
+ IsEquivalence _≈₁_ → _<₁_ Respects₂ _≈₁_ → Transitive _<₁_ →
+ ∀ {_≤₂_} →
+ Transitive _≤₂_ →
+ Transitive (×-Lex _≈₁_ _<₁_ _≤₂_)
+ ×-transitive {_≈₁_ = _≈₁_} {_<₁_ = _<₁_} eq₁ resp₁ trans₁
+ {_≤₂_ = _≤₂_} trans₂ {x} {y} {z} = trans {x} {y} {z}
where
module Eq₁ = IsEquivalence eq₁
- trans : Transitive (×-Lex ≈₁ <₁ ≤₂)
+ trans : Transitive (×-Lex _≈₁_ _<₁_ _≤₂_)
trans (inj₁ x₁<y₁) (inj₁ y₁<z₁) = inj₁ (trans₁ x₁<y₁ y₁<z₁)
trans (inj₁ x₁<y₁) (inj₂ y≈≤z) =
inj₁ (proj₁ resp₁ (proj₁ y≈≤z) x₁<y₁)
@@ -61,14 +67,16 @@ private
, trans₂ (proj₂ x≈≤y) (proj₂ y≈≤z) )
×-antisymmetric :
- ∀ {≈₁ <₁} → Symmetric ≈₁ → Irreflexive ≈₁ <₁ → Asymmetric <₁ →
- ∀ {≈₂ ≤₂} → Antisymmetric ≈₂ ≤₂ →
- Antisymmetric (≈₁ ×-Rel ≈₂) (×-Lex ≈₁ <₁ ≤₂)
- ×-antisymmetric {≈₁ = ≈₁} {<₁ = <₁} sym₁ irrefl₁ asym₁
- {≈₂ = ≈₂} {≤₂ = ≤₂} antisym₂ {x} {y} =
+ ∀ {_≈₁_ _<₁_} →
+ Symmetric _≈₁_ → Irreflexive _≈₁_ _<₁_ → Asymmetric _<₁_ →
+ ∀ {_≈₂_ _≤₂_ : Rel A₂ ℓ₂} →
+ Antisymmetric _≈₂_ _≤₂_ →
+ Antisymmetric (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _≤₂_)
+ ×-antisymmetric {_≈₁_ = _≈₁_} {_<₁_ = _<₁_} sym₁ irrefl₁ asym₁
+ {_≈₂_ = _≈₂_} {_≤₂_ = _≤₂_} antisym₂ {x} {y} =
antisym {x} {y}
where
- antisym : Antisymmetric (≈₁ ×-Rel ≈₂) (×-Lex ≈₁ <₁ ≤₂)
+ antisym : Antisymmetric (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _≤₂_)
antisym (inj₁ x₁<y₁) (inj₁ y₁<x₁) =
⊥-elim {Whatever = _ × _} $ asym₁ x₁<y₁ y₁<x₁
antisym (inj₁ x₁<y₁) (inj₂ y≈≤x) =
@@ -79,49 +87,51 @@ private
proj₁ x≈≤y , antisym₂ (proj₂ x≈≤y) (proj₂ y≈≤x)
×-asymmetric :
- ∀ {≈₁ <₁} → Symmetric ≈₁ → <₁ Respects₂ ≈₁ → Asymmetric <₁ →
- ∀ {<₂} → Asymmetric <₂ →
- Asymmetric (×-Lex ≈₁ <₁ <₂)
- ×-asymmetric {≈₁ = ≈₁} {<₁ = <₁} sym₁ resp₁ asym₁
- {<₂ = <₂} asym₂ {x} {y} = asym {x} {y}
+ ∀ {_≈₁_ _<₁_} →
+ Symmetric _≈₁_ → _<₁_ Respects₂ _≈₁_ → Asymmetric _<₁_ →
+ ∀ {_<₂_} →
+ Asymmetric _<₂_ →
+ Asymmetric (×-Lex _≈₁_ _<₁_ _<₂_)
+ ×-asymmetric {_≈₁_ = _≈₁_} {_<₁_ = _<₁_} sym₁ resp₁ asym₁
+ {_<₂_ = _<₂_} asym₂ {x} {y} = asym {x} {y}
where
- irrefl₁ : Irreflexive ≈₁ <₁
+ irrefl₁ : Irreflexive _≈₁_ _<₁_
irrefl₁ = asym⟶irr resp₁ sym₁ asym₁
- asym : Asymmetric (×-Lex ≈₁ <₁ <₂)
+ asym : Asymmetric (×-Lex _≈₁_ _<₁_ _<₂_)
asym (inj₁ x₁<y₁) (inj₁ y₁<x₁) = asym₁ x₁<y₁ y₁<x₁
asym (inj₁ x₁<y₁) (inj₂ y≈<x) = irrefl₁ (sym₁ $ proj₁ y≈<x) x₁<y₁
asym (inj₂ x≈<y) (inj₁ y₁<x₁) = irrefl₁ (sym₁ $ proj₁ x≈<y) y₁<x₁
asym (inj₂ x≈<y) (inj₂ y≈<x) = asym₂ (proj₂ x≈<y) (proj₂ y≈<x)
×-≈-respects₂ :
- ∀ {≈₁ <₁} → IsEquivalence ≈₁ → <₁ Respects₂ ≈₁ →
- ∀ {≈₂ <₂} → <₂ Respects₂ ≈₂ →
- (×-Lex ≈₁ <₁ <₂) Respects₂ (≈₁ ×-Rel ≈₂)
- ×-≈-respects₂ {≈₁ = ≈₁} {<₁ = <₁} eq₁ resp₁
- {≈₂ = ≈₂} {<₂ = <₂} resp₂ =
+ ∀ {_≈₁_ _<₁_} → IsEquivalence _≈₁_ → _<₁_ Respects₂ _≈₁_ →
+ ∀ {_≈₂_ _<₂_} → _<₂_ Respects₂ _≈₂_ →
+ (×-Lex _≈₁_ _<₁_ _<₂_) Respects₂ (_≈₁_ ×-Rel _≈₂_)
+ ×-≈-respects₂ {_≈₁_ = _≈₁_} {_<₁_ = _<₁_} eq₁ resp₁
+ {_≈₂_ = _≈₂_} {_<₂_ = _<₂_} resp₂ =
(λ {x y z} → resp¹ {x} {y} {z}) ,
(λ {x y z} → resp² {x} {y} {z})
where
- < = ×-Lex ≈₁ <₁ <₂
+ _<_ = ×-Lex _≈₁_ _<₁_ _<₂_
open IsEquivalence eq₁ renaming (sym to sym₁; trans to trans₁)
- resp¹ : ∀ {x} → (< x) Respects (≈₁ ×-Rel ≈₂)
+ resp¹ : ∀ {x} → (_<_ x) Respects (_≈₁_ ×-Rel _≈₂_)
resp¹ y≈y' (inj₁ x₁<y₁) = inj₁ (proj₁ resp₁ (proj₁ y≈y') x₁<y₁)
resp¹ y≈y' (inj₂ x≈<y) =
inj₂ ( trans₁ (proj₁ x≈<y) (proj₁ y≈y')
, proj₁ resp₂ (proj₂ y≈y') (proj₂ x≈<y) )
- resp² : ∀ {y} → (flip < y) Respects (≈₁ ×-Rel ≈₂)
+ resp² : ∀ {y} → (flip _<_ y) Respects (_≈₁_ ×-Rel _≈₂_)
resp² x≈x' (inj₁ x₁<y₁) = inj₁ (proj₂ resp₁ (proj₁ x≈x') x₁<y₁)
resp² x≈x' (inj₂ x≈<y) =
inj₂ ( trans₁ (sym₁ $ proj₁ x≈x') (proj₁ x≈<y)
, proj₂ resp₂ (proj₂ x≈x') (proj₂ x≈<y) )
- ×-decidable : ∀ {≈₁ <₁} → Decidable ≈₁ → Decidable <₁ →
- ∀ {≤₂} → Decidable ≤₂ →
- Decidable (×-Lex ≈₁ <₁ ≤₂)
+ ×-decidable : ∀ {_≈₁_ _<₁_} → Decidable _≈₁_ → Decidable _<₁_ →
+ ∀ {_≤₂_} → Decidable _≤₂_ →
+ Decidable (×-Lex _≈₁_ _<₁_ _≤₂_)
×-decidable dec-≈₁ dec-<₁ dec-≤₂ = λ x y →
dec-<₁ (proj₁ x) (proj₁ y)
⊎-dec
@@ -129,93 +139,111 @@ private
×-dec
dec-≤₂ (proj₂ x) (proj₂ y))
- ×-total : ∀ {≈₁ <₁} → Total <₁ →
- ∀ {≤₂} →
- Total (×-Lex ≈₁ <₁ ≤₂)
- ×-total {≈₁ = ≈₁} {<₁ = <₁} total₁ {≤₂ = ≤₂} = total
+ ×-total : ∀ {_≈₁_ _<₁_} → Total _<₁_ →
+ ∀ {_≤₂_} →
+ Total (×-Lex _≈₁_ _<₁_ _≤₂_)
+ ×-total {_≈₁_ = _≈₁_} {_<₁_ = _<₁_} total₁ {_≤₂_ = _≤₂_} = total
where
- total : Total (×-Lex ≈₁ <₁ ≤₂)
+ total : Total (×-Lex _≈₁_ _<₁_ _≤₂_)
total x y with total₁ (proj₁ x) (proj₁ y)
... | inj₁ x₁<y₁ = inj₁ (inj₁ x₁<y₁)
... | inj₂ x₁>y₁ = inj₂ (inj₁ x₁>y₁)
- ×-compare : ∀ {≈₁ <₁} → Symmetric ≈₁ → Trichotomous ≈₁ <₁ →
- ∀ {≈₂ <₂} → Trichotomous ≈₂ <₂ →
- Trichotomous (≈₁ ×-Rel ≈₂) (×-Lex ≈₁ <₁ <₂)
- ×-compare {≈₁} {<₁} sym₁ compare₁ {≈₂} {<₂} compare₂ = cmp
+ ×-compare : ∀ {_≈₁_ _<₁_} → Symmetric _≈₁_ → Trichotomous _≈₁_ _<₁_ →
+ ∀ {_≈₂_ _<₂_} → Trichotomous _≈₂_ _<₂_ →
+ Trichotomous (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_)
+ ×-compare {_≈₁_} {_<₁_} sym₁ compare₁ {_≈₂_} {_<₂_} compare₂ = cmp
where
- cmp : Trichotomous (≈₁ ×-Rel ≈₂) (×-Lex ≈₁ <₁ <₂)
- cmp (x₁ , x₂) (y₁ , y₂) with compare₁ x₁ y₁
- ... | tri< x₁<y₁ x₁≉y₁ x₁≯y₁ = tri< (inj₁ x₁<y₁) (x₁≉y₁ ∘ proj₁)
- [ x₁≯y₁ , x₁≉y₁ ∘ sym₁ ∘ proj₁ ]
- ... | tri> x₁≮y₁ x₁≉y₁ x₁>y₁ = tri> [ x₁≮y₁ , x₁≉y₁ ∘ proj₁ ]
- (x₁≉y₁ ∘ proj₁) (inj₁ x₁>y₁)
- ... | tri≈ x₁≮y₁ x₁≈y₁ x₁≯y₁ with compare₂ x₂ y₂
- ... | tri< x₂<y₂ x₂≉y₂ x₂≯y₂ = tri< (inj₂ (x₁≈y₁ , x₂<y₂))
- (x₂≉y₂ ∘ proj₂)
- [ x₁≯y₁ , x₂≯y₂ ∘ proj₂ ]
- ... | tri> x₂≮y₂ x₂≉y₂ x₂>y₂ = tri> [ x₁≮y₁ , x₂≮y₂ ∘ proj₂ ]
- (x₂≉y₂ ∘ proj₂)
- (inj₂ (sym₁ x₁≈y₁ , x₂>y₂))
- ... | tri≈ x₂≮y₂ x₂≈y₂ x₂≯y₂ = tri≈ [ x₁≮y₁ , x₂≮y₂ ∘ proj₂ ]
- (x₁≈y₁ , x₂≈y₂)
- [ x₁≯y₁ , x₂≯y₂ ∘ proj₂ ]
+ cmp″ : ∀ {x₁ y₁ x₂ y₂} →
+ ¬ (x₁ <₁ y₁) → x₁ ≈₁ y₁ → ¬ (y₁ <₁ x₁) →
+ Tri (x₂ <₂ y₂) (x₂ ≈₂ y₂) (y₂ <₂ x₂) →
+ Tri (×-Lex _≈₁_ _<₁_ _<₂_ (x₁ , x₂) (y₁ , y₂))
+ ((_≈₁_ ×-Rel _≈₂_) (x₁ , x₂) (y₁ , y₂))
+ (×-Lex _≈₁_ _<₁_ _<₂_ (y₁ , y₂) (x₁ , x₂))
+ cmp″ x₁≮y₁ x₁≈y₁ x₁≯y₁ (tri< x₂<y₂ x₂≉y₂ x₂≯y₂) =
+ tri< (inj₂ (x₁≈y₁ , x₂<y₂))
+ (x₂≉y₂ ∘ proj₂)
+ [ x₁≯y₁ , x₂≯y₂ ∘ proj₂ ]
+ cmp″ x₁≮y₁ x₁≈y₁ x₁≯y₁ (tri> x₂≮y₂ x₂≉y₂ x₂>y₂) =
+ tri> [ x₁≮y₁ , x₂≮y₂ ∘ proj₂ ]
+ (x₂≉y₂ ∘ proj₂)
+ (inj₂ (sym₁ x₁≈y₁ , x₂>y₂))
+ cmp″ x₁≮y₁ x₁≈y₁ x₁≯y₁ (tri≈ x₂≮y₂ x₂≈y₂ x₂≯y₂) =
+ tri≈ [ x₁≮y₁ , x₂≮y₂ ∘ proj₂ ]
+ (x₁≈y₁ , x₂≈y₂)
+ [ x₁≯y₁ , x₂≯y₂ ∘ proj₂ ]
+
+ cmp′ : ∀ {x₁ y₁} → Tri (x₁ <₁ y₁) (x₁ ≈₁ y₁) (y₁ <₁ x₁) →
+ ∀ x₂ y₂ →
+ Tri (×-Lex _≈₁_ _<₁_ _<₂_ (x₁ , x₂) (y₁ , y₂))
+ ((_≈₁_ ×-Rel _≈₂_) (x₁ , x₂) (y₁ , y₂))
+ (×-Lex _≈₁_ _<₁_ _<₂_ (y₁ , y₂) (x₁ , x₂))
+ cmp′ (tri< x₁<y₁ x₁≉y₁ x₁≯y₁) x₂ y₂ =
+ tri< (inj₁ x₁<y₁) (x₁≉y₁ ∘ proj₁) [ x₁≯y₁ , x₁≉y₁ ∘ sym₁ ∘ proj₁ ]
+ cmp′ (tri> x₁≮y₁ x₁≉y₁ x₁>y₁) x₂ y₂ =
+ tri> [ x₁≮y₁ , x₁≉y₁ ∘ proj₁ ] (x₁≉y₁ ∘ proj₁) (inj₁ x₁>y₁)
+ cmp′ (tri≈ x₁≮y₁ x₁≈y₁ x₁≯y₁) x₂ y₂ =
+ cmp″ x₁≮y₁ x₁≈y₁ x₁≯y₁ (compare₂ x₂ y₂)
+
+ cmp : Trichotomous (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_)
+ cmp (x₁ , x₂) (y₁ , y₂) = cmp′ (compare₁ x₁ y₁) x₂ y₂
-- Some collections of properties which are preserved by ×-Lex.
- _×-isPreorder_ : ∀ {≈₁ ∼₁} → IsPreorder ≈₁ ∼₁ →
- ∀ {≈₂ ∼₂} → IsPreorder ≈₂ ∼₂ →
- IsPreorder (≈₁ ×-Rel ≈₂) (×-Lex ≈₁ ∼₁ ∼₂)
- _×-isPreorder_ {≈₁ = ≈₁} {∼₁ = ∼₁} pre₁ {∼₂ = ∼₂} pre₂ = record
- { isEquivalence = Pointwise._×-isEquivalence_
- (isEquivalence pre₁) (isEquivalence pre₂)
- ; reflexive = λ {x y} →
- ×-reflexive ≈₁ ∼₁ ∼₂ (reflexive pre₂) {x} {y}
- ; trans = λ {x y z} →
- ×-transitive
- (isEquivalence pre₁) (∼-resp-≈ pre₁) (trans pre₁)
- {≤₂ = ∼₂} (trans pre₂) {x} {y} {z}
- ; ∼-resp-≈ = ×-≈-respects₂ (isEquivalence pre₁)
- (∼-resp-≈ pre₁)
- (∼-resp-≈ pre₂)
- }
+ _×-isPreorder_ : ∀ {_≈₁_ _∼₁_} → IsPreorder _≈₁_ _∼₁_ →
+ ∀ {_≈₂_ _∼₂_} → IsPreorder _≈₂_ _∼₂_ →
+ IsPreorder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _∼₁_ _∼₂_)
+ _×-isPreorder_ {_≈₁_ = _≈₁_} {_∼₁_ = _∼₁_} pre₁ {_∼₂_ = _∼₂_} pre₂ =
+ record
+ { isEquivalence = Pointwise._×-isEquivalence_
+ (isEquivalence pre₁) (isEquivalence pre₂)
+ ; reflexive = λ {x y} →
+ ×-reflexive _≈₁_ _∼₁_ _∼₂_ (reflexive pre₂)
+ {x} {y}
+ ; trans = λ {x y z} →
+ ×-transitive
+ (isEquivalence pre₁) (∼-resp-≈ pre₁)
+ (trans pre₁) {_≤₂_ = _∼₂_} (trans pre₂)
+ {x} {y} {z}
+ }
where open IsPreorder
_×-isStrictPartialOrder_ :
- ∀ {≈₁ <₁} → IsStrictPartialOrder ≈₁ <₁ →
- ∀ {≈₂ <₂} → IsStrictPartialOrder ≈₂ <₂ →
- IsStrictPartialOrder (≈₁ ×-Rel ≈₂) (×-Lex ≈₁ <₁ <₂)
- _×-isStrictPartialOrder_ {<₁ = <₁} spo₁ {<₂ = <₂} spo₂ =
+ ∀ {_≈₁_ _<₁_} → IsStrictPartialOrder _≈₁_ _<₁_ →
+ ∀ {_≈₂_ _<₂_} → IsStrictPartialOrder _≈₂_ _<₂_ →
+ IsStrictPartialOrder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_)
+ _×-isStrictPartialOrder_ {_<₁_ = _<₁_} spo₁ {_<₂_ = _<₂_} spo₂ =
record
{ isEquivalence = Pointwise._×-isEquivalence_
(isEquivalence spo₁) (isEquivalence spo₂)
; irrefl = λ {x y} →
- _×-irreflexive_ {<₁ = <₁} (irrefl spo₁)
- {<₂ = <₂} (irrefl spo₂) {x} {y}
+ _×-irreflexive_ {_<₁_ = _<₁_} (irrefl spo₁)
+ {_<₂_ = _<₂_} (irrefl spo₂)
+ {x} {y}
; trans = λ {x y z} →
×-transitive
- {<₁ = <₁} (isEquivalence spo₁) (<-resp-≈ spo₁)
- (trans spo₁)
- {≤₂ = <₂} (trans spo₂) {x} {y} {z}
+ {_<₁_ = _<₁_} (isEquivalence spo₁)
+ (<-resp-≈ spo₁) (trans spo₁)
+ {_≤₂_ = _<₂_} (trans spo₂) {x} {y} {z}
; <-resp-≈ = ×-≈-respects₂ (isEquivalence spo₁)
- (<-resp-≈ spo₁)
- (<-resp-≈ spo₂)
+ (<-resp-≈ spo₁)
+ (<-resp-≈ spo₂)
}
where open IsStrictPartialOrder
_×-isStrictTotalOrder_ :
- ∀ {≈₁ <₁} → IsStrictTotalOrder ≈₁ <₁ →
- ∀ {≈₂ <₂} → IsStrictTotalOrder ≈₂ <₂ →
- IsStrictTotalOrder (≈₁ ×-Rel ≈₂) (×-Lex ≈₁ <₁ <₂)
- _×-isStrictTotalOrder_ {<₁ = <₁} spo₁ {<₂ = <₂} spo₂ =
+ ∀ {_≈₁_ _<₁_} → IsStrictTotalOrder _≈₁_ _<₁_ →
+ ∀ {_≈₂_ _<₂_} → IsStrictTotalOrder _≈₂_ _<₂_ →
+ IsStrictTotalOrder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_)
+ _×-isStrictTotalOrder_ {_<₁_ = _<₁_} spo₁ {_<₂_ = _<₂_} spo₂ =
record
{ isEquivalence = Pointwise._×-isEquivalence_
(isEquivalence spo₁) (isEquivalence spo₂)
; trans = λ {x y z} →
×-transitive
- {<₁ = <₁} (isEquivalence spo₁) (<-resp-≈ spo₁)
- (trans spo₁)
- {≤₂ = <₂} (trans spo₂) {x} {y} {z}
+ {_<₁_ = _<₁_} (isEquivalence spo₁)
+ (<-resp-≈ spo₁) (trans spo₁)
+ {_≤₂_ = _<₂_} (trans spo₂) {x} {y} {z}
; compare = ×-compare (Eq.sym spo₁) (compare spo₁)
(compare spo₂)
; <-resp-≈ = ×-≈-respects₂ (isEquivalence spo₁)
@@ -228,13 +256,16 @@ open Dummy public
-- "Packages" (e.g. preorders) can also be combined.
-_×-preorder_ : Preorder _ _ _ → Preorder _ _ _ → Preorder _ _ _
+_×-preorder_ :
+ ∀ {p₁ p₂ p₃ p₄} →
+ Preorder p₁ p₂ _ → Preorder p₃ p₄ _ → Preorder _ _ _
p₁ ×-preorder p₂ = record
{ isPreorder = isPreorder p₁ ×-isPreorder isPreorder p₂
} where open Preorder
_×-strictPartialOrder_ :
- StrictPartialOrder _ _ _ → StrictPartialOrder _ _ _ →
+ ∀ {s₁ s₂ s₃ s₄} →
+ StrictPartialOrder s₁ s₂ _ → StrictPartialOrder s₃ s₄ _ →
StrictPartialOrder _ _ _
s₁ ×-strictPartialOrder s₂ = record
{ isStrictPartialOrder = isStrictPartialOrder s₁
@@ -243,7 +274,8 @@ s₁ ×-strictPartialOrder s₂ = record
} where open StrictPartialOrder
_×-strictTotalOrder_ :
- StrictTotalOrder _ _ _ → StrictTotalOrder _ _ _ →
+ ∀ {s₁ s₂ s₃ s₄} →
+ StrictTotalOrder s₁ s₂ _ → StrictTotalOrder s₃ s₄ _ →
StrictTotalOrder _ _ _
s₁ ×-strictTotalOrder s₂ = record
{ isStrictTotalOrder = isStrictTotalOrder s₁
diff --git a/src/Relation/Binary/PropositionalEquality.agda b/src/Relation/Binary/PropositionalEquality.agda
index b14fad0..680a3b1 100644
--- a/src/Relation/Binary/PropositionalEquality.agda
+++ b/src/Relation/Binary/PropositionalEquality.agda
@@ -6,11 +6,12 @@
module Relation.Binary.PropositionalEquality where
-open import Data.Function
-open import Data.Function.Equality using (_≡⇨_; _⟶_)
+open import Function
+open import Function.Equality using (Π; _⟶_; ≡-setoid)
open import Data.Product
open import Level
open import Relation.Binary
+import Relation.Binary.Indexed as I
open import Relation.Binary.Consequences
-- Some of the definitions can be found in the following modules:
@@ -57,7 +58,6 @@ isPreorder = record
{ isEquivalence = isEquivalence
; reflexive = id
; trans = trans
- ; ∼-resp-≈ = resp₂ _≡_
}
preorder : ∀ {a} → Set a → Preorder _ _ _
@@ -74,15 +74,23 @@ preorder A = record
infix 4 _≗_
_→-setoid_ : ∀ {a b} (A : Set a) (B : Set b) → Setoid _ _
-A →-setoid B = A ≡⇨ λ _ → setoid B
+A →-setoid B = ≡-setoid A (Setoid.indexedSetoid (setoid B))
_≗_ : ∀ {a b} {A : Set a} {B : Set b} (f g : A → B) → Set _
_≗_ {A = A} {B} = Setoid._≈_ (A →-setoid B)
+:→-to-Π : ∀ {a b₁ b₂} {A : Set a} {B : I.Setoid _ b₁ b₂} →
+ ((x : A) → I.Setoid.Carrier B x) → Π (setoid A) B
+:→-to-Π {B = B} f = record { _⟨$⟩_ = f; cong = cong′ }
+ where
+ open I.Setoid B using (_≈_)
+
+ cong′ : ∀ {x y} → x ≡ y → f x ≈ f y
+ cong′ refl = I.Setoid.refl B
+
→-to-⟶ : ∀ {a b₁ b₂} {A : Set a} {B : Setoid b₁ b₂} →
(A → Setoid.Carrier B) → setoid A ⟶ B
-→-to-⟶ {B = B} f =
- record { _⟨$⟩_ = f; cong = Setoid.reflexive B ∘ cong f }
+→-to-⟶ = :→-to-Π
------------------------------------------------------------------------
-- The inspect idiom
@@ -92,7 +100,7 @@ _≗_ {A = A} {B} = Setoid._≈_ (A →-setoid B)
-- r ≡ e.
data Inspect {a} {A : Set a} (x : A) : Set a where
- _with-≡_ : (y : A) (eq : y ≡ x) → Inspect x
+ _with-≡_ : (y : A) (eq : x ≡ y) → Inspect x
inspect : ∀ {a} {A : Set a} (x : A) → Inspect x
inspect x = x with-≡ refl
@@ -118,3 +126,14 @@ module ≡-Reasoning where
open EqR (setoid A) public
hiding (_≡⟨_⟩_) renaming (_≈⟨_⟩_ to _≡⟨_⟩_)
open Dummy public
+
+------------------------------------------------------------------------
+-- Definition of functional extensionality
+
+-- If _≡_ were extensional, then the following statement could be
+-- proved.
+
+Extensionality : (a b : Level) → Set _
+Extensionality a b =
+ {A : Set a} {B : A → Set b} {f g : (x : A) → B x} →
+ (∀ x → f x ≡ g x) → f ≡ g
diff --git a/src/Relation/Binary/Props/Preorder.agda b/src/Relation/Binary/Props/Preorder.agda
index 5813dad..82bfe38 100644
--- a/src/Relation/Binary/Props/Preorder.agda
+++ b/src/Relation/Binary/Props/Preorder.agda
@@ -9,7 +9,7 @@ open import Relation.Binary
module Relation.Binary.Props.Preorder
{p₁ p₂ p₃} (P : Preorder p₁ p₂ p₃) where
-open import Data.Function
+open import Function
open import Data.Product as Prod
open Relation.Binary.Preorder P
@@ -17,8 +17,8 @@ open Relation.Binary.Preorder P
------------------------------------------------------------------------
-- For every preorder there is an induced equivalence
-inducedEquivalence : Setoid _ _
-inducedEquivalence = record
+InducedEquivalence : Setoid _ _
+InducedEquivalence = record
{ _≈_ = λ x y → x ∼ y × y ∼ x
; isEquivalence = record
{ refl = (refl , refl)
diff --git a/src/Relation/Binary/Props/StrictPartialOrder.agda b/src/Relation/Binary/Props/StrictPartialOrder.agda
index 8db8d4b..9c65caa 100644
--- a/src/Relation/Binary/Props/StrictPartialOrder.agda
+++ b/src/Relation/Binary/Props/StrictPartialOrder.agda
@@ -24,7 +24,6 @@ poset = record
{ isEquivalence = isEquivalence
; reflexive = reflexive
; trans = trans isEquivalence <-resp-≈ <-trans
- ; ∼-resp-≈ = ≤-resp-≈ isEquivalence <-resp-≈
}
; antisym = antisym isEquivalence <-trans irrefl
}
diff --git a/src/Relation/Binary/Reflection.agda b/src/Relation/Binary/Reflection.agda
index 496eb30..149f897 100644
--- a/src/Relation/Binary/Reflection.agda
+++ b/src/Relation/Binary/Reflection.agda
@@ -6,7 +6,7 @@
{-# OPTIONS --universe-polymorphism #-}
open import Data.Fin
-open import Data.Function
+open import Function
open import Data.Nat
open import Data.Vec as Vec
open import Level
diff --git a/src/Relation/Binary/Sigma/Pointwise.agda b/src/Relation/Binary/Sigma/Pointwise.agda
new file mode 100644
index 0000000..dfac94e
--- /dev/null
+++ b/src/Relation/Binary/Sigma/Pointwise.agda
@@ -0,0 +1,179 @@
+------------------------------------------------------------------------
+-- Pointwise lifting of binary relations to sigma types
+------------------------------------------------------------------------
+
+{-# OPTIONS --universe-polymorphism #-}
+
+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.Equivalence as Eq
+ using (Equivalent; _⇔_; module Equivalent)
+ renaming (_∘_ to _⟨∘⟩_)
+open import Function.Inverse as Inv
+ using (Inverse; _⇿_; module Inverse; Isomorphism)
+ renaming (_∘_ to _⟪∘⟫_)
+open import Function.LeftInverse
+ using (_LeftInverseOf_; _RightInverseOf_)
+import Relation.Binary as B
+open import Relation.Binary.Indexed as I using (_at_)
+import Relation.Binary.HeterogeneousEquality as H
+import Relation.Binary.PropositionalEquality as P
+
+------------------------------------------------------------------------
+-- Pointwise lifting
+
+infixr 4 _,_
+
+data REL {a₁ a₂ b₁ b₂ ℓ₁ ℓ₂}
+ {A₁ : Set a₁} (B₁ : A₁ → Set b₁)
+ {A₂ : Set a₂} (B₂ : A₂ → Set b₂)
+ (_R₁_ : B.REL A₁ A₂ ℓ₁) (_R₂_ : I.REL B₁ B₂ ℓ₂) :
+ B.REL (Σ A₁ B₁) (Σ A₂ B₂) (a₁ ⊔ a₂ ⊔ b₁ ⊔ b₂ ⊔ ℓ₁ ⊔ ℓ₂) where
+ _,_ : {x₁ : A₁} {y₁ : B₁ x₁} {x₂ : A₂} {y₂ : B₂ x₂}
+ (x₁Rx₂ : x₁ R₁ x₂) (y₁Ry₂ : y₁ R₂ y₂) →
+ REL B₁ B₂ _R₁_ _R₂_ (x₁ , y₁) (x₂ , y₂)
+
+Rel : ∀ {a b ℓ₁ ℓ₂} {A : Set a} (B : A → Set b)
+ (_R₁_ : B.Rel A ℓ₁) (_R₂_ : I.Rel B ℓ₂) → B.Rel (Σ A B) _
+Rel B = REL B B
+
+------------------------------------------------------------------------
+-- Rel preserves many properties
+
+private
+ module Dummy {a b ℓ₁ ℓ₂} {A : Set a} {B : A → Set b}
+ {R₁ : B.Rel A ℓ₁} {R₂ : I.Rel B ℓ₂} where
+
+ refl : B.Reflexive R₁ → I.Reflexive B R₂ →
+ B.Reflexive (Rel B R₁ R₂)
+ refl refl₁ refl₂ {x = (x , y)} = (refl₁ , refl₂)
+
+ symmetric : B.Symmetric R₁ → I.Symmetric B R₂ →
+ B.Symmetric (Rel B R₁ R₂)
+ symmetric sym₁ sym₂ (x₁Rx₂ , y₁Ry₂) = (sym₁ x₁Rx₂ , sym₂ y₁Ry₂)
+
+ transitive : B.Transitive R₁ → I.Transitive B R₂ →
+ B.Transitive (Rel B R₁ R₂)
+ transitive trans₁ trans₂ (x₁Rx₂ , y₁Ry₂) (x₂Rx₃ , y₂Ry₃) =
+ (trans₁ x₁Rx₂ x₂Rx₃ , trans₂ y₁Ry₂ y₂Ry₃)
+
+ isEquivalence : B.IsEquivalence R₁ → I.IsEquivalence B R₂ →
+ B.IsEquivalence (Rel B R₁ R₂)
+ isEquivalence eq₁ eq₂ = record
+ { refl = refl (B.IsEquivalence.refl eq₁)
+ (I.IsEquivalence.refl eq₂)
+ ; sym = symmetric (B.IsEquivalence.sym eq₁)
+ (I.IsEquivalence.sym eq₂)
+ ; trans = transitive (B.IsEquivalence.trans eq₁)
+ (I.IsEquivalence.trans eq₂)
+ }
+
+open Dummy public
+
+setoid : ∀ {b₁ b₂ i₁ i₂} →
+ (A : B.Setoid b₁ b₂) → I.Setoid (B.Setoid.Carrier A) i₁ i₂ →
+ B.Setoid _ _
+setoid s₁ s₂ = record
+ { isEquivalence = isEquivalence (B.Setoid.isEquivalence s₁)
+ (I.Setoid.isEquivalence s₂)
+ }
+
+------------------------------------------------------------------------
+-- The propositional equality setoid over sigma types can be
+-- decomposed using Rel
+
+Rel⇿≡ : ∀ {a b} {A : Set a} {B : A → Set b} →
+ Inverse (setoid (P.setoid A) (H.indexedSetoid B))
+ (P.setoid (Σ A B))
+Rel⇿≡ {a} {b} {A} {B} = record
+ { to = record { _⟨$⟩_ = id; cong = to-cong }
+ ; from = record { _⟨$⟩_ = id; cong = from-cong }
+ ; inverse-of = record
+ { left-inverse-of = uncurry (λ _ _ → (P.refl , H.refl))
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+ where
+ open I using (_=[_]⇒_)
+
+ to-cong : Rel B P._≡_ (λ x y → H._≅_ x y) =[ id {a = a ⊔ b} ]⇒ P._≡_
+ to-cong (P.refl , H.refl) = P.refl
+
+ from-cong : P._≡_ =[ id {a = a ⊔ b} ]⇒ Rel B P._≡_ (λ x y → H._≅_ x y)
+ from-cong {i = (x , y)} P.refl = (P.refl , H.refl)
+
+------------------------------------------------------------------------
+-- 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
+ { 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 using (_=[_]⇒_)
+
+ to = Prod.map id (F._⟨$⟩_ (Equivalent.to F⇔T))
+
+ to-cong : _≈F_ =[ to ]⇒ _≈T_
+ to-cong (P.refl , ∼) = (P.refl , F.cong (Equivalent.to F⇔T) ∼)
+
+ from = Prod.map id (F._⟨$⟩_ (Equivalent.from F⇔T))
+
+ from-cong : _≈T_ =[ from ]⇒ _≈F_
+ from-cong (P.refl , ∼) = (P.refl , F.cong (Equivalent.from F⇔T) ∼)
+
+⇔ : ∀ {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₂ =
+ Inverse.equivalent (Rel⇿≡ {B = B₂}) ⟨∘⟩
+ equivalent (λ {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
+ { to = Equivalent.to eq
+ ; from = Equivalent.from eq
+ ; inverse-of = record
+ { left-inverse-of = left
+ ; right-inverse-of = right
+ }
+ }
+ where
+ eq = equivalent (Inverse.equivalent F⇿T)
+
+ left : Equivalent.from eq LeftInverseOf Equivalent.to eq
+ left (x , y) = (P.refl , Inverse.left-inverse-of F⇿T y)
+
+ right : Equivalent.from eq RightInverseOf Equivalent.to eq
+ right (x , y) = (P.refl , Inverse.right-inverse-of F⇿T y)
+
+⇿ : ∀ {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₂ =
+ Rel⇿≡ {B = B₂} ⟪∘⟫
+ inverse (λ {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 = Inv.inverse} = ⇿
diff --git a/src/Relation/Binary/Simple.agda b/src/Relation/Binary/Simple.agda
index 3b17b91..a1e6da4 100644
--- a/src/Relation/Binary/Simple.agda
+++ b/src/Relation/Binary/Simple.agda
@@ -21,13 +21,14 @@ Const I = λ _ _ → I
Always : ∀ {a} {A : Set a} → Rel A zero
Always = Const ⊤
+Always-setoid : ∀ {a} (A : Set a) → Setoid _ _
+Always-setoid A = record
+ { Carrier = A
+ ; _≈_ = Always
+ ; isEquivalence = record {}
+ }
+
-- The universally false relation.
Never : ∀ {a} {A : Set a} → Rel A zero
Never = Const ⊥
-
--- Always is an equivalence.
-
-Always-isEquivalence : ∀ {a} {A : Set a} →
- IsEquivalence (Always {A = A})
-Always-isEquivalence = record {}
diff --git a/src/Relation/Binary/StrictToNonStrict.agda b/src/Relation/Binary/StrictToNonStrict.agda
index eb34d61..be7bb1f 100644
--- a/src/Relation/Binary/StrictToNonStrict.agda
+++ b/src/Relation/Binary/StrictToNonStrict.agda
@@ -17,7 +17,7 @@ module Relation.Binary.StrictToNonStrict
open import Relation.Nullary
open import Relation.Binary.Consequences
-open import Data.Function
+open import Function
open import Data.Product
open import Data.Sum
open import Data.Empty
@@ -49,7 +49,7 @@ antisym eq trans irrefl = as
as (inj₂ x≈y) _ = x≈y
as (inj₁ _) (inj₂ y≈x) = Eq.sym y≈x
as (inj₁ x<y) (inj₁ y<x) =
- ⊥-elim (trans∧irr⟶asym {≈ = _≈_} Eq.refl trans irrefl x<y y<x)
+ ⊥-elim (trans∧irr⟶asym {_≈_ = _≈_} Eq.refl trans irrefl x<y y<x)
trans : IsEquivalence _≈_ → _<_ Respects₂ _≈_ →
Transitive _<_ → Transitive _≤_
diff --git a/src/Relation/Binary/Sum.agda b/src/Relation/Binary/Sum.agda
index 7d31a54..6f40f9c 100644
--- a/src/Relation/Binary/Sum.agda
+++ b/src/Relation/Binary/Sum.agda
@@ -2,19 +2,29 @@
-- Sums of binary relations
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Relation.Binary.Sum where
-open import Data.Function
open import Data.Sum as Sum
open import Data.Product
open import Data.Unit using (⊤)
open import Data.Empty
+open import Function
+open import Function.Equality as F using (_⟨$⟩_)
+open import Function.Equivalence as Eq
+ using (Equivalent; _⇔_; module Equivalent)
+ renaming (_∘_ to _⟨∘⟩_)
+open import Function.Inverse as Inv
+ using (Inverse; _⇿_; module Inverse)
+ renaming (_∘_ to _⟪∘⟫_)
open import Level
open import Relation.Nullary
open import Relation.Binary
+import Relation.Binary.PropositionalEquality as P
private
- module Dummy {A₁ A₂ : Set} where
+ module Dummy {a₁ a₂} {A₁ : Set a₁} {A₂ : Set a₂} where
----------------------------------------------------------------------
-- Sums of relations
@@ -23,21 +33,21 @@ private
-- Generalised sum.
- data ⊎ʳ (P : Set) (_∼₁_ : Rel A₁ zero) (_∼₂_ : Rel A₂ zero) :
- A₁ ⊎ A₂ → A₁ ⊎ A₂ → Set where
+ data ⊎ʳ {ℓ₁ ℓ₂} (P : Set) (_∼₁_ : Rel A₁ ℓ₁) (_∼₂_ : Rel A₂ ℓ₂) :
+ A₁ ⊎ A₂ → A₁ ⊎ A₂ → Set (a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂) where
₁∼₂ : ∀ {x y} (p : P) → ⊎ʳ P _∼₁_ _∼₂_ (inj₁ x) (inj₂ y)
₁∼₁ : ∀ {x y} (x∼₁y : x ∼₁ y) → ⊎ʳ P _∼₁_ _∼₂_ (inj₁ x) (inj₁ y)
₂∼₂ : ∀ {x y} (x∼₂y : x ∼₂ y) → ⊎ʳ P _∼₁_ _∼₂_ (inj₂ x) (inj₂ y)
-- Pointwise sum.
- _⊎-Rel_ : Rel A₁ zero → Rel A₂ zero → Rel (A₁ ⊎ A₂) zero
+ _⊎-Rel_ : ∀ {ℓ₁ ℓ₂} → Rel A₁ ℓ₁ → Rel A₂ ℓ₂ → Rel (A₁ ⊎ A₂) _
_⊎-Rel_ = ⊎ʳ ⊥
- -- All things to the left are smaller than (or equal to, depending
- -- on the underlying equality) all things to the right.
+ -- All things to the left are "smaller than" all things to the
+ -- right.
- _⊎-<_ : Rel A₁ zero → Rel A₂ zero → Rel (A₁ ⊎ A₂) zero
+ _⊎-<_ : ∀ {ℓ₁ ℓ₂} → Rel A₁ ℓ₁ → Rel A₂ ℓ₂ → Rel (A₁ ⊎ A₂) _
_⊎-<_ = ⊎ʳ ⊤
----------------------------------------------------------------------
@@ -45,23 +55,24 @@ private
private
- ₁≁₂ : {∼₁ : Rel A₁ zero} {∼₂ : Rel A₂ zero} →
+ ₁≁₂ : ∀ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} →
∀ {x y} → ¬ (inj₁ x ⟨ ∼₁ ⊎-Rel ∼₂ ⟩ inj₂ y)
₁≁₂ (₁∼₂ ())
- drop-inj₁ : {∼₁ : Rel A₁ zero} {∼₂ : Rel A₂ zero} →
+ drop-inj₁ : ∀ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} →
∀ {P x y} → inj₁ x ⟨ ⊎ʳ P ∼₁ ∼₂ ⟩ inj₁ y → ∼₁ x y
drop-inj₁ (₁∼₁ x∼y) = x∼y
- drop-inj₂ : {∼₁ : Rel A₁ zero} {∼₂ : Rel A₂ zero} →
+ drop-inj₂ : ∀ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} →
∀ {P x y} → inj₂ x ⟨ ⊎ʳ P ∼₁ ∼₂ ⟩ inj₂ y → ∼₂ x y
drop-inj₂ (₂∼₂ x∼y) = x∼y
----------------------------------------------------------------------
-- Some properties which are preserved by the relation formers above
- _⊎-reflexive_ : {≈₁ ∼₁ : Rel A₁ zero} → ≈₁ ⇒ ∼₁ →
- {≈₂ ∼₂ : Rel A₂ zero} → ≈₂ ⇒ ∼₂ →
+ _⊎-reflexive_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {∼₁ : Rel A₁ ℓ₁′}
+ {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {∼₂ : Rel A₂ ℓ₂′} →
+ ≈₁ ⇒ ∼₁ → ≈₂ ⇒ ∼₂ →
∀ {P} → (≈₁ ⊎-Rel ≈₂) ⇒ (⊎ʳ P ∼₁ ∼₂)
refl₁ ⊎-reflexive refl₂ = refl
where
@@ -70,17 +81,17 @@ private
refl (₂∼₂ x₂≈y₂) = ₂∼₂ (refl₂ x₂≈y₂)
refl (₁∼₂ ())
- _⊎-refl_ : {∼₁ : Rel A₁ zero} → Reflexive ∼₁ →
- {∼₂ : Rel A₂ zero} → Reflexive ∼₂ →
- Reflexive (∼₁ ⊎-Rel ∼₂)
+ _⊎-refl_ : ∀ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} →
+ Reflexive ∼₁ → Reflexive ∼₂ → Reflexive (∼₁ ⊎-Rel ∼₂)
refl₁ ⊎-refl refl₂ = refl
where
refl : Reflexive (_ ⊎-Rel _)
refl {x = inj₁ _} = ₁∼₁ refl₁
refl {x = inj₂ _} = ₂∼₂ refl₂
- _⊎-irreflexive_ : {≈₁ <₁ : Rel A₁ zero} → Irreflexive ≈₁ <₁ →
- {≈₂ <₂ : Rel A₂ zero} → Irreflexive ≈₂ <₂ →
+ _⊎-irreflexive_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {<₁ : Rel A₁ ℓ₁′}
+ {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {<₂ : Rel A₂ ℓ₂′} →
+ Irreflexive ≈₁ <₁ → Irreflexive ≈₂ <₂ →
∀ {P} → Irreflexive (≈₁ ⊎-Rel ≈₂) (⊎ʳ P <₁ <₂)
irrefl₁ ⊎-irreflexive irrefl₂ = irrefl
where
@@ -89,9 +100,8 @@ private
irrefl (₂∼₂ x₂≈y₂) (₂∼₂ x₂<y₂) = irrefl₂ x₂≈y₂ x₂<y₂
irrefl (₁∼₂ ()) _
- _⊎-symmetric_ : {∼₁ : Rel A₁ zero} → Symmetric ∼₁ →
- {∼₂ : Rel A₂ zero} → Symmetric ∼₂ →
- Symmetric (∼₁ ⊎-Rel ∼₂)
+ _⊎-symmetric_ : ∀ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} →
+ Symmetric ∼₁ → Symmetric ∼₂ → Symmetric (∼₁ ⊎-Rel ∼₂)
sym₁ ⊎-symmetric sym₂ = sym
where
sym : Symmetric (_ ⊎-Rel _)
@@ -99,8 +109,8 @@ private
sym (₂∼₂ x₂∼y₂) = ₂∼₂ (sym₂ x₂∼y₂)
sym (₁∼₂ ())
- _⊎-transitive_ : {∼₁ : Rel A₁ zero} → Transitive ∼₁ →
- {∼₂ : Rel A₂ zero} → Transitive ∼₂ →
+ _⊎-transitive_ : ∀ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} →
+ Transitive ∼₁ → Transitive ∼₂ →
∀ {P} → Transitive (⊎ʳ P ∼₁ ∼₂)
trans₁ ⊎-transitive trans₂ = trans
where
@@ -110,8 +120,9 @@ private
trans (₁∼₂ p) (₂∼₂ _) = ₁∼₂ p
trans (₁∼₁ _) (₁∼₂ p) = ₁∼₂ p
- _⊎-antisymmetric_ : {≈₁ ≤₁ : Rel A₁ zero} → Antisymmetric ≈₁ ≤₁ →
- {≈₂ ≤₂ : Rel A₂ zero} → Antisymmetric ≈₂ ≤₂ →
+ _⊎-antisymmetric_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {≤₁ : Rel A₁ ℓ₁′}
+ {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {≤₂ : Rel A₂ ℓ₂′} →
+ Antisymmetric ≈₁ ≤₁ → Antisymmetric ≈₂ ≤₂ →
∀ {P} → Antisymmetric (≈₁ ⊎-Rel ≈₂) (⊎ʳ P ≤₁ ≤₂)
antisym₁ ⊎-antisymmetric antisym₂ = antisym
where
@@ -120,8 +131,8 @@ private
antisym (₂∼₂ x≤y) (₂∼₂ y≤x) = ₂∼₂ (antisym₂ x≤y y≤x)
antisym (₁∼₂ _) ()
- _⊎-asymmetric_ : {<₁ : Rel A₁ zero} → Asymmetric <₁ →
- {<₂ : Rel A₂ zero} → Asymmetric <₂ →
+ _⊎-asymmetric_ : ∀ {ℓ₁ ℓ₂} {<₁ : Rel A₁ ℓ₁} {<₂ : Rel A₂ ℓ₂} →
+ Asymmetric <₁ → Asymmetric <₂ →
∀ {P} → Asymmetric (⊎ʳ P <₁ <₂)
asym₁ ⊎-asymmetric asym₂ = asym
where
@@ -130,11 +141,12 @@ private
asym (₂∼₂ x<y) (₂∼₂ y<x) = asym₂ x<y y<x
asym (₁∼₂ _) ()
- _⊎-≈-respects₂_ : {≈₁ ∼₁ : Rel A₁ zero} → ∼₁ Respects₂ ≈₁ →
- {≈₂ ∼₂ : Rel A₂ zero} → ∼₂ Respects₂ ≈₂ →
+ _⊎-≈-respects₂_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {∼₁ : Rel A₁ ℓ₁′}
+ {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {∼₂ : Rel A₂ ℓ₂′} →
+ ∼₁ Respects₂ ≈₁ → ∼₂ Respects₂ ≈₂ →
∀ {P} → (⊎ʳ P ∼₁ ∼₂) Respects₂ (≈₁ ⊎-Rel ≈₂)
- _⊎-≈-respects₂_ {≈₁ = ≈₁} {∼₁ = ∼₁} resp₁
- {≈₂ = ≈₂} {∼₂ = ∼₂} resp₂ {P} =
+ _⊎-≈-respects₂_ {≈₁ = ≈₁} {∼₁ = ∼₁}{≈₂ = ≈₂} {∼₂ = ∼₂}
+ resp₁ resp₂ {P} =
(λ {_ _ _} → resp¹) ,
(λ {_ _ _} → resp²)
where
@@ -151,21 +163,21 @@ private
resp² (₁∼₁ x≈x') (₁∼₂ p) = (₁∼₂ p)
resp² (₁∼₂ ()) _
- _⊎-substitutive_ : {∼₁ : Rel A₁ zero} → Substitutive ∼₁ zero →
- {∼₂ : Rel A₂ zero} → Substitutive ∼₂ zero →
- Substitutive (∼₁ ⊎-Rel ∼₂) zero
+ _⊎-substitutive_ : ∀ {ℓ₁ ℓ₂ ℓ₃} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} →
+ Substitutive ∼₁ ℓ₃ → Substitutive ∼₂ ℓ₃ →
+ Substitutive (∼₁ ⊎-Rel ∼₂) ℓ₃
subst₁ ⊎-substitutive subst₂ = subst
where
- subst : Substitutive (_ ⊎-Rel _) zero
+ subst : Substitutive (_ ⊎-Rel _) _
subst P (₁∼₁ x∼y) Px = subst₁ (λ z → P (inj₁ z)) x∼y Px
subst P (₂∼₂ x∼y) Px = subst₂ (λ z → P (inj₂ z)) x∼y Px
subst P (₁∼₂ ()) Px
- ⊎-decidable : {∼₁ : Rel A₁ zero} → Decidable ∼₁ →
- {∼₂ : Rel A₂ zero} → Decidable ∼₂ →
+ ⊎-decidable : ∀ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} →
+ Decidable ∼₁ → Decidable ∼₂ →
∀ {P} → (∀ {x y} → Dec (inj₁ x ⟨ ⊎ʳ P ∼₁ ∼₂ ⟩ inj₂ y)) →
Decidable (⊎ʳ P ∼₁ ∼₂)
- ⊎-decidable {∼₁ = ∼₁} dec₁ {∼₂ = ∼₂} dec₂ {P} dec₁₂ = dec
+ ⊎-decidable {∼₁ = ∼₁} {∼₂ = ∼₂} dec₁ dec₂ {P} dec₁₂ = dec
where
dec : Decidable (⊎ʳ P ∼₁ ∼₂)
dec (inj₁ x) (inj₁ y) with dec₁ x y
@@ -177,9 +189,8 @@ private
dec (inj₁ x) (inj₂ y) = dec₁₂
dec (inj₂ x) (inj₁ y) = no (λ())
- _⊎-<-total_ : {≤₁ : Rel A₁ zero} → Total ≤₁ →
- {≤₂ : Rel A₂ zero} → Total ≤₂ →
- Total (≤₁ ⊎-< ≤₂)
+ _⊎-<-total_ : ∀ {ℓ₁ ℓ₂} {≤₁ : Rel A₁ ℓ₁} {≤₂ : Rel A₂ ℓ₂} →
+ Total ≤₁ → Total ≤₂ → Total (≤₁ ⊎-< ≤₂)
total₁ ⊎-<-total total₂ = total
where
total : Total (_ ⊎-< _)
@@ -188,11 +199,12 @@ private
total (inj₁ x) (inj₂ y) = inj₁ (₁∼₂ _)
total (inj₂ x) (inj₁ y) = inj₂ (₁∼₂ _)
- _⊎-<-trichotomous_ : {≈₁ <₁ : Rel A₁ zero} → Trichotomous ≈₁ <₁ →
- {≈₂ <₂ : Rel A₂ zero} → Trichotomous ≈₂ <₂ →
+ _⊎-<-trichotomous_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {<₁ : Rel A₁ ℓ₁′}
+ {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {<₂ : Rel A₂ ℓ₂′} →
+ Trichotomous ≈₁ <₁ → Trichotomous ≈₂ <₂ →
Trichotomous (≈₁ ⊎-Rel ≈₂) (<₁ ⊎-< <₂)
- _⊎-<-trichotomous_ {≈₁ = ≈₁} {<₁ = <₁} tri₁
- {≈₂ = ≈₂} {<₂ = <₂} tri₂ = tri
+ _⊎-<-trichotomous_ {≈₁ = ≈₁} {<₁ = <₁} {≈₂ = ≈₂} {<₂ = <₂}
+ tri₁ tri₂ = tri
where
tri : Trichotomous (≈₁ ⊎-Rel ≈₂) (<₁ ⊎-< <₂)
tri (inj₁ x) (inj₂ y) = tri< (₁∼₂ _) ₁≁₂ (λ())
@@ -215,8 +227,8 @@ private
----------------------------------------------------------------------
-- Some collections of properties which are preserved
- _⊎-isEquivalence_ : {≈₁ : Rel A₁ zero} → IsEquivalence ≈₁ →
- {≈₂ : Rel A₂ zero} → IsEquivalence ≈₂ →
+ _⊎-isEquivalence_ : ∀ {ℓ₁ ℓ₂} {≈₁ : Rel A₁ ℓ₁} {≈₂ : Rel A₂ ℓ₂} →
+ IsEquivalence ≈₁ → IsEquivalence ≈₂ →
IsEquivalence (≈₁ ⊎-Rel ≈₂)
eq₁ ⊎-isEquivalence eq₂ = record
{ refl = refl eq₁ ⊎-refl refl eq₂
@@ -225,20 +237,20 @@ private
}
where open IsEquivalence
- _⊎-isPreorder_ : {≈₁ ∼₁ : Rel A₁ zero} → IsPreorder ≈₁ ∼₁ →
- {≈₂ ∼₂ : Rel A₂ zero} → IsPreorder ≈₂ ∼₂ →
+ _⊎-isPreorder_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {∼₁ : Rel A₁ ℓ₁′}
+ {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {∼₂ : Rel A₂ ℓ₂′} →
+ IsPreorder ≈₁ ∼₁ → IsPreorder ≈₂ ∼₂ →
∀ {P} → IsPreorder (≈₁ ⊎-Rel ≈₂) (⊎ʳ P ∼₁ ∼₂)
pre₁ ⊎-isPreorder pre₂ = record
{ isEquivalence = isEquivalence pre₁ ⊎-isEquivalence
isEquivalence pre₂
; reflexive = reflexive pre₁ ⊎-reflexive reflexive pre₂
; trans = trans pre₁ ⊎-transitive trans pre₂
- ; ∼-resp-≈ = ∼-resp-≈ pre₁ ⊎-≈-respects₂ ∼-resp-≈ pre₂
}
where open IsPreorder
- _⊎-isDecEquivalence_ : {≈₁ : Rel A₁ zero} → IsDecEquivalence ≈₁ →
- {≈₂ : Rel A₂ zero} → IsDecEquivalence ≈₂ →
+ _⊎-isDecEquivalence_ : ∀ {ℓ₁ ℓ₂} {≈₁ : Rel A₁ ℓ₁} {≈₂ : Rel A₂ ℓ₂} →
+ IsDecEquivalence ≈₁ → IsDecEquivalence ≈₂ →
IsDecEquivalence (≈₁ ⊎-Rel ≈₂)
eq₁ ⊎-isDecEquivalence eq₂ = record
{ isEquivalence = isEquivalence eq₁ ⊎-isEquivalence
@@ -247,8 +259,9 @@ private
}
where open IsDecEquivalence
- _⊎-isPartialOrder_ : {≈₁ ≤₁ : Rel A₁ zero} → IsPartialOrder ≈₁ ≤₁ →
- {≈₂ ≤₂ : Rel A₂ zero} → IsPartialOrder ≈₂ ≤₂ →
+ _⊎-isPartialOrder_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {≤₁ : Rel A₁ ℓ₁′}
+ {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {≤₂ : Rel A₂ ℓ₂′} →
+ IsPartialOrder ≈₁ ≤₁ → IsPartialOrder ≈₂ ≤₂ →
∀ {P} → IsPartialOrder (≈₁ ⊎-Rel ≈₂) (⊎ʳ P ≤₁ ≤₂)
po₁ ⊎-isPartialOrder po₂ = record
{ isPreorder = isPreorder po₁ ⊎-isPreorder isPreorder po₂
@@ -257,8 +270,9 @@ private
where open IsPartialOrder
_⊎-isStrictPartialOrder_ :
- {≈₁ <₁ : Rel A₁ zero} → IsStrictPartialOrder ≈₁ <₁ →
- {≈₂ <₂ : Rel A₂ zero} → IsStrictPartialOrder ≈₂ <₂ →
+ ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {<₁ : Rel A₁ ℓ₁′}
+ {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {<₂ : Rel A₂ ℓ₂′} →
+ IsStrictPartialOrder ≈₁ <₁ → IsStrictPartialOrder ≈₂ <₂ →
∀ {P} → IsStrictPartialOrder (≈₁ ⊎-Rel ≈₂) (⊎ʳ P <₁ <₂)
spo₁ ⊎-isStrictPartialOrder spo₂ = record
{ isEquivalence = isEquivalence spo₁ ⊎-isEquivalence
@@ -269,8 +283,9 @@ private
}
where open IsStrictPartialOrder
- _⊎-<-isTotalOrder_ : {≈₁ ≤₁ : Rel A₁ zero} → IsTotalOrder ≈₁ ≤₁ →
- {≈₂ ≤₂ : Rel A₂ zero} → IsTotalOrder ≈₂ ≤₂ →
+ _⊎-<-isTotalOrder_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {≤₁ : Rel A₁ ℓ₁′}
+ {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {≤₂ : Rel A₂ ℓ₂′} →
+ IsTotalOrder ≈₁ ≤₁ → IsTotalOrder ≈₂ ≤₂ →
IsTotalOrder (≈₁ ⊎-Rel ≈₂) (≤₁ ⊎-< ≤₂)
to₁ ⊎-<-isTotalOrder to₂ = record
{ isPartialOrder = isPartialOrder to₁ ⊎-isPartialOrder
@@ -280,8 +295,9 @@ private
where open IsTotalOrder
_⊎-<-isDecTotalOrder_ :
- {≈₁ ≤₁ : Rel A₁ zero} → IsDecTotalOrder ≈₁ ≤₁ →
- {≈₂ ≤₂ : Rel A₂ zero} → IsDecTotalOrder ≈₂ ≤₂ →
+ ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {≤₁ : Rel A₁ ℓ₁′}
+ {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {≤₂ : Rel A₂ ℓ₂′} →
+ IsDecTotalOrder ≈₁ ≤₁ → IsDecTotalOrder ≈₂ ≤₂ →
IsDecTotalOrder (≈₁ ⊎-Rel ≈₂) (≤₁ ⊎-< ≤₂)
to₁ ⊎-<-isDecTotalOrder to₂ = record
{ isTotalOrder = isTotalOrder to₁ ⊎-<-isTotalOrder isTotalOrder to₂
@@ -295,31 +311,36 @@ open Dummy public
------------------------------------------------------------------------
-- The game can be taken even further...
-_⊎-setoid_ : Setoid _ _ → Setoid _ _ → Setoid _ _
+_⊎-setoid_ : ∀ {s₁ s₂ s₃ s₄} →
+ Setoid s₁ s₂ → Setoid s₃ s₄ → Setoid _ _
s₁ ⊎-setoid s₂ = record
{ isEquivalence = isEquivalence s₁ ⊎-isEquivalence isEquivalence s₂
} where open Setoid
-_⊎-preorder_ : Preorder _ _ _ → Preorder _ _ _ → Preorder _ _ _
+_⊎-preorder_ : ∀ {p₁ p₂ p₃ p₄ p₅ p₆} →
+ Preorder p₁ p₂ p₃ → Preorder p₄ p₅ p₆ → Preorder _ _ _
p₁ ⊎-preorder p₂ = record
{ _∼_ = _∼_ p₁ ⊎-Rel _∼_ p₂
; isPreorder = isPreorder p₁ ⊎-isPreorder isPreorder p₂
} where open Preorder
-_⊎-decSetoid_ : DecSetoid _ _ → DecSetoid _ _ → DecSetoid _ _
+_⊎-decSetoid_ : ∀ {s₁ s₂ s₃ s₄} →
+ DecSetoid s₁ s₂ → DecSetoid s₃ s₄ → DecSetoid _ _
ds₁ ⊎-decSetoid ds₂ = record
{ isDecEquivalence = isDecEquivalence ds₁ ⊎-isDecEquivalence
isDecEquivalence ds₂
} where open DecSetoid
-_⊎-poset_ : Poset _ _ _ → Poset _ _ _ → Poset _ _ _
+_⊎-poset_ : ∀ {p₁ p₂ p₃ p₄ p₅ p₆} →
+ Poset p₁ p₂ p₃ → Poset p₄ p₅ p₆ → Poset _ _ _
po₁ ⊎-poset po₂ = record
{ _≤_ = _≤_ po₁ ⊎-Rel _≤_ po₂
; isPartialOrder = isPartialOrder po₁ ⊎-isPartialOrder
isPartialOrder po₂
} where open Poset
-_⊎-<-poset_ : Poset _ _ _ → Poset _ _ _ → Poset _ _ _
+_⊎-<-poset_ : ∀ {p₁ p₂ p₃ p₄ p₅ p₆} →
+ Poset p₁ p₂ p₃ → Poset p₄ p₅ p₆ → Poset _ _ _
po₁ ⊎-<-poset po₂ = record
{ _≤_ = _≤_ po₁ ⊎-< _≤_ po₂
; isPartialOrder = isPartialOrder po₁ ⊎-isPartialOrder
@@ -327,7 +348,8 @@ po₁ ⊎-<-poset po₂ = record
} where open Poset
_⊎-<-strictPartialOrder_ :
- StrictPartialOrder _ _ _ → StrictPartialOrder _ _ _ →
+ ∀ {p₁ p₂ p₃ p₄ p₅ p₆} →
+ StrictPartialOrder p₁ p₂ p₃ → StrictPartialOrder p₄ p₅ p₆ →
StrictPartialOrder _ _ _
spo₁ ⊎-<-strictPartialOrder spo₂ = record
{ _<_ = _<_ spo₁ ⊎-< _<_ spo₂
@@ -337,14 +359,101 @@ spo₁ ⊎-<-strictPartialOrder spo₂ = record
} where open StrictPartialOrder
_⊎-<-totalOrder_ :
- TotalOrder _ _ _ → TotalOrder _ _ _ → TotalOrder _ _ _
+ ∀ {t₁ t₂ t₃ t₄ t₅ t₆} →
+ TotalOrder t₁ t₂ t₃ → TotalOrder t₄ t₅ t₆ → TotalOrder _ _ _
to₁ ⊎-<-totalOrder to₂ = record
{ isTotalOrder = isTotalOrder to₁ ⊎-<-isTotalOrder isTotalOrder to₂
} where open TotalOrder
_⊎-<-decTotalOrder_ :
- DecTotalOrder _ _ _ → DecTotalOrder _ _ _ → DecTotalOrder _ _ _
+ ∀ {t₁ t₂ t₃ t₄ t₅ t₆} →
+ DecTotalOrder t₁ t₂ t₃ → DecTotalOrder t₄ t₅ t₆ → DecTotalOrder _ _ _
to₁ ⊎-<-decTotalOrder to₂ = record
{ isDecTotalOrder = isDecTotalOrder to₁ ⊎-<-isDecTotalOrder
isDecTotalOrder to₂
} where open DecTotalOrder
+
+------------------------------------------------------------------------
+-- Some properties related to equivalences and inverses
+
+⊎-Rel⇿≡ : ∀ {a b} (A : Set a) (B : Set b) →
+ Inverse (P.setoid A ⊎-setoid P.setoid B) (P.setoid (A ⊎ B))
+⊎-Rel⇿≡ _ _ = record
+ { to = record { _⟨$⟩_ = id; cong = to-cong }
+ ; from = record { _⟨$⟩_ = id; cong = from-cong }
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.refl ⊎-refl P.refl
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+ where
+ to-cong : (P._≡_ ⊎-Rel P._≡_) ⇒ P._≡_
+ to-cong (₁∼₂ ())
+ to-cong (₁∼₁ P.refl) = P.refl
+ to-cong (₂∼₂ P.refl) = P.refl
+
+ from-cong : P._≡_ ⇒ (P._≡_ ⊎-Rel P._≡_)
+ from-cong P.refl = P.refl ⊎-refl P.refl
+
+_⊎-equivalent_ :
+ ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈}
+ {A : Setoid s₁ s₂} {B : Setoid s₃ s₄}
+ {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} →
+ Equivalent A B → Equivalent C D →
+ Equivalent (A ⊎-setoid C) (B ⊎-setoid D)
+_⊎-equivalent_ {A = A} {B} {C} {D} A⇔B C⇔D = record
+ { to = record { _⟨$⟩_ = to; cong = to-cong }
+ ; from = record { _⟨$⟩_ = from; cong = from-cong }
+ }
+ where
+ open Setoid (A ⊎-setoid C) using () renaming (_≈_ to _≈AC_)
+ open Setoid (B ⊎-setoid D) using () renaming (_≈_ to _≈BD_)
+
+ to = Sum.map (_⟨$⟩_ (Equivalent.to A⇔B))
+ (_⟨$⟩_ (Equivalent.to C⇔D))
+
+ to-cong : _≈AC_ =[ to ]⇒ _≈BD_
+ to-cong (₁∼₂ ())
+ to-cong (₁∼₁ x∼₁y) = ₁∼₁ $ F.cong (Equivalent.to A⇔B) x∼₁y
+ to-cong (₂∼₂ x∼₂y) = ₂∼₂ $ F.cong (Equivalent.to C⇔D) x∼₂y
+
+ from = Sum.map (_⟨$⟩_ (Equivalent.from A⇔B))
+ (_⟨$⟩_ (Equivalent.from C⇔D))
+
+ from-cong : _≈BD_ =[ from ]⇒ _≈AC_
+ from-cong (₁∼₂ ())
+ from-cong (₁∼₁ x∼₁y) = ₁∼₁ $ F.cong (Equivalent.from A⇔B) x∼₁y
+ from-cong (₂∼₂ x∼₂y) = ₂∼₂ $ F.cong (Equivalent.from C⇔D) x∼₂y
+
+_⊎-⇔_ : ∀ {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 = A} {B} {C} {D} A⇔B C⇔D =
+ Inverse.equivalent (⊎-Rel⇿≡ B D) ⟨∘⟩
+ A⇔B ⊎-equivalent C⇔D ⟨∘⟩
+ Eq.sym (Inverse.equivalent (⊎-Rel⇿≡ A C))
+
+_⊎-inverse_ :
+ ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈}
+ {A : Setoid s₁ s₂} {B : Setoid s₃ s₄}
+ {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} →
+ Inverse A B → Inverse C D → Inverse (A ⊎-setoid C) (B ⊎-setoid D)
+A⇿B ⊎-inverse C⇿D = record
+ { to = Equivalent.to eq
+ ; from = Equivalent.from eq
+ ; inverse-of = record
+ { left-inverse-of = [ ₁∼₁ ∘ left-inverse-of A⇿B
+ , ₂∼₂ ∘ left-inverse-of C⇿D
+ ]
+ ; right-inverse-of = [ ₁∼₁ ∘ right-inverse-of A⇿B
+ , ₂∼₂ ∘ right-inverse-of C⇿D
+ ]
+ }
+ }
+ where
+ open Inverse
+ eq = equivalent A⇿B ⊎-equivalent equivalent C⇿D
+
+_⊎-⇿_ : ∀ {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 = A} {B} {C} {D} A⇿B C⇿D =
+ ⊎-Rel⇿≡ B D ⟪∘⟫ A⇿B ⊎-inverse C⇿D ⟪∘⟫ Inv.sym (⊎-Rel⇿≡ A C)
diff --git a/src/Relation/Nullary.agda b/src/Relation/Nullary.agda
index c201382..c3efa69 100644
--- a/src/Relation/Nullary.agda
+++ b/src/Relation/Nullary.agda
@@ -2,14 +2,10 @@
-- Operations on nullary relations (like negation and decidability)
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- Some operations on/properties of nullary relations, i.e. sets.
module Relation.Nullary where
-open import Data.Product
-open import Level
import Relation.Nullary.Core as Core
------------------------------------------------------------------------
@@ -18,14 +14,6 @@ import Relation.Nullary.Core as Core
open Core public using (¬_)
------------------------------------------------------------------------
--- Equivalence
-
-infix 3 _⇔_
-
-_⇔_ : ∀ {ℓ₁ ℓ₂} → Set ℓ₁ → Set ℓ₂ → Set (ℓ₁ ⊔ ℓ₂)
-P ⇔ Q = (P → Q) × (Q → P)
-
-------------------------------------------------------------------------
-- Decidable relations
open Core public using (Dec; yes; no)
diff --git a/src/Relation/Nullary/Decidable.agda b/src/Relation/Nullary/Decidable.agda
index 3d466b0..d8493cc 100644
--- a/src/Relation/Nullary/Decidable.agda
+++ b/src/Relation/Nullary/Decidable.agda
@@ -7,7 +7,10 @@
module Relation.Nullary.Decidable where
open import Data.Empty
-open import Data.Function
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence
+ using (_⇔_; equivalent; module Equivalent)
open import Data.Bool
open import Data.Product hiding (map)
open import Relation.Nullary
@@ -36,15 +39,19 @@ fromWitness {Q = yes p} = const _
fromWitness {Q = no ¬p} = ¬p
map : ∀ {p q} {P : Set p} {Q : Set q} → P ⇔ Q → Dec P → Dec Q
-map eq (yes p) = yes (proj₁ eq p)
-map eq (no ¬p) = no (¬p ∘ proj₂ eq)
+map P⇔Q (yes p) = yes (Equivalent.to P⇔Q ⟨$⟩ p)
+map P⇔Q (no ¬p) = no (¬p ∘ _⟨$⟩_ (Equivalent.from P⇔Q))
+
+map′ : ∀ {p q} {P : Set p} {Q : Set q} →
+ (P → Q) → (Q → P) → Dec P → Dec Q
+map′ P→Q Q→P = map (equivalent P→Q Q→P)
fromYes : ∀ {p} {P : Set p} → P → Dec P → P
fromYes _ (yes p) = p
fromYes p (no ¬p) = ⊥-elim (¬p p)
fromYes-map-commute :
- ∀ {p q} {P : Set p} {Q : Set q} {x y} (eq : P ⇔ Q) (d : Dec P) →
- fromYes y (map eq d) ≡ proj₁ eq (fromYes x d)
+ ∀ {p q} {P : Set p} {Q : Set q} {x y} (P⇔Q : P ⇔ Q) (d : Dec P) →
+ fromYes y (map P⇔Q d) ≡ Equivalent.to P⇔Q ⟨$⟩ fromYes x d
fromYes-map-commute _ (yes p) = refl
fromYes-map-commute {x = p} _ (no ¬p) = ⊥-elim (¬p p)
diff --git a/src/Relation/Nullary/Negation.agda b/src/Relation/Nullary/Negation.agda
index c97d007..173ec65 100644
--- a/src/Relation/Nullary/Negation.agda
+++ b/src/Relation/Nullary/Negation.agda
@@ -10,7 +10,7 @@ open import Relation.Nullary
open import Relation.Unary
open import Data.Bool
open import Data.Empty
-open import Data.Function
+open import Function
open import Data.Product as Prod
open import Data.Fin
open import Data.Fin.Dec
@@ -132,7 +132,7 @@ call/cc hyp ¬p = hyp (λ p → ⊥-elim (¬p p)) ¬p
independence-of-premise
: ∀ {p q r} {P : Set p} {Q : Set q} {R : Q → Set r} →
- Q → (P → Σ Q R) → ¬ ¬ Σ Q (λ x → P → R x)
+ Q → (P → Σ Q R) → ¬ ¬ (Σ[ x ∶ Q ] (P → R x))
independence-of-premise {P = P} q f = ¬¬-map helper excluded-middle
where
helper : Dec P → _
@@ -165,3 +165,23 @@ private
helper : ∃ (λ b → P → if b then Q else R) → (P → Q) ⊎ (P → R)
helper (true , f) = inj₁ f
helper (false , f) = inj₂ f
+
+-- The classical statements of excluded middle and double-negation
+-- elimination.
+
+Excluded-Middle : (ℓ : Level) → Set (suc ℓ)
+Excluded-Middle p = {P : Set p} → Dec P
+
+Double-Negation-Elimination : (ℓ : Level) → Set (suc ℓ)
+Double-Negation-Elimination p = {P : Set p} → Stable P
+
+private
+
+ -- The two statements above are equivalent. The proofs are so
+ -- simple, given the definitions above, that they are not exported.
+
+ em⇒dne : ∀ {ℓ} → Excluded-Middle ℓ → Double-Negation-Elimination ℓ
+ em⇒dne em = decidable-stable em
+
+ dne⇒em : ∀ {ℓ} → Double-Negation-Elimination ℓ → Excluded-Middle ℓ
+ dne⇒em dne = dne excluded-middle
diff --git a/src/Relation/Nullary/Product.agda b/src/Relation/Nullary/Product.agda
index 7751802..2cedd3e 100644
--- a/src/Relation/Nullary/Product.agda
+++ b/src/Relation/Nullary/Product.agda
@@ -2,15 +2,18 @@
-- Products of nullary relations
------------------------------------------------------------------------
+{-# OPTIONS --universe-polymorphism #-}
+
module Relation.Nullary.Product where
open import Data.Product
-open import Data.Function
+open import Function
open import Relation.Nullary
-- Some properties which are preserved by _×_.
-_×-dec_ : {P Q : Set} → Dec P → Dec Q → Dec (P × Q)
+_×-dec_ : ∀ {p q} {P : Set p} {Q : Set q} →
+ Dec P → Dec Q → Dec (P × Q)
yes p ×-dec yes q = yes (p , q)
no ¬p ×-dec _ = no (¬p ∘ proj₁)
_ ×-dec no ¬q = no (¬q ∘ proj₂)
diff --git a/src/Relation/Nullary/Universe.agda b/src/Relation/Nullary/Universe.agda
index 5f9953f..fb85597 100644
--- a/src/Relation/Nullary/Universe.agda
+++ b/src/Relation/Nullary/Universe.agda
@@ -14,8 +14,8 @@ open import Relation.Binary.Sum
open import Relation.Binary.Product.Pointwise
open import Data.Sum as Sum hiding (map)
open import Data.Product as Prod hiding (map)
-open import Data.Function
-import Data.Function.Equality as FunS
+open import Function
+import Function.Equality as FunS
open import Data.Empty
open import Category.Applicative
open import Category.Monad
@@ -37,36 +37,24 @@ data PropF : Set₁ where
_⇒_ : (P₁ : Set) (F₂ : PropF) → PropF
¬¬_ : (F : PropF) → PropF
-⟦_⟧ : PropF → (Set → Set)
-⟦ Id ⟧ = λ x → x
-⟦ K P ⟧ = λ _ → P
-⟦ F₁ ∨ F₂ ⟧ = λ x → ⟦ F₁ ⟧ x ⊎ ⟦ F₂ ⟧ x
-⟦ F₁ ∧ F₂ ⟧ = λ x → ⟦ F₁ ⟧ x × ⟦ F₂ ⟧ x
-⟦ P₁ ⇒ F₂ ⟧ = λ x → P₁ → ⟦ F₂ ⟧ x
-⟦ ¬¬ F ⟧ = λ x → ¬ ¬ ⟦ F ⟧ x
-
-- Equalities for universe inhabitants.
-Eq : (F : PropF) {P : Set} → Rel (⟦ F ⟧ P) zero
-Eq Id = _≡_
-Eq (K P) = _≡_
-Eq (F₁ ∨ F₂) = Eq F₁ ⊎-Rel Eq F₂
-Eq (F₁ ∧ F₂) = Eq F₁ ×-Rel Eq F₂
-Eq (P₁ ⇒ F₂) = FunS.≡↝ (λ _ → Eq F₂)
-Eq (¬¬ F) = Always
+mutual
+
+ setoid : PropF → {P : Set} → Setoid zero zero
+ setoid Id {P} = PropEq.setoid P
+ setoid (K P) = PropEq.setoid P
+ setoid (F₁ ∨ F₂) = setoid F₁ ⊎-setoid setoid F₂
+ setoid (F₁ ∧ F₂) = setoid F₁ ×-setoid setoid F₂
+ setoid (P₁ ⇒ F₂) = FunS.≡-setoid P₁
+ (Setoid.indexedSetoid (setoid F₂))
+ setoid (¬¬ F) {P} = Always-setoid (¬ ¬ ⟦ F ⟧ P)
+
+ ⟦_⟧ : PropF → (Set → Set)
+ ⟦ F ⟧ P = Setoid.Carrier (setoid F {P})
⟨_⟩_≈_ : (F : PropF) {P : Set} → Rel (⟦ F ⟧ P) zero
-⟨_⟩_≈_ = Eq
-
-isEquivalence : ∀ F {P} → IsEquivalence (Eq F {P})
-isEquivalence Id = PropEq.isEquivalence
-isEquivalence (K P) = PropEq.isEquivalence
-isEquivalence (F₁ ∨ F₂) = isEquivalence F₁ ⊎-isEquivalence
- isEquivalence F₂
-isEquivalence (F₁ ∧ F₂) = isEquivalence F₁ ×-isEquivalence
- isEquivalence F₂
-isEquivalence (P₁ ⇒ F₂) = FunS.≡↝-isEquivalence (λ _ → isEquivalence F₂)
-isEquivalence (¬¬ F) = Always-isEquivalence
+⟨_⟩_≈_ F = Setoid._≈_ (setoid F)
-- ⟦ F ⟧ is functorial.
diff --git a/src/Relation/Unary.agda b/src/Relation/Unary.agda
index efd0ea2..562415a 100644
--- a/src/Relation/Unary.agda
+++ b/src/Relation/Unary.agda
@@ -7,7 +7,7 @@
module Relation.Unary where
open import Data.Empty
-open import Data.Function
+open import Function
open import Data.Unit
open import Data.Product
open import Data.Sum
@@ -123,12 +123,11 @@ infixr 0 _⟨→⟩_
_⟨×⟩_ : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b} →
Pred A ℓ₁ → Pred B ℓ₂ → Pred (A × B) _
-(P ⟨×⟩ Q) p = P (proj₁ p) × Q (proj₂ p)
+P ⟨×⟩ Q = uncurry (λ p q → P p × Q q)
_⟨⊎⟩_ : ∀ {a b ℓ} {A : Set a} {B : Set b} →
Pred A ℓ → Pred B ℓ → Pred (A ⊎ B) _
-(P ⟨⊎⟩ Q) (inj₁ p) = P p
-(P ⟨⊎⟩ Q) (inj₂ q) = Q q
+P ⟨⊎⟩ Q = [ P , Q ]
_⟨→⟩_ : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b} →
Pred A ℓ₁ → Pred B ℓ₂ → Pred (A → B) _