summaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorIain Lane <laney@debian.org>2011-11-29 16:51:32 +0000
committerIain Lane <laney@debian.org>2011-11-29 16:51:32 +0000
commitbecce68e992af760dbdc6f17e77bc4dfeecccef7 (patch)
treefe929e2d8863fe17c1c2909283e1442b6f4a3286
parent9251e0b00e309e805a4576ae558b3d1e321d8f2c (diff)
Imported Upstream version 0.6~darcs20111129t1640
-rw-r--r--Everything.agda27
-rw-r--r--GNUmakefile6
-rw-r--r--GenerateEverything.hs23
-rw-r--r--Header2
-rw-r--r--LICENCE2
-rw-r--r--README.agda33
-rw-r--r--README.txt7
-rw-r--r--README/AVL.agda128
-rw-r--r--README/Case.agda38
-rw-r--r--README/Nat.agda2
-rw-r--r--README/Record.agda41
-rw-r--r--ffi/Data/FFI.hs11
-rw-r--r--ffi/IO/FFI.hs3
-rw-r--r--ffi/Setup.hs3
-rw-r--r--ffi/agda-lib-ffi.cabal12
-rw-r--r--lib.cabal4
-rw-r--r--src/Algebra.agda4
-rw-r--r--src/Algebra/FunctionProperties.agda4
-rw-r--r--src/Algebra/FunctionProperties/Core.agda4
-rw-r--r--src/Algebra/Morphism.agda4
-rw-r--r--src/Algebra/Operations.agda4
-rw-r--r--src/Algebra/Props/AbelianGroup.agda4
-rw-r--r--src/Algebra/Props/BooleanAlgebra.agda8
-rw-r--r--src/Algebra/Props/BooleanAlgebra/Expression.agda4
-rw-r--r--src/Algebra/Props/DistributiveLattice.agda8
-rw-r--r--src/Algebra/Props/Group.agda4
-rw-r--r--src/Algebra/Props/Lattice.agda8
-rw-r--r--src/Algebra/Props/Ring.agda4
-rw-r--r--src/Algebra/RingSolver.agda6
-rw-r--r--src/Algebra/RingSolver/AlmostCommutativeRing.agda4
-rw-r--r--src/Algebra/RingSolver/Lemmas.agda4
-rw-r--r--src/Algebra/RingSolver/Simple.agda4
-rw-r--r--src/Algebra/Structures.agda6
-rw-r--r--src/Category/Applicative.agda4
-rw-r--r--src/Category/Applicative/Indexed.agda4
-rw-r--r--src/Category/Functor.agda4
-rw-r--r--src/Category/Monad.agda4
-rw-r--r--src/Category/Monad/Continuation.agda4
-rw-r--r--src/Category/Monad/Identity.agda4
-rw-r--r--src/Category/Monad/Indexed.agda4
-rw-r--r--src/Category/Monad/Partiality.agda30
-rw-r--r--src/Category/Monad/State.agda4
-rw-r--r--src/Coinduction.agda4
-rw-r--r--src/Data/AVL.agda379
-rw-r--r--src/Data/AVL/IndexedMap.agda43
-rw-r--r--src/Data/AVL/Sets.agda28
-rw-r--r--src/Data/Bin.agda2
-rw-r--r--src/Data/Bool.agda4
-rw-r--r--src/Data/Bool/Properties.agda70
-rw-r--r--src/Data/Bool/Show.agda2
-rw-r--r--src/Data/BoundedVec.agda7
-rw-r--r--src/Data/BoundedVec/Inefficient.agda4
-rw-r--r--src/Data/Char.agda2
-rw-r--r--src/Data/Cofin.agda2
-rw-r--r--src/Data/Colist.agda11
-rw-r--r--src/Data/Conat.agda2
-rw-r--r--src/Data/Container.agda83
-rw-r--r--src/Data/Container/Any.agda117
-rw-r--r--src/Data/Container/Combinator.agda36
-rw-r--r--src/Data/Context.agda188
-rw-r--r--src/Data/Covec.agda2
-rw-r--r--src/Data/DifferenceList.agda32
-rw-r--r--src/Data/DifferenceNat.agda2
-rw-r--r--src/Data/DifferenceVec.agda27
-rw-r--r--src/Data/Digit.agda2
-rw-r--r--src/Data/Empty.agda7
-rw-r--r--src/Data/Fin.agda8
-rw-r--r--src/Data/Fin/Dec.agda10
-rw-r--r--src/Data/Fin/Props.agda6
-rw-r--r--src/Data/Fin/Subset.agda2
-rw-r--r--src/Data/Fin/Subset/Props.agda14
-rw-r--r--src/Data/Fin/Substitution.agda2
-rw-r--r--src/Data/Fin/Substitution/Example.agda2
-rw-r--r--src/Data/Fin/Substitution/Lemmas.agda2
-rw-r--r--src/Data/Fin/Substitution/List.agda2
-rw-r--r--src/Data/Graph/Acyclic.agda28
-rw-r--r--src/Data/Integer.agda2
-rw-r--r--src/Data/Integer/Divisibility.agda2
-rw-r--r--src/Data/Integer/Properties.agda2
-rw-r--r--src/Data/List.agda6
-rw-r--r--src/Data/List/All.agda4
-rw-r--r--src/Data/List/All/Properties.agda10
-rw-r--r--src/Data/List/Any.agda54
-rw-r--r--src/Data/List/Any/BagAndSetEquality.agda358
-rw-r--r--src/Data/List/Any/Membership.agda122
-rw-r--r--src/Data/List/Any/Properties.agda186
-rw-r--r--src/Data/List/Countdown.agda8
-rw-r--r--src/Data/List/NonEmpty.agda4
-rw-r--r--src/Data/List/NonEmpty/Properties.agda4
-rw-r--r--src/Data/List/Properties.agda4
-rw-r--r--src/Data/List/Reverse.agda2
-rw-r--r--src/Data/Maybe.agda15
-rw-r--r--src/Data/Maybe/Core.agda6
-rw-r--r--src/Data/Nat.agda22
-rw-r--r--src/Data/Nat/Coprimality.agda4
-rw-r--r--src/Data/Nat/DivMod.agda9
-rw-r--r--src/Data/Nat/Divisibility.agda12
-rw-r--r--src/Data/Nat/GCD.agda2
-rw-r--r--src/Data/Nat/GCD/Lemmas.agda2
-rw-r--r--src/Data/Nat/InfinitelyOften.agda5
-rw-r--r--src/Data/Nat/LCM.agda4
-rw-r--r--src/Data/Nat/Primality.agda37
-rw-r--r--src/Data/Nat/Properties.agda4
-rw-r--r--src/Data/Nat/Show.agda12
-rw-r--r--src/Data/Plus.agda6
-rw-r--r--src/Data/Product.agda4
-rw-r--r--src/Data/Product/N-ary.agda10
-rw-r--r--src/Data/Rational.agda6
-rw-r--r--src/Data/ReflexiveClosure.agda4
-rw-r--r--src/Data/Sign.agda2
-rw-r--r--src/Data/Sign/Properties.agda2
-rw-r--r--src/Data/Star.agda4
-rw-r--r--src/Data/Star/BoundedVec.agda2
-rw-r--r--src/Data/Star/Decoration.agda4
-rw-r--r--src/Data/Star/Environment.agda2
-rw-r--r--src/Data/Star/Fin.agda2
-rw-r--r--src/Data/Star/List.agda2
-rw-r--r--src/Data/Star/Nat.agda10
-rw-r--r--src/Data/Star/Pointer.agda2
-rw-r--r--src/Data/Star/Properties.agda4
-rw-r--r--src/Data/Star/Vec.agda2
-rw-r--r--src/Data/Stream.agda29
-rw-r--r--src/Data/String.agda2
-rw-r--r--src/Data/Sum.agda7
-rw-r--r--src/Data/Unit.agda10
-rw-r--r--src/Data/Unit/Core.agda42
-rw-r--r--src/Data/Vec.agda7
-rw-r--r--src/Data/Vec/Equality.agda34
-rw-r--r--src/Data/Vec/N-ary.agda12
-rw-r--r--src/Data/Vec/Properties.agda13
-rw-r--r--src/Data/W.agda4
-rw-r--r--src/Foreign/Haskell.agda34
-rw-r--r--src/Function.agda31
-rw-r--r--src/Function/Bijection.agda6
-rw-r--r--src/Function/Equality.agda4
-rw-r--r--src/Function/Equivalence.agda46
-rw-r--r--src/Function/Injection.agda12
-rw-r--r--src/Function/Inverse.agda144
-rw-r--r--src/Function/LeftInverse.agda19
-rw-r--r--src/Function/Related.agda382
-rw-r--r--src/Function/Related/TypeIsomorphisms.agda (renamed from src/Function/Inverse/TypeIsomorphisms.agda)226
-rw-r--r--src/Function/Surjection.agda23
-rw-r--r--src/IO.agda63
-rw-r--r--src/IO/Primitive.agda21
-rw-r--r--src/Induction.agda4
-rw-r--r--src/Induction/Lexicographic.agda4
-rw-r--r--src/Induction/Nat.agda2
-rw-r--r--src/Induction/WellFounded.agda4
-rw-r--r--src/Level.agda34
-rw-r--r--src/Record.agda230
-rw-r--r--src/Reflection.agda344
-rw-r--r--src/Relation/Binary.agda4
-rw-r--r--src/Relation/Binary/Consequences.agda4
-rw-r--r--src/Relation/Binary/Consequences/Core.agda4
-rw-r--r--src/Relation/Binary/Core.agda4
-rw-r--r--src/Relation/Binary/EqReasoning.agda4
-rw-r--r--src/Relation/Binary/Flip.agda4
-rw-r--r--src/Relation/Binary/HeterogeneousEquality.agda60
-rw-r--r--src/Relation/Binary/HeterogeneousEquality/Core.agda26
-rw-r--r--src/Relation/Binary/Indexed.agda4
-rw-r--r--src/Relation/Binary/Indexed/Core.agda4
-rw-r--r--src/Relation/Binary/InducedPreorders.agda4
-rw-r--r--src/Relation/Binary/List/NonStrictLex.agda2
-rw-r--r--src/Relation/Binary/List/Pointwise.agda4
-rw-r--r--src/Relation/Binary/List/StrictLex.agda2
-rw-r--r--src/Relation/Binary/NonStrictToStrict.agda4
-rw-r--r--src/Relation/Binary/On.agda4
-rw-r--r--src/Relation/Binary/OrderMorphism.agda4
-rw-r--r--src/Relation/Binary/PartialOrderReasoning.agda4
-rw-r--r--src/Relation/Binary/PreorderReasoning.agda4
-rw-r--r--src/Relation/Binary/Product/NonStrictLex.agda4
-rw-r--r--src/Relation/Binary/Product/Pointwise.agda215
-rw-r--r--src/Relation/Binary/Product/StrictLex.agda13
-rw-r--r--src/Relation/Binary/PropositionalEquality.agda84
-rw-r--r--src/Relation/Binary/PropositionalEquality/Core.agda4
-rw-r--r--src/Relation/Binary/PropositionalEquality/TrustMe.agda8
-rw-r--r--src/Relation/Binary/Props/DecTotalOrder.agda4
-rw-r--r--src/Relation/Binary/Props/Poset.agda4
-rw-r--r--src/Relation/Binary/Props/Preorder.agda4
-rw-r--r--src/Relation/Binary/Props/StrictPartialOrder.agda4
-rw-r--r--src/Relation/Binary/Props/StrictTotalOrder.agda4
-rw-r--r--src/Relation/Binary/Props/TotalOrder.agda4
-rw-r--r--src/Relation/Binary/Reflection.agda10
-rw-r--r--src/Relation/Binary/Sigma/Pointwise.agda398
-rw-r--r--src/Relation/Binary/Simple.agda14
-rw-r--r--src/Relation/Binary/StrictPartialOrderReasoning.agda4
-rw-r--r--src/Relation/Binary/StrictToNonStrict.agda4
-rw-r--r--src/Relation/Binary/Sum.agda219
-rw-r--r--src/Relation/Binary/Vec/Pointwise.agda32
-rw-r--r--src/Relation/Nullary.agda2
-rw-r--r--src/Relation/Nullary/Core.agda4
-rw-r--r--src/Relation/Nullary/Decidable.agda48
-rw-r--r--src/Relation/Nullary/Negation.agda12
-rw-r--r--src/Relation/Nullary/Product.agda6
-rw-r--r--src/Relation/Nullary/Sum.agda4
-rw-r--r--src/Relation/Nullary/Universe.agda67
-rw-r--r--src/Relation/Unary.agda4
-rw-r--r--src/Size.agda2
-rw-r--r--src/Universe.agda27
199 files changed, 3644 insertions, 2098 deletions
diff --git a/Everything.agda b/Everything.agda
index 7e92f6c..68c42b8 100644
--- a/Everything.agda
+++ b/Everything.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- All library modules, along with short descriptions
------------------------------------------------------------------------
@@ -136,9 +138,6 @@ import Data.Container.Any
-- Container combinators
import Data.Container.Combinator
--- Contexts, variables, substitutions, etc.
-import Data.Context
-
-- Coinductive vectors
import Data.Covec
@@ -263,6 +262,9 @@ import Data.Nat.InfinitelyOften
-- Least common multiple
import Data.Nat.LCM
+-- Primality
+import Data.Nat.Primality
+
-- A bunch of properties about natural number operations
import Data.Nat.Properties
@@ -329,7 +331,7 @@ import Data.String
-- Sums (disjoint unions)
import Data.Sum
--- The unit type
+-- Some unit types
import Data.Unit
-- Vectors
@@ -347,7 +349,7 @@ import Data.Vec.Properties
-- W-types
import Data.W
--- Types used (only) when calling out to Haskell via the FFI
+-- Type(s) used (only) when calling out to Haskell via the FFI
import Foreign.Haskell
-- Simple combinators working solely on and with functions
@@ -368,12 +370,17 @@ import Function.Injection
-- Inverses
import Function.Inverse
--- Various basic type isomorphisms
-import Function.Inverse.TypeIsomorphisms
-
-- Left inverses
import Function.LeftInverse
+-- A universe which includes several kinds of "relatedness" for sets,
+-- such as equivalences, surjections and bijections
+import Function.Related
+
+-- Basic lemmas showing that various types are related (isomorphic or
+-- equivalent or…)
+import Function.Related.TypeIsomorphisms
+
-- Surjections
import Function.Surjection
@@ -398,6 +405,10 @@ import Induction.WellFounded
-- Universe levels
import Level
+-- Record types with manifest fields and "with", based on Randy
+-- Pollack's "Dependently Typed Records in Type Theory"
+import Record
+
-- Support for reflection
import Reflection
diff --git a/GNUmakefile b/GNUmakefile
index a180864..342f661 100644
--- a/GNUmakefile
+++ b/GNUmakefile
@@ -3,7 +3,13 @@ AGDA=agda
test: Everything.agda
$(AGDA) -i. -isrc README.agda
+setup: Everything.agda agda-lib-ffi
+
.PHONY: Everything.agda
Everything.agda:
cabal install
GenerateEverything
+
+.PHONY: agda-lib-ffi
+agda-lib-ffi:
+ cd ffi && cabal install
diff --git a/GenerateEverything.hs b/GenerateEverything.hs
index 4515f86..dbd42ef 100644
--- a/GenerateEverything.hs
+++ b/GenerateEverything.hs
@@ -18,14 +18,14 @@ main = do
[] -> return ()
_ -> hPutStr stderr usage >> exitFailure
- header <- readFile headerFile
+ header <- readFileUTF8 headerFile
modules <- filter isLibraryModule . List.sort <$>
find always
(extension ==? ".agda" ||? extension ==? ".lagda")
srcDir
headers <- mapM extractHeader modules
- writeFile outputFile $
+ writeFileUTF8 outputFile $
header ++ format (zip modules headers)
-- | Usage info.
@@ -54,11 +54,11 @@ isLibraryModule f =
-- | Reads a module and extracts the header.
extractHeader :: FilePath -> IO [String]
-extractHeader mod = fmap (extract . lines) $ readFile mod
+extractHeader mod = fmap (extract . lines) $ readFileUTF8 mod
where
delimiter = all (== '-')
- extract (d1 : ss)
+ extract (d1 : "-- The Agda standard library" : "--" : ss)
| delimiter d1
, (info, d2 : rest) <- span ("-- " `List.isPrefixOf`) ss
, delimiter d2
@@ -84,3 +84,18 @@ fileToMod = map slashToDot . dropExtension . makeRelative srcDir
where
slashToDot c | isPathSeparator c = '.'
| otherwise = c
+
+-- | A variant of 'readFile' which uses the 'utf8' encoding.
+
+readFileUTF8 :: FilePath -> IO String
+readFileUTF8 f = do
+ h <- openFile f ReadMode
+ hSetEncoding h utf8
+ hGetContents h
+
+-- | A variant of 'writeFile' which uses the 'utf8' encoding.
+
+writeFileUTF8 :: FilePath -> String -> IO ()
+writeFileUTF8 f s = withFile f WriteMode $ \h -> do
+ hSetEncoding h utf8
+ hPutStr h s
diff --git a/Header b/Header
index e62d4b2..dd0cdfa 100644
--- a/Header
+++ b/Header
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- All library modules, along with short descriptions
------------------------------------------------------------------------
diff --git a/LICENCE b/LICENCE
index 507653d..158f19b 100644
--- a/LICENCE
+++ b/LICENCE
@@ -2,7 +2,7 @@ Copyright (c) 2007-2011 Nils Anders Danielsson, Ulf Norell, Shin-Cheng
Mu, Samuel Bronson, Dan Doel, Patrik Jansson, Liang-Ting Chen,
Jean-Philippe Bernardy, Andrés Sicard-Ramírez, Nicolas Pouillard,
Darin Morrison, Peter Berry, Daniel Brown, Simon Foster, Dominique
-Devriese
+Devriese, Andreas Abel, Alcatel-Lucent
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 e995e8e..578c7eb 100644
--- a/README.agda
+++ b/README.agda
@@ -3,11 +3,11 @@ module README where
------------------------------------------------------------------------
-- The Agda standard library
--
--- Author: Nils Anders Danielsson, with contributions from
--- Jean-Philippe Bernardy, Peter Berry, Samuel Bronson, Daniel Brown,
--- Liang-Ting Chen, Dominique Devriese, Dan Doel, Simon Foster, Patrik
--- Jansson, Darin Morrison, Shin-Cheng Mu, Ulf Norell, Nicolas
--- Pouillard and Andrés Sicard-Ramírez
+-- Author: Nils Anders Danielsson, with contributions from Andreas
+-- Abel, Jean-Philippe Bernardy, Peter Berry, Samuel Bronson, Daniel
+-- Brown, Liang-Ting Chen, Dominique Devriese, Dan Doel, Simon Foster,
+-- Patrik Jansson, Alan Jeffrey, Darin Morrison, Shin-Cheng Mu, Ulf
+-- Norell, Nicolas Pouillard and Andrés Sicard-Ramírez
------------------------------------------------------------------------
-- Note that the development version of the library often requires the
@@ -22,6 +22,15 @@ module README where
-- --include-path flag or by customising the Emacs mode variable
-- agda2-include-dirs (M-x customize-group RET agda2 RET).
+-- To compile the library using the MAlonzo compiler you first need to
+-- install some supporting Haskell code, for instance as follows:
+--
+-- cd ffi
+-- cabal install
+--
+-- Currently the library does not support the Epic or JavaScript
+-- compiler backends.
+
-- Contributions to this library are welcome (but to avoid wasted work
-- it is suggested that you discuss large changes before implementing
-- them). Please send contributions in the form of darcs patches (run
@@ -61,6 +70,8 @@ module README where
-- Input/output-related functions.
-- • Level
-- Universe levels.
+-- • Record
+-- An encoding of record types with manifest fields and "with".
-- • Reflection
-- Support for reflection.
-- • Relation
@@ -246,6 +257,18 @@ import IO
import README.Nat
+-- Some examples showing how the AVL tree module can be used.
+
+import README.AVL
+
+-- An example showing how the Record module can be used.
+
+import README.Record
+
+-- An example showing how the case expression can be used.
+
+import README.Case
+
------------------------------------------------------------------------
-- Core modules
------------------------------------------------------------------------
diff --git a/README.txt b/README.txt
index 6361460..e02b7a8 100644
--- a/README.txt
+++ b/README.txt
@@ -9,6 +9,7 @@ generated automatically; if you have downloaded the library from its
darcs repository and want to type check README you can construct
Everything by running "cabal install && GenerateEverything".
-Note that all library sources are located under src. The modules
-README, README.* and Everything are not really part of the library, so
-these modules are located in the top-level directory instead.
+Note that all library sources are located under src or ffi. The
+modules 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/AVL.agda b/README/AVL.agda
new file mode 100644
index 0000000..ed9509c
--- /dev/null
+++ b/README/AVL.agda
@@ -0,0 +1,128 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Some examples showing how the AVL tree module can be used
+------------------------------------------------------------------------
+
+module README.AVL where
+
+------------------------------------------------------------------------
+-- Setup
+
+-- AVL trees are defined in Data.AVL.
+
+import Data.AVL
+
+-- This module is parametrised by keys, which have to form a (strict)
+-- total order, and values, which are indexed by keys. Let us use
+-- natural numbers as keys and vectors of strings as values.
+
+import Data.Nat.Properties as ℕ
+open import Data.String using (String)
+open import Data.Vec using (Vec; _∷_; [])
+open import Relation.Binary using (module StrictTotalOrder)
+
+open Data.AVL (Vec String)
+ (StrictTotalOrder.isStrictTotalOrder ℕ.strictTotalOrder)
+
+------------------------------------------------------------------------
+-- Construction of trees
+
+-- Some values.
+
+v₁ = "cepa" ∷ []
+v₂ = "apa" ∷ "bepa" ∷ []
+
+-- Empty and singleton trees.
+
+t₀ : Tree
+t₀ = empty
+
+t₁ : Tree
+t₁ = singleton 2 v₂
+
+-- Insertion of a key-value pair into a tree.
+
+t₂ = insert 1 v₁ t₁
+
+-- Deletion of the mapping for a certain key.
+
+t₃ = delete 2 t₂
+
+-- Conversion of a list of key-value mappings to a tree.
+
+open import Data.List using (_∷_; [])
+open import Data.Product as Prod using (_,_; _,′_)
+
+t₄ = fromList ((2 , v₂) ∷ (1 , v₁) ∷ [])
+
+------------------------------------------------------------------------
+-- Queries
+
+-- Let us formulate queries as unit tests.
+
+open import Relation.Binary.PropositionalEquality using (_≡_; refl)
+
+-- Searching for a key.
+
+open import Data.Bool using (true; false)
+open import Data.Maybe as Maybe using (just; nothing)
+
+q₀ : lookup 2 t₂ ≡ just v₂
+q₀ = refl
+
+q₁ : lookup 2 t₃ ≡ nothing
+q₁ = refl
+
+q₂ : 3 ∈? t₂ ≡ false
+q₂ = refl
+
+q₃ : 1 ∈? t₄ ≡ true
+q₃ = refl
+
+-- Turning a tree into a sorted list of key-value pairs.
+
+q₄ : toList t₁ ≡ (2 , v₂) ∷ []
+q₄ = refl
+
+q₅ : toList t₂ ≡ (1 , v₁) ∷ (2 , v₂) ∷ []
+q₅ = refl
+
+------------------------------------------------------------------------
+-- Views
+
+-- Partitioning a tree into the smallest element plus the rest, or the
+-- largest element plus the rest.
+
+open import Category.Functor using (module RawFunctor)
+open import Function using (id)
+import Level
+
+open RawFunctor (Maybe.functor {f = Level.zero}) using (_<$>_)
+
+v₆ : headTail t₀ ≡ nothing
+v₆ = refl
+
+v₇ : Prod.map id toList <$> headTail t₂ ≡
+ just ((1 , v₁) , ((2 , v₂) ∷ []))
+v₇ = refl
+
+v₈ : initLast t₀ ≡ nothing
+v₈ = refl
+
+v₉ : Prod.map toList id <$> initLast t₄ ≡
+ just (((1 , v₁) ∷ []) ,′ (2 , v₂))
+v₉ = refl
+
+------------------------------------------------------------------------
+-- Further reading
+
+-- Variations of the AVL tree module are available:
+
+-- • Finite maps with indexed keys and values.
+
+import Data.AVL.IndexedMap
+
+-- • Finite sets.
+
+import Data.AVL.Sets
diff --git a/README/Case.agda b/README/Case.agda
new file mode 100644
index 0000000..e42f55a
--- /dev/null
+++ b/README/Case.agda
@@ -0,0 +1,38 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Examples showing how the case expressions can be used
+------------------------------------------------------------------------
+
+module README.Case where
+
+open import Data.Fin hiding (pred)
+open import Data.Maybe hiding (from-just)
+open import Data.Nat hiding (pred)
+open import Function
+
+-- Some simple examples.
+
+empty : ∀ {a} {A : Set a} → Fin 0 → A
+empty i = case i of λ()
+
+pred : ℕ → ℕ
+pred n = case n of λ
+ { zero → zero
+ ; (suc n) → n
+ }
+
+from-just : ∀ {a} {A : Set a} (x : Maybe A) → From-just A x
+from-just x = case x return From-just _ of λ
+ { (just x) → x
+ ; nothing → _
+ }
+
+-- Note that some natural uses of case are rejected by the termination
+-- checker:
+--
+-- plus : ℕ → ℕ → ℕ
+-- plus m n = case m of λ
+-- { zero → n
+-- ; (suc m) → suc (plus m n)
+-- }
diff --git a/README/Nat.agda b/README/Nat.agda
index 2717365..5e6496b 100644
--- a/README/Nat.agda
+++ b/README/Nat.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some examples showing where the natural numbers and some related
-- operations and properties are defined, and how they can be used
------------------------------------------------------------------------
diff --git a/README/Record.agda b/README/Record.agda
new file mode 100644
index 0000000..ff680f1
--- /dev/null
+++ b/README/Record.agda
@@ -0,0 +1,41 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- An example of how the Record module can be used
+------------------------------------------------------------------------
+
+-- Taken from Randy Pollack's paper "Dependently Typed Records in Type
+-- Theory".
+
+module README.Record where
+
+open import Data.Product
+open import Data.String
+open import Function using (flip)
+open import Level
+import Record
+open import Relation.Binary
+
+-- Let us use strings as labels.
+
+open Record String _≟_
+
+-- Partial equivalence relations.
+
+PER : Signature _
+PER = ∅ , "S" ∶ (λ _ → Set)
+ , "R" ∶ (λ r → r · "S" → r · "S" → Set)
+ , "sym" ∶ (λ r → Lift (Symmetric (r · "R")))
+ , "trans" ∶ (λ r → Lift (Transitive (r · "R")))
+
+-- Given a PER the converse relation is also a PER.
+
+converse : (P : Record PER) →
+ Record (PER With "S" ≔ (λ _ → P · "S")
+ With "R" ≔ (λ _ → flip (P · "R")))
+converse P =
+ Rec′⇒Rec
+ (PER With "S" ≔ (λ _ → P · "S")
+ With "R" ≔ (λ _ → flip (P · "R")))
+ ((((_ ,) ,) , lift λ {_} → lower (P · "sym")) ,
+ (lift λ {_} yRx zRy → lower (P · "trans") zRy yRx))
diff --git a/ffi/Data/FFI.hs b/ffi/Data/FFI.hs
new file mode 100644
index 0000000..6c8bf54
--- /dev/null
+++ b/ffi/Data/FFI.hs
@@ -0,0 +1,11 @@
+{-# LANGUAGE EmptyDataDecls #-}
+
+module Data.FFI where
+
+type AgdaList a b = [b]
+type AgdaMaybe a b = Maybe b
+type AgdaEither a b c d = Either c d
+
+data AgdaEmpty
+
+data AgdaStream a = Cons a (AgdaStream a)
diff --git a/ffi/IO/FFI.hs b/ffi/IO/FFI.hs
new file mode 100644
index 0000000..5c6fac3
--- /dev/null
+++ b/ffi/IO/FFI.hs
@@ -0,0 +1,3 @@
+module IO.FFI where
+
+type AgdaIO a b = IO b
diff --git a/ffi/Setup.hs b/ffi/Setup.hs
new file mode 100644
index 0000000..e8ef27d
--- /dev/null
+++ b/ffi/Setup.hs
@@ -0,0 +1,3 @@
+import Distribution.Simple
+
+main = defaultMain
diff --git a/ffi/agda-lib-ffi.cabal b/ffi/agda-lib-ffi.cabal
new file mode 100644
index 0000000..1af663d
--- /dev/null
+++ b/ffi/agda-lib-ffi.cabal
@@ -0,0 +1,12 @@
+name: agda-lib-ffi
+version: 0.0.1
+cabal-version: >= 1.8
+build-type: Simple
+description: Auxiliary Haskell code used by Agda's standard library.
+
+library
+ build-depends:
+ -- Phony dependency, to avoid a bug in cabal-install.
+ base >= 0 && <= 99999999999
+ exposed-modules: Data.FFI
+ IO.FFI
diff --git a/lib.cabal b/lib.cabal
index a70149c..a0c40cd 100644
--- a/lib.cabal
+++ b/lib.cabal
@@ -7,12 +7,12 @@ description: Helper programs.
executable GenerateEverything
hs-source-dirs: .
main-is: GenerateEverything.hs
- build-depends: base >= 4.2 && < 4.4,
+ build-depends: base >= 4.2 && < 4.5,
filemanip == 0.3.*,
filepath >= 1.1 && < 1.3
executable AllNonAsciiChars
hs-source-dirs: .
main-is: AllNonAsciiChars.hs
- build-depends: base >= 4.2 && < 4.4,
+ build-depends: base >= 4.2 && < 4.5,
filemanip == 0.3.*
diff --git a/src/Algebra.agda b/src/Algebra.agda
index bbc9e9b..0607fbc 100644
--- a/src/Algebra.agda
+++ b/src/Algebra.agda
@@ -1,10 +1,10 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Definitions of algebraic structures like monoids and rings
-- (packed in records together with sets, operations, etc.)
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Algebra where
open import Relation.Binary
diff --git a/src/Algebra/FunctionProperties.agda b/src/Algebra/FunctionProperties.agda
index df1d94c..1d49422 100644
--- a/src/Algebra/FunctionProperties.agda
+++ b/src/Algebra/FunctionProperties.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties of functions, such as associativity and commutativity
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- These properties can (for instance) be used to define algebraic
-- structures.
diff --git a/src/Algebra/FunctionProperties/Core.agda b/src/Algebra/FunctionProperties/Core.agda
index 1b604f0..7814470 100644
--- a/src/Algebra/FunctionProperties/Core.agda
+++ b/src/Algebra/FunctionProperties/Core.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- 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 some of
diff --git a/src/Algebra/Morphism.agda b/src/Algebra/Morphism.agda
index 3bf8c49..6da18bb 100644
--- a/src/Algebra/Morphism.agda
+++ b/src/Algebra/Morphism.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Morphisms between algebraic structures
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Algebra.Morphism where
open import Relation.Binary
diff --git a/src/Algebra/Operations.agda b/src/Algebra/Operations.agda
index d8b6e13..4aac58c 100644
--- a/src/Algebra/Operations.agda
+++ b/src/Algebra/Operations.agda
@@ -1,10 +1,10 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some defined operations (multiplication by natural number and
-- exponentiation)
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Algebra
module Algebra.Operations {s₁ s₂} (S : Semiring s₁ s₂) where
diff --git a/src/Algebra/Props/AbelianGroup.agda b/src/Algebra/Props/AbelianGroup.agda
index 774c061..37ecb5c 100644
--- a/src/Algebra/Props/AbelianGroup.agda
+++ b/src/Algebra/Props/AbelianGroup.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some derivable properties
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Algebra
module Algebra.Props.AbelianGroup
diff --git a/src/Algebra/Props/BooleanAlgebra.agda b/src/Algebra/Props/BooleanAlgebra.agda
index d0ec453..de20cb2 100644
--- a/src/Algebra/Props/BooleanAlgebra.agda
+++ b/src/Algebra/Props/BooleanAlgebra.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some derivable properties
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Algebra
module Algebra.Props.BooleanAlgebra
@@ -22,7 +22,7 @@ import Relation.Binary.EqReasoning as EqR; open EqR setoid
open import Relation.Binary
open import Function
open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence using (_⇔_; module Equivalent)
+open import Function.Equivalence using (_⇔_; module Equivalence)
open import Data.Product
------------------------------------------------------------------------
@@ -261,7 +261,7 @@ replace-equality {_≈′_} ≈⇔≈′ = record
; ∧-complementʳ = λ x → to ⟨$⟩ ∧-complementʳ x
; ¬-cong = λ i≈j → to ⟨$⟩ ¬-cong (from ⟨$⟩ i≈j)
}
- } where open module E {x y} = Equivalent (≈⇔≈′ {x} {y})
+ } where open module E {x y} = Equivalence (≈⇔≈′ {x} {y})
------------------------------------------------------------------------
-- (⊕, ∧, id, ⊥, ⊤) is a commutative ring
diff --git a/src/Algebra/Props/BooleanAlgebra/Expression.agda b/src/Algebra/Props/BooleanAlgebra/Expression.agda
index 4c0319d..5a3814f 100644
--- a/src/Algebra/Props/BooleanAlgebra/Expression.agda
+++ b/src/Algebra/Props/BooleanAlgebra/Expression.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Boolean algebra expressions
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Algebra
module Algebra.Props.BooleanAlgebra.Expression
diff --git a/src/Algebra/Props/DistributiveLattice.agda b/src/Algebra/Props/DistributiveLattice.agda
index b48fc76..37d8f98 100644
--- a/src/Algebra/Props/DistributiveLattice.agda
+++ b/src/Algebra/Props/DistributiveLattice.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some derivable properties
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Algebra
module Algebra.Props.DistributiveLattice
@@ -21,7 +21,7 @@ open import Relation.Binary
import Relation.Binary.EqReasoning as EqR; open EqR setoid
open import Function
open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence using (_⇔_; module Equivalent)
+open import Function.Equivalence using (_⇔_; module Equivalence)
open import Data.Product
∨-∧-distrib : _∨_ DistributesOver _∧_
@@ -82,4 +82,4 @@ replace-equality {_≈′_} ≈⇔≈′ = record
{ isLattice = Lattice.isLattice (L.replace-equality ≈⇔≈′)
; ∨-∧-distribʳ = λ x y z → to ⟨$⟩ ∨-∧-distribʳ x y z
}
- } where open module E {x y} = Equivalent (≈⇔≈′ {x} {y})
+ } where open module E {x y} = Equivalence (≈⇔≈′ {x} {y})
diff --git a/src/Algebra/Props/Group.agda b/src/Algebra/Props/Group.agda
index e7dfe84..31283db 100644
--- a/src/Algebra/Props/Group.agda
+++ b/src/Algebra/Props/Group.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some derivable properties
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Algebra
module Algebra.Props.Group {g₁ g₂} (G : Group g₁ g₂) where
diff --git a/src/Algebra/Props/Lattice.agda b/src/Algebra/Props/Lattice.agda
index c0a7954..4f9f142 100644
--- a/src/Algebra/Props/Lattice.agda
+++ b/src/Algebra/Props/Lattice.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some derivable properties
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Algebra
module Algebra.Props.Lattice {l₁ l₂} (L : Lattice l₁ l₂) where
@@ -15,7 +15,7 @@ open import Relation.Binary
import Relation.Binary.EqReasoning as EqR; open EqR setoid
open import Function
open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence using (_⇔_; module Equivalent)
+open import Function.Equivalence using (_⇔_; module Equivalence)
open import Data.Product
∧-idempotent : Idempotent _∧_
@@ -103,4 +103,4 @@ replace-equality {_≈′_} ≈⇔≈′ = record
; absorptive = (λ x y → to ⟨$⟩ proj₁ absorptive x y)
, (λ x y → to ⟨$⟩ proj₂ absorptive x y)
}
- } where open module E {x y} = Equivalent (≈⇔≈′ {x} {y})
+ } where open module E {x y} = Equivalence (≈⇔≈′ {x} {y})
diff --git a/src/Algebra/Props/Ring.agda b/src/Algebra/Props/Ring.agda
index a2c05d8..61f1278 100644
--- a/src/Algebra/Props/Ring.agda
+++ b/src/Algebra/Props/Ring.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some derivable properties
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Algebra
module Algebra.Props.Ring {r₁ r₂} (R : Ring r₁ r₂) where
diff --git a/src/Algebra/RingSolver.agda b/src/Algebra/RingSolver.agda
index c678408..31d4d52 100644
--- a/src/Algebra/RingSolver.agda
+++ b/src/Algebra/RingSolver.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- 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.
@@ -34,7 +34,7 @@ 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 Function hiding (type-signature)
-open import Level
+open import Level using (_⊔_)
infix 9 _↑ :-_ -‿NF_
infixr 9 _:^_ _^-NF_ _:↑_
diff --git a/src/Algebra/RingSolver/AlmostCommutativeRing.agda b/src/Algebra/RingSolver/AlmostCommutativeRing.agda
index 946ff9c..802da61 100644
--- a/src/Algebra/RingSolver/AlmostCommutativeRing.agda
+++ b/src/Algebra/RingSolver/AlmostCommutativeRing.agda
@@ -1,10 +1,10 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Commutative semirings with some additional structure ("almost"
-- commutative rings), used by the ring solver
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Algebra.RingSolver.AlmostCommutativeRing where
open import Relation.Binary
diff --git a/src/Algebra/RingSolver/Lemmas.agda b/src/Algebra/RingSolver/Lemmas.agda
index 2003794..9748803 100644
--- a/src/Algebra/RingSolver/Lemmas.agda
+++ b/src/Algebra/RingSolver/Lemmas.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- 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.
diff --git a/src/Algebra/RingSolver/Simple.agda b/src/Algebra/RingSolver/Simple.agda
index 077d910..d969e87 100644
--- a/src/Algebra/RingSolver/Simple.agda
+++ b/src/Algebra/RingSolver/Simple.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Instantiates the ring solver with two copies of the same ring
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Algebra.RingSolver.AlmostCommutativeRing
module Algebra.RingSolver.Simple
diff --git a/src/Algebra/Structures.agda b/src/Algebra/Structures.agda
index 314683a..a93254b 100644
--- a/src/Algebra/Structures.agda
+++ b/src/Algebra/Structures.agda
@@ -1,10 +1,10 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some algebraic structures (not packed up with sets, operations,
-- etc.)
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Algebra.Structures where
@@ -12,7 +12,7 @@ module Algebra.Structures where
import Algebra.FunctionProperties as FunctionProperties
open import Data.Product
open import Function
-open import Level hiding (zero)
+open import Level using (_⊔_)
import Relation.Binary.EqReasoning as EqR
open FunctionProperties using (Op₁; Op₂)
diff --git a/src/Category/Applicative.agda b/src/Category/Applicative.agda
index 6cafb2e..85e5178 100644
--- a/src/Category/Applicative.agda
+++ b/src/Category/Applicative.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Applicative functors
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- Note that currently the applicative functor laws are not included
-- here.
diff --git a/src/Category/Applicative/Indexed.agda b/src/Category/Applicative/Indexed.agda
index 20be955..2012cf0 100644
--- a/src/Category/Applicative/Indexed.agda
+++ b/src/Category/Applicative/Indexed.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Indexed applicative functors
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- Note that currently the applicative functor laws are not included
-- here.
diff --git a/src/Category/Functor.agda b/src/Category/Functor.agda
index 33a3814..3fc7676 100644
--- a/src/Category/Functor.agda
+++ b/src/Category/Functor.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Functors
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- Note that currently the functor laws are not included here.
module Category.Functor where
diff --git a/src/Category/Monad.agda b/src/Category/Monad.agda
index 3ffb879..8e2ddf3 100644
--- a/src/Category/Monad.agda
+++ b/src/Category/Monad.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Monads
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- Note that currently the monad laws are not included here.
module Category.Monad where
diff --git a/src/Category/Monad/Continuation.agda b/src/Category/Monad/Continuation.agda
index ee1ab33..396f75c 100644
--- a/src/Category/Monad/Continuation.agda
+++ b/src/Category/Monad/Continuation.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- A delimited continuation monad
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Category.Monad.Continuation where
open import Category.Applicative
diff --git a/src/Category/Monad/Identity.agda b/src/Category/Monad/Identity.agda
index 86cb6ef..f01c249 100644
--- a/src/Category/Monad/Identity.agda
+++ b/src/Category/Monad/Identity.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- The identity monad
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Category.Monad.Identity where
open import Category.Monad
diff --git a/src/Category/Monad/Indexed.agda b/src/Category/Monad/Indexed.agda
index c5163d8..d286f92 100644
--- a/src/Category/Monad/Indexed.agda
+++ b/src/Category/Monad/Indexed.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Indexed monads
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- Note that currently the monad laws are not included here.
module Category.Monad.Indexed where
diff --git a/src/Category/Monad/Partiality.agda b/src/Category/Monad/Partiality.agda
index bbda0e4..ab79741 100644
--- a/src/Category/Monad/Partiality.agda
+++ b/src/Category/Monad/Partiality.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- The partiality monad
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Category.Monad.Partiality where
open import Coinduction
@@ -13,8 +13,8 @@ open import Data.Nat using (ℕ; zero; suc; _+_)
open import Data.Product as Prod hiding (map)
open import Data.Sum hiding (map)
open import Function
-open import Function.Equivalence using (_⇔_; equivalent)
-open import Level
+open import Function.Equivalence using (_⇔_; equivalence)
+open import Level using (_⊔_)
open import Relation.Binary as B hiding (Rel)
open import Relation.Binary.PropositionalEquality as P using (_≡_)
open import Relation.Nullary
@@ -146,7 +146,7 @@ module Equality {a ℓ} {A : Set a} -- The "return type".
private
module Dummy {a ℓ} {A : Set a} {_∼_ : A → A → Set ℓ} where
- open Equality _∼_
+ open Equality _∼_ using (Rel; _≅_; _≳_; _≲_; _≈_; _⇓[_]_; _⇑[_])
open Equality.Rel
-- All relations include strong equality.
@@ -233,8 +233,8 @@ private
now-trans : ∀ {k x y} {v : A} →
Rel k x y → Rel k y (now v) → Rel k x (now v)
now-trans (now x∼y) (now y∼z) = now (trans-∼ x∼y y∼z)
- now-trans x∼ly (laterˡ y∼z) = now-trans (laterʳ⁻¹ x∼ly) y∼z
now-trans (laterˡ x∼y) y∼z = laterˡ (now-trans x∼y y∼z)
+ now-trans x∼ly (laterˡ y∼z) = now-trans (laterʳ⁻¹ x∼ly) y∼z
mutual
@@ -376,7 +376,7 @@ private
-- Map.
- map : ∀ {_∼′_ k} →
+ map : ∀ {_∼′_ : A → A → Set a} {k} →
_∼′_ ⇒ _∼_ → Equality.Rel _∼′_ k ⇒ Equality.Rel _∼_ k
map ∼′⇒∼ (now x∼y) = now (∼′⇒∼ x∼y)
map ∼′⇒∼ (later x∼y) = later (♯ map ∼′⇒∼ (♭ x∼y))
@@ -424,7 +424,7 @@ private
-- Never is a left and right "zero" of bind.
- left-zero : ∀ {B} (f : B → A ⊥) → let open M in
+ left-zero : {B : Set a} (f : B → A ⊥) → let open M in
(never >>= f) ≅ never
left-zero f = later (♯ left-zero f)
@@ -557,14 +557,14 @@ module Workaround {a} where
infixl 1 _>>=_
- data _⊥P : Set a → Set (suc a) where
+ data _⊥P : Set a → Set (Level.suc a) where
now : ∀ {A} (x : A) → A ⊥P
later : ∀ {A} (x : ∞ (A ⊥P)) → A ⊥P
_>>=_ : ∀ {A B} (x : A ⊥P) (f : A → B ⊥P) → B ⊥P
private
- data _⊥W : Set a → Set (suc a) where
+ data _⊥W : Set a → Set (Level.suc a) where
now : ∀ {A} (x : A) → A ⊥W
later : ∀ {A} (x : A ⊥P) → A ⊥W
@@ -596,8 +596,8 @@ module Workaround {a} where
module Correct where
private
- open module Eq {A} = Equality {A = A} _≡_
- open module R {A} = Reasoning {A = A} P.isEquivalence
+ open module Eq {A : Set a} = Equality {A = A} _≡_
+ open module R {A : Set a} = Reasoning (P.isEquivalence {A = A})
now-hom : ∀ {A} (x : A) → ⟦ now x ⟧P ≅ now x
now-hom x = now x ∎
@@ -653,7 +653,7 @@ module AlternativeEquality {a ℓ} where
_∣_≈P_ : ∀ S → B.Rel (El S ⊥) _
_∣_≈P_ = flip RelP (other weak)
- data RelP S : Kind → B.Rel (El S ⊥) (suc (a ⊔ ℓ)) where
+ data RelP S : Kind → B.Rel (El S ⊥) (Level.suc (a ⊔ ℓ)) where
-- Congruences.
@@ -721,7 +721,7 @@ module AlternativeEquality {a ℓ} where
-- Proof WHNFs.
- data RelW S : Kind → B.Rel (El S ⊥) (suc (a ⊔ ℓ)) where
+ data RelW S : Kind → B.Rel (El S ⊥) (Level.suc (a ⊔ ℓ)) where
now : ∀ {k x y} (xRy : x ⟨ Eq S ⟩ y) → RelW S k (now x) (now y)
later : ∀ {k x y} (x∼y : RelP S k (♭ x) (♭ y)) → RelW S k (later x) (later y)
laterʳ : ∀ {x y} (x≈y : RelW S (other weak) x (♭ y)) → RelW S (other weak) x (later y)
@@ -861,7 +861,7 @@ module AlternativeEquality {a ℓ} where
-- equivalence).
correct : ∀ {S k x y} → RelP S k x y ⇔ Rel (Eq S) k x y
- correct = equivalent sound complete
+ correct = equivalence sound complete
------------------------------------------------------------------------
-- Example
diff --git a/src/Category/Monad/State.agda b/src/Category/Monad/State.agda
index 670a910..4a45571 100644
--- a/src/Category/Monad/State.agda
+++ b/src/Category/Monad/State.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- The state monad
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Category.Monad.State where
open import Category.Applicative.Indexed
diff --git a/src/Coinduction.agda b/src/Coinduction.agda
index 3fe1592..39d1e87 100644
--- a/src/Coinduction.agda
+++ b/src/Coinduction.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Basic types related to coinduction
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Coinduction where
import Level
diff --git a/src/Data/AVL.agda b/src/Data/AVL.agda
index 2f6b051..e2307d6 100644
--- a/src/Data/AVL.agda
+++ b/src/Data/AVL.agda
@@ -1,31 +1,90 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- AVL trees
------------------------------------------------------------------------
--- AVL trees are balanced binary search trees. The search tree
--- invariant is not statically enforced in the current implementation,
--- just the balance invariant.
+-- AVL trees are balanced binary search trees.
+
+-- The search tree invariant is specified using the technique
+-- described by Conor McBride in his talk "Pivotal pragmatism".
-open import Level
open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
-module Data.AVL (OrderedKeySet : StrictTotalOrder zero zero zero)
- (Value : StrictTotalOrder.Carrier OrderedKeySet → Set)
- where
+module Data.AVL
+ {k v ℓ}
+ {Key : Set k} (Value : Key → Set v)
+ {_<_ : Rel Key ℓ}
+ (isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_)
+ where
-open import Data.Nat hiding (compare)
-open StrictTotalOrder OrderedKeySet renaming (Carrier to Key)
-open import Data.Product
-open import Data.Maybe
open import Data.Bool
-open import Data.List as List using (List)
import Data.DifferenceList as DiffList
-open import Relation.Binary.PropositionalEquality
+open import Data.Empty
+open import Data.List as List using (List)
+open import Data.Maybe
+open import Data.Nat hiding (_<_; compare; _⊔_)
+open import Data.Product
+open import Data.Unit
+open import Function
+open import Level using (_⊔_; Lift; lift)
+
+open IsStrictTotalOrder isStrictTotalOrder
------------------------------------------------------------------------
--- Types and functions which are used to keep track of invariants
+-- Extended keys
+
+module Extended-key where
+
+ -- The key type extended with a new minimum and maximum.
+
+ data Key⁺ : Set k where
+ ⊥⁺ ⊤⁺ : Key⁺
+ [_] : (k : Key) → Key⁺
+
+ -- An extended strict ordering relation.
+
+ infix 4 _<⁺_
+
+ _<⁺_ : Key⁺ → Key⁺ → Set ℓ
+ ⊥⁺ <⁺ [ _ ] = Lift ⊤
+ ⊥⁺ <⁺ ⊤⁺ = Lift ⊤
+ [ x ] <⁺ [ y ] = x < y
+ [ _ ] <⁺ ⊤⁺ = Lift ⊤
+ _ <⁺ _ = Lift ⊥
+
+ -- A pair of ordering constraints.
-module Invariants where
+ infix 4 _<_<_
+
+ _<_<_ : Key⁺ → Key → Key⁺ → Set ℓ
+ l < x < u = l <⁺ [ x ] × [ x ] <⁺ u
+
+ -- _<⁺_ is transitive.
+
+ trans⁺ : ∀ l {m u} → l <⁺ m → m <⁺ u → l <⁺ u
+
+ trans⁺ [ l ] {m = [ m ]} {u = [ u ]} l<m m<u = trans l<m m<u
+
+ trans⁺ ⊥⁺ {u = [ _ ]} _ _ = _
+ trans⁺ ⊥⁺ {u = ⊤⁺} _ _ = _
+ trans⁺ [ _ ] {u = ⊤⁺} _ _ = _
+
+ trans⁺ _ {m = ⊥⁺} {u = ⊥⁺} _ (lift ())
+ trans⁺ _ {m = [ _ ]} {u = ⊥⁺} _ (lift ())
+ trans⁺ _ {m = ⊤⁺} {u = ⊥⁺} _ (lift ())
+ trans⁺ [ _ ] {m = ⊥⁺} {u = [ _ ]} (lift ()) _
+ trans⁺ [ _ ] {m = ⊤⁺} {u = [ _ ]} _ (lift ())
+ trans⁺ ⊤⁺ {m = ⊥⁺} (lift ()) _
+ trans⁺ ⊤⁺ {m = [ _ ]} (lift ()) _
+ trans⁺ ⊤⁺ {m = ⊤⁺} (lift ()) _
+
+------------------------------------------------------------------------
+-- Types and functions which are used to keep track of height
+-- invariants
+
+module Height-invariants where
-- Bits. (I would use Fin 2 instead if Agda had "defined patterns",
-- so that I could pattern match on 1# instead of suc zero; the text
@@ -92,205 +151,241 @@ module Invariants where
max-lemma : ∀ {m n} (bal : m ∼ n) →
1 + max (1+ (max∼max bal)) ≡ 2 + max bal
- max-lemma ∼+ = refl
- max-lemma ∼0 = refl
- max-lemma ∼- = refl
+ max-lemma ∼+ = P.refl
+ max-lemma ∼0 = P.refl
+ max-lemma ∼- = P.refl
------------------------------------------------------------------------
-- AVL trees
-- Key/value pairs.
-KV : Set
+KV : Set (k ⊔ v)
KV = Σ Key Value
module Indexed where
- open Invariants
+ open Extended-key
+ open Height-invariants
- -- The trees are indexed on their height. (bal is the balance
- -- factor.)
+ -- The trees have three parameters/indices: a lower bound on the
+ -- keys, an upper bound, and a height.
+ --
+ -- (The bal argument is the balance factor.)
- data Tree : ℕ → Set where
- leaf : Tree 0
+ data Tree (l u : Key⁺) : ℕ → Set (k ⊔ v ⊔ ℓ) where
+ leaf : (l<u : l <⁺ u) → Tree l u 0
node : ∀ {hˡ hʳ}
- (l : Tree hˡ) (k : KV) (r : Tree hʳ) (bal : hˡ ∼ hʳ) →
- Tree (1 + max bal)
+ (k : KV)
+ (lk : Tree l [ proj₁ k ] hˡ)
+ (ku : Tree [ proj₁ k ] u hʳ) (bal : hˡ ∼ hʳ) →
+ Tree l u (1 + max bal)
+
+ -- Cast operations. Logarithmic in the size of the tree, if we don't
+ -- count the time needed to construct the new proofs in the leaf
+ -- cases. (The same kind of caveat applies to other operations
+ -- below.)
+ --
+ -- Perhaps it would be worthwhile changing the data structure so
+ -- that the casts could be implemented in constant time (excluding
+ -- proof manipulation). However, note that this would not change the
+ -- worst-case time complexity of the operations below (up to θ).
+
+ castˡ : ∀ {l m u h} → l <⁺ m → Tree m u h → Tree l u h
+ castˡ {l} l<m (leaf m<u) = leaf (trans⁺ l l<m m<u)
+ castˡ l<m (node k mk ku bal) = node k (castˡ l<m mk) ku bal
+
+ castʳ : ∀ {l m u h} → Tree l m h → m <⁺ u → Tree l u h
+ castʳ {l} (leaf l<m) m<u = leaf (trans⁺ l l<m m<u)
+ castʳ (node k lk km bal) m<u = node k lk (castʳ km m<u) bal
-- Various constant-time functions which construct trees out of
-- smaller pieces, sometimes using rotation.
- joinˡ⁺ : ∀ {hˡ hʳ} →
- (∃ λ i → Tree (i ⊕ hˡ)) → KV → Tree hʳ → (bal : hˡ ∼ hʳ) →
- ∃ λ i → Tree (i ⊕ (1 + max bal))
- joinˡ⁺ (1# , node t₁ k₂
- (node t₃ k₄ t₅ bal)
- ∼+) k₆ t₇ ∼- = (0# , subst Tree (max-lemma bal)
- (node (node t₁ k₂ t₃ (max∼ bal))
- k₄
- (node t₅ k₆ t₇ (∼max bal))
+ joinˡ⁺ : ∀ {l u hˡ hʳ} →
+ (k : KV) →
+ (∃ λ i → Tree l [ proj₁ k ] (i ⊕ hˡ)) →
+ Tree [ proj₁ k ] u hʳ →
+ (bal : hˡ ∼ hʳ) →
+ ∃ λ i → Tree l u (i ⊕ (1 + max bal))
+ joinˡ⁺ k₆ (1# , node k₂ t₁
+ (node k₄ t₃ t₅ bal)
+ ∼+) t₇ ∼- = (0# , P.subst (Tree _ _) (max-lemma bal)
+ (node k₄
+ (node k₂ t₁ t₃ (max∼ bal))
+ (node k₆ t₅ t₇ (∼max bal))
(1+ (max∼max bal))))
- joinˡ⁺ (1# , node t₁ k₂ t₃ ∼-) k₄ t₅ ∼- = (0# , node t₁ k₂ (node t₃ k₄ t₅ ∼0) ∼0)
- joinˡ⁺ (1# , node t₁ k₂ t₃ ∼0) k₄ t₅ ∼- = (1# , node t₁ k₂ (node t₃ k₄ t₅ ∼-) ∼+)
- joinˡ⁺ (1# , t₁) k₂ t₃ ∼0 = (1# , node t₁ k₂ t₃ ∼-)
- joinˡ⁺ (1# , t₁) k₂ t₃ ∼+ = (0# , node t₁ k₂ t₃ ∼0)
- joinˡ⁺ (0# , t₁) k₂ t₃ bal = (0# , node t₁ k₂ t₃ bal)
-
- joinʳ⁺ : ∀ {hˡ hʳ} →
- Tree hˡ → KV → (∃ λ i → Tree (i ⊕ hʳ)) → (bal : hˡ ∼ hʳ) →
- ∃ λ i → Tree (i ⊕ (1 + max bal))
- joinʳ⁺ t₁ k₂ (1# , node
- (node t₃ k₄ t₅ bal)
- k₆ t₇ ∼-) ∼+ = (0# , subst Tree (max-lemma bal)
- (node (node t₁ k₂ t₃ (max∼ bal))
- k₄
- (node t₅ k₆ t₇ (∼max bal))
+ joinˡ⁺ k₄ (1# , node k₂ t₁ t₃ ∼-) t₅ ∼- = (0# , node k₂ t₁ (node k₄ t₃ t₅ ∼0) ∼0)
+ joinˡ⁺ k₄ (1# , node k₂ t₁ t₃ ∼0) t₅ ∼- = (1# , node k₂ t₁ (node k₄ t₃ t₅ ∼-) ∼+)
+ joinˡ⁺ k₂ (1# , t₁) t₃ ∼0 = (1# , node k₂ t₁ t₃ ∼-)
+ joinˡ⁺ k₂ (1# , t₁) t₃ ∼+ = (0# , node k₂ t₁ t₃ ∼0)
+ joinˡ⁺ k₂ (0# , t₁) t₃ bal = (0# , node k₂ t₁ t₃ bal)
+
+ joinʳ⁺ : ∀ {l u hˡ hʳ} →
+ (k : KV) →
+ Tree l [ proj₁ k ] hˡ →
+ (∃ λ i → Tree [ proj₁ k ] u (i ⊕ hʳ)) →
+ (bal : hˡ ∼ hʳ) →
+ ∃ λ i → Tree l u (i ⊕ (1 + max bal))
+ joinʳ⁺ k₂ t₁ (1# , node k₆
+ (node k₄ t₃ t₅ bal)
+ t₇ ∼-) ∼+ = (0# , P.subst (Tree _ _) (max-lemma bal)
+ (node k₄
+ (node k₂ t₁ t₃ (max∼ bal))
+ (node k₆ t₅ t₇ (∼max bal))
(1+ (max∼max bal))))
- joinʳ⁺ t₁ k₂ (1# , node t₃ k₄ t₅ ∼+) ∼+ = (0# , node (node t₁ k₂ t₃ ∼0) k₄ t₅ ∼0)
- joinʳ⁺ t₁ k₂ (1# , node t₃ k₄ t₅ ∼0) ∼+ = (1# , node (node t₁ k₂ t₃ ∼+) k₄ t₅ ∼-)
- joinʳ⁺ t₁ k₂ (1# , t₃) ∼0 = (1# , node t₁ k₂ t₃ ∼+)
- joinʳ⁺ t₁ k₂ (1# , t₃) ∼- = (0# , node t₁ k₂ t₃ ∼0)
- joinʳ⁺ t₁ k₂ (0# , t₃) bal = (0# , node t₁ k₂ t₃ bal)
-
- joinˡ⁻ : ∀ hˡ {hʳ} →
- (∃ λ i → Tree (i ⊕ hˡ -1)) → KV → Tree hʳ → (bal : hˡ ∼ hʳ) →
- ∃ λ i → Tree (i ⊕ max bal)
- joinˡ⁻ zero (0# , t₁) k₂ t₃ bal = (1# , node t₁ k₂ t₃ bal)
- joinˡ⁻ zero (1# , t₁) k₂ t₃ bal = (1# , node t₁ k₂ t₃ bal)
- joinˡ⁻ (suc _) (0# , t₁) k₂ t₃ ∼+ = joinʳ⁺ t₁ k₂ (1# , t₃) ∼+
- joinˡ⁻ (suc _) (0# , t₁) k₂ t₃ ∼0 = (1# , node t₁ k₂ t₃ ∼+)
- joinˡ⁻ (suc _) (0# , t₁) k₂ t₃ ∼- = (0# , node t₁ k₂ t₃ ∼0)
- joinˡ⁻ (suc _) (1# , t₁) k₂ t₃ bal = (1# , node t₁ k₂ t₃ bal)
-
- joinʳ⁻ : ∀ {hˡ} hʳ →
- Tree hˡ → KV → (∃ λ i → Tree (i ⊕ hʳ -1)) → (bal : hˡ ∼ hʳ) →
- ∃ λ i → Tree (i ⊕ max bal)
- joinʳ⁻ zero t₁ k₂ (0# , t₃) bal = (1# , node t₁ k₂ t₃ bal)
- joinʳ⁻ zero t₁ k₂ (1# , t₃) bal = (1# , node t₁ k₂ t₃ bal)
- joinʳ⁻ (suc _) t₁ k₂ (0# , t₃) ∼- = joinˡ⁺ (1# , t₁) k₂ t₃ ∼-
- joinʳ⁻ (suc _) t₁ k₂ (0# , t₃) ∼0 = (1# , node t₁ k₂ t₃ ∼-)
- joinʳ⁻ (suc _) t₁ k₂ (0# , t₃) ∼+ = (0# , node t₁ k₂ t₃ ∼0)
- joinʳ⁻ (suc _) t₁ k₂ (1# , t₃) bal = (1# , node t₁ k₂ t₃ bal)
+ joinʳ⁺ k₂ t₁ (1# , node k₄ t₃ t₅ ∼+) ∼+ = (0# , node k₄ (node k₂ t₁ t₃ ∼0) t₅ ∼0)
+ joinʳ⁺ k₂ t₁ (1# , node k₄ t₃ t₅ ∼0) ∼+ = (1# , node k₄ (node k₂ t₁ t₃ ∼+) t₅ ∼-)
+ joinʳ⁺ k₂ t₁ (1# , t₃) ∼0 = (1# , node k₂ t₁ t₃ ∼+)
+ joinʳ⁺ k₂ t₁ (1# , t₃) ∼- = (0# , node k₂ t₁ t₃ ∼0)
+ joinʳ⁺ k₂ t₁ (0# , t₃) bal = (0# , node k₂ t₁ t₃ bal)
+
+ joinˡ⁻ : ∀ {l u} hˡ {hʳ} →
+ (k : KV) →
+ (∃ λ i → Tree l [ proj₁ k ] (i ⊕ hˡ -1)) →
+ Tree [ proj₁ k ] u hʳ →
+ (bal : hˡ ∼ hʳ) →
+ ∃ λ i → Tree l u (i ⊕ max bal)
+ joinˡ⁻ zero k₂ (0# , t₁) t₃ bal = (1# , node k₂ t₁ t₃ bal)
+ joinˡ⁻ zero k₂ (1# , t₁) t₃ bal = (1# , node k₂ t₁ t₃ bal)
+ joinˡ⁻ (suc _) k₂ (0# , t₁) t₃ ∼+ = joinʳ⁺ k₂ t₁ (1# , t₃) ∼+
+ joinˡ⁻ (suc _) k₂ (0# , t₁) t₃ ∼0 = (1# , node k₂ t₁ t₃ ∼+)
+ joinˡ⁻ (suc _) k₂ (0# , t₁) t₃ ∼- = (0# , node k₂ t₁ t₃ ∼0)
+ joinˡ⁻ (suc _) k₂ (1# , t₁) t₃ bal = (1# , node k₂ t₁ t₃ bal)
+
+ joinʳ⁻ : ∀ {l u hˡ} hʳ →
+ (k : KV) →
+ Tree l [ proj₁ k ] hˡ →
+ (∃ λ i → Tree [ proj₁ k ] u (i ⊕ hʳ -1)) →
+ (bal : hˡ ∼ hʳ) →
+ ∃ λ i → Tree l u (i ⊕ max bal)
+ joinʳ⁻ zero k₂ t₁ (0# , t₃) bal = (1# , node k₂ t₁ t₃ bal)
+ joinʳ⁻ zero k₂ t₁ (1# , t₃) bal = (1# , node k₂ t₁ t₃ bal)
+ joinʳ⁻ (suc _) k₂ t₁ (0# , t₃) ∼- = joinˡ⁺ k₂ (1# , t₁) t₃ ∼-
+ joinʳ⁻ (suc _) k₂ t₁ (0# , t₃) ∼0 = (1# , node k₂ t₁ t₃ ∼-)
+ joinʳ⁻ (suc _) k₂ t₁ (0# , t₃) ∼+ = (0# , node k₂ t₁ t₃ ∼0)
+ joinʳ⁻ (suc _) k₂ t₁ (1# , t₃) bal = (1# , node k₂ t₁ t₃ bal)
-- Extracts the smallest element from the tree, plus the rest.
-- Logarithmic in the size of the tree.
- headTail : ∀ {h} → Tree (1 + h) →
- KV × ∃ λ i → Tree (i ⊕ h)
- headTail (node leaf k₁ t₂ ∼+) = (k₁ , 0# , t₂)
- headTail (node leaf k₁ t₂ ∼0) = (k₁ , 0# , t₂)
- headTail (node {hˡ = suc _} t₁₂ k₃ t₄ bal) with headTail t₁₂
- ... | (k₁ , t₂) = (k₁ , joinˡ⁻ _ t₂ k₃ t₄ bal)
+ headTail : ∀ {l u h} → Tree l u (1 + h) →
+ ∃ λ (k : KV) → l <⁺ [ proj₁ k ] ×
+ ∃ λ i → Tree [ proj₁ k ] u (i ⊕ h)
+ headTail (node k₁ (leaf l<k₁) t₂ ∼+) = (k₁ , l<k₁ , 0# , t₂)
+ headTail (node k₁ (leaf l<k₁) t₂ ∼0) = (k₁ , l<k₁ , 0# , t₂)
+ headTail (node {hˡ = suc _} k₃ t₁₂ t₄ bal) with headTail t₁₂
+ ... | (k₁ , l<k₁ , t₂) = (k₁ , l<k₁ , joinˡ⁻ _ k₃ t₂ t₄ bal)
-- Extracts the largest element from the tree, plus the rest.
-- Logarithmic in the size of the tree.
- initLast : ∀ {h} → Tree (1 + h) →
- ∃ (λ i → Tree (i ⊕ h)) × KV
- initLast (node t₁ k₂ leaf ∼-) = ((0# , t₁) , k₂)
- initLast (node t₁ k₂ leaf ∼0) = ((0# , t₁) , k₂)
- initLast (node {hʳ = suc _} t₁ k₂ t₃₄ bal) with initLast t₃₄
- ... | (t₃ , k₄) = (joinʳ⁻ _ t₁ k₂ t₃ bal , k₄)
-
- -- Another joining function. Logarithmic in the size of the tree.
-
- join : ∀ {hˡ hʳ} →
- Tree hˡ → Tree hʳ → (bal : hˡ ∼ hʳ) →
- ∃ λ i → Tree (i ⊕ max bal)
- join t₁ leaf ∼0 = (0# , t₁)
- join t₁ leaf ∼- = (0# , t₁)
+ initLast : ∀ {l u h} → Tree l u (1 + h) →
+ ∃ λ (k : KV) → [ proj₁ k ] <⁺ u ×
+ ∃ λ i → Tree l [ proj₁ k ] (i ⊕ h)
+ initLast (node k₂ t₁ (leaf k₂<u) ∼-) = (k₂ , k₂<u , (0# , t₁))
+ initLast (node k₂ t₁ (leaf k₂<u) ∼0) = (k₂ , k₂<u , (0# , t₁))
+ initLast (node {hʳ = suc _} k₂ t₁ t₃₄ bal) with initLast t₃₄
+ ... | (k₄ , k₄<u , t₃) = (k₄ , k₄<u , joinʳ⁻ _ k₂ t₁ t₃ bal)
+
+ -- Another joining function. Logarithmic in the size of either of
+ -- the input trees (which need to have almost equal heights).
+
+ join : ∀ {l m u hˡ hʳ} →
+ Tree l m hˡ → Tree m u hʳ → (bal : hˡ ∼ hʳ) →
+ ∃ λ i → Tree l u (i ⊕ max bal)
+ join t₁ (leaf m<u) ∼0 = (0# , castʳ t₁ m<u)
+ join t₁ (leaf m<u) ∼- = (0# , castʳ t₁ m<u)
join {hʳ = suc _} t₁ t₂₃ bal with headTail t₂₃
- ... | (k₂ , t₃) = joinʳ⁻ _ t₁ k₂ t₃ bal
+ ... | (k₂ , m<k₂ , t₃) = joinʳ⁻ _ k₂ (castʳ t₁ m<k₂) t₃ bal
-- An empty tree.
- empty : Tree 0
+ empty : ∀ {l u} → l <⁺ u → Tree l u 0
empty = leaf
-- A singleton tree.
- singleton : (k : Key) → Value k → Tree 1
- singleton k v = node leaf (k , v) leaf ∼0
+ singleton : ∀ {l u} (k : Key) → Value k → l < k < u → Tree l u 1
+ singleton k v (l<k , k<u) = node (k , v) (leaf l<k) (leaf k<u) ∼0
-- Inserts a key into the tree. If the key already exists, then it
-- is replaced. Logarithmic in the size of the tree (assuming
-- constant-time comparisons).
- insert : ∀ {h} → (k : Key) → Value k → Tree h →
- ∃ λ i → Tree (i ⊕ h)
- insert k v leaf = (1# , singleton k v)
- insert k v (node l p r bal) with compare k (proj₁ p)
- ... | tri< _ _ _ = joinˡ⁺ (insert k v l) p r bal
- ... | tri≈ _ _ _ = (0# , node l (k , v) r bal)
- ... | tri> _ _ _ = joinʳ⁺ l p (insert k v r) bal
+ insert : ∀ {l u h} → (k : Key) → Value k → Tree l u h → l < k < u →
+ ∃ λ i → Tree l u (i ⊕ h)
+ insert k v (leaf l<u) l<k<u = (1# , singleton k v l<k<u)
+ insert k v (node p lp pu bal) (l<k , k<u) with compare k (proj₁ p)
+ ... | tri< k<p _ _ = joinˡ⁺ p (insert k v lp (l<k , k<p)) pu bal
+ ... | tri> _ _ p<k = joinʳ⁺ p lp (insert k v pu (p<k , k<u)) bal
+ ... | tri≈ _ k≡p _ rewrite P.sym k≡p = (0# , node (k , v) lp pu bal)
-- Deletes the key/value pair containing the given key, if any.
-- Logarithmic in the size of the tree (assuming constant-time
-- comparisons).
- delete : ∀ {h} → Key → Tree h →
- ∃ λ i → Tree (i ⊕ h -1)
- delete k leaf = (0# , leaf)
- delete k (node l p r bal) with compare k (proj₁ p)
- ... | tri< _ _ _ = joinˡ⁻ _ (delete k l) p r bal
- ... | tri≈ _ _ _ = join l r bal
- ... | tri> _ _ _ = joinʳ⁻ _ l p (delete k r) bal
-
- -- Looks up a key in the tree. Logarithmic in the size of the tree
- -- (assuming constant-time comparisons).
-
- lookup : ∀ {h} → (k : Key) → Tree h →
- Maybe (∃ λ k′ → Value k′ × k ≈ k′)
- lookup k leaf = nothing
- lookup k (node l (k′ , v) r _) with compare k k′
- ... | tri< _ _ _ = lookup k l
- ... | tri≈ _ eq _ = just (k′ , v , eq)
- ... | tri> _ _ _ = lookup k r
+ delete : ∀ {l u h} → Key → Tree l u h →
+ ∃ λ i → Tree l u (i ⊕ h -1)
+ delete k (leaf l<u) = (0# , leaf l<u)
+ delete k (node p lp pu bal) with compare k (proj₁ p)
+ ... | tri< _ _ _ = joinˡ⁻ _ p (delete k lp) pu bal
+ ... | tri> _ _ _ = joinʳ⁻ _ p lp (delete k pu) bal
+ ... | tri≈ _ _ _ = join lp pu bal
+
+ -- Looks up a key. Logarithmic in the size of the tree (assuming
+ -- constant-time comparisons).
+
+ lookup : ∀ {l u h} → (k : Key) → Tree l u h → Maybe (Value k)
+ lookup k (leaf _) = nothing
+ lookup k (node (k′ , v) lk′ k′u _) with compare k k′
+ ... | tri< _ _ _ = lookup k lk′
+ ... | tri> _ _ _ = lookup k k′u
+ ... | tri≈ _ eq _ rewrite eq = just v
-- Converts the tree to an ordered list. Linear in the size of the
-- tree.
open DiffList
- toDiffList : ∀ {h} → Tree h → DiffList KV
- toDiffList leaf = []
- toDiffList (node l k r _) = toDiffList l ++ [ k ] ++ toDiffList r
+ toDiffList : ∀ {l u h} → Tree l u h → DiffList KV
+ toDiffList (leaf _) = []
+ toDiffList (node k l r _) = toDiffList l ++ k ∷ toDiffList r
------------------------------------------------------------------------
-- Types and functions with hidden indices
-data Tree : Set where
- tree : ∀ {h} → Indexed.Tree h → Tree
+data Tree : Set (k ⊔ v ⊔ ℓ) where
+ tree : let open Extended-key in
+ ∀ {h} → Indexed.Tree ⊥⁺ ⊤⁺ h → Tree
empty : Tree
-empty = tree Indexed.empty
+empty = tree (Indexed.empty _)
singleton : (k : Key) → Value k → Tree
-singleton k v = tree (Indexed.singleton k v)
+singleton k v = tree (Indexed.singleton k v _)
insert : (k : Key) → Value k → Tree → Tree
-insert k v (tree t) with Indexed.insert k v t
-... | (_ , t′) = tree t′
+insert k v (tree t) = tree $ proj₂ $ Indexed.insert k v t _
delete : Key → Tree → Tree
-delete k (tree t) with Indexed.delete k t
-... | (_ , t′) = tree t′
+delete k (tree t) = tree $ proj₂ $ Indexed.delete k t
-lookup : (k : Key) → Tree → Maybe (∃ λ k′ → Value k′ × k ≈ k′)
+lookup : (k : Key) → Tree → Maybe (Value k)
lookup k (tree t) = Indexed.lookup k t
_∈?_ : Key → Tree → Bool
k ∈? t = maybeToBool (lookup k t)
headTail : Tree → Maybe (KV × Tree)
-headTail (tree Indexed.leaf) = nothing
-headTail (tree {h = suc _} t) with Indexed.headTail t
-... | (k , _ , t′) = just (k , tree t′)
+headTail (tree (Indexed.leaf _)) = nothing
+headTail (tree {h = suc _} t) with Indexed.headTail t
+... | (k , _ , _ , t′) = just (k , tree (Indexed.castˡ _ t′))
initLast : Tree → Maybe (Tree × KV)
-initLast (tree Indexed.leaf) = nothing
-initLast (tree {h = suc _} t) with Indexed.initLast t
-... | ((_ , t′) , k) = just (tree t′ , k)
+initLast (tree (Indexed.leaf _)) = nothing
+initLast (tree {h = suc _} t) with Indexed.initLast t
+... | (k , _ , _ , t′) = just (tree (Indexed.castʳ t′ _) , k)
-- The input does not need to be ordered.
diff --git a/src/Data/AVL/IndexedMap.agda b/src/Data/AVL/IndexedMap.agda
index f5f2d59..b733338 100644
--- a/src/Data/AVL/IndexedMap.agda
+++ b/src/Data/AVL/IndexedMap.agda
@@ -1,31 +1,33 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Finite maps with indexed keys and values, based on AVL trees
------------------------------------------------------------------------
+open import Data.Product as Prod
open import Relation.Binary
-open import Relation.Binary.PropositionalEquality
-open import Data.Product as Prod hiding (map)
-open import Level
+open import Relation.Binary.PropositionalEquality using (_≡_)
module Data.AVL.IndexedMap
- {Index : Set} {Key : Index → Set} {_≈_ _<_ : Rel (∃ Key) zero}
- (isOrderedKeySet : IsStrictTotalOrder _≈_ _<_)
- -- Equal keys must have equal indices.
- (indicesEqual : _≈_ =[ proj₁ ]⇒ _≡_)
- (Value : Index → Set)
- where
+ {i k v ℓ}
+ {Index : Set i} {Key : Index → Set k} (Value : Index → Set v)
+ {_<_ : Rel (∃ Key) ℓ}
+ (isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_)
+ where
+open import Category.Functor
import Data.AVL
-open import Function
-open import Data.Maybe as Maybe
open import Data.Bool
open import Data.List as List using (List)
-open import Category.Functor
-open RawFunctor Maybe.functor
+open import Data.Maybe as Maybe
+open import Function
+open import Level
+
+open RawFunctor (Maybe.functor {f = i ⊔ k ⊔ v ⊔ ℓ})
-- Key/value pairs.
-KV : Set
+KV : Set (i ⊔ k ⊔ v)
KV = ∃ λ i → Key i × Value i
-- Conversions.
@@ -38,15 +40,11 @@ private
toKV : Σ[ ik ∶ ∃ Key ] Value (proj₁ ik) → KV
toKV ((i , k) , v) = (i , k , v)
-private
-
- Order : StrictTotalOrder _ _ _
- Order = record { isStrictTotalOrder = isOrderedKeySet }
-
-- The map type.
private
- open module AVL = Data.AVL Order (λ ik → Value (proj₁ ik))
+ open module AVL =
+ Data.AVL (λ ik → Value (proj₁ ik)) isStrictTotalOrder
public using () renaming (Tree to Map)
-- Repackaged functions.
@@ -64,10 +62,7 @@ delete : ∀ {i} → Key i → Map → Map
delete k = AVL.delete (, k)
lookup : ∀ {i} → Key i → Map → Maybe (Value i)
-lookup k m with AVL.lookup (_ , k) m
-... | nothing = nothing
-... | just ((i′ , k′) , v′ , eq) with indicesEqual eq
-... | refl = just v′
+lookup k m = AVL.lookup (, k) m
_∈?_ : ∀ {i} → Key i → Map → Bool
_∈?_ k = AVL._∈?_ (, k)
diff --git a/src/Data/AVL/Sets.agda b/src/Data/AVL/Sets.agda
index 93fd8fa..bf74ae9 100644
--- a/src/Data/AVL/Sets.agda
+++ b/src/Data/AVL/Sets.agda
@@ -1,28 +1,33 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Finite sets, based on AVL trees
------------------------------------------------------------------------
-open import Level
open import Relation.Binary
+open import Relation.Binary.PropositionalEquality using (_≡_)
module Data.AVL.Sets
- (OrderedKeySet : StrictTotalOrder zero zero zero) where
+ {k ℓ} {Key : Set k} {_<_ : Rel Key ℓ}
+ (isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_)
+ where
+open import Category.Functor
import Data.AVL as AVL
-open StrictTotalOrder OrderedKeySet renaming (Carrier to Key)
-open import Data.Unit
-open import Function
-open import Data.Product as Prod using (_×_; _,_; proj₁)
-open import Data.Maybe as Maybe
open import Data.Bool
open import Data.List as List using (List)
-open import Category.Functor
-open RawFunctor Maybe.functor
+open import Data.Maybe as Maybe
+open import Data.Product as Prod using (_×_; _,_; proj₁)
+open import Data.Unit
+open import Function
+open import Level
+
+open RawFunctor (Maybe.functor {f = k ⊔ ℓ})
-- The set type. (Note that Set is a reserved word.)
private
- open module S = AVL OrderedKeySet (const ⊤)
+ open module S = AVL (const ⊤) isStrictTotalOrder
public using () renaming (Tree to ⟨Set⟩)
-- Repackaged functions.
@@ -39,9 +44,6 @@ insert k = S.insert k _
delete : Key → ⟨Set⟩ → ⟨Set⟩
delete = S.delete
-lookup : Key → ⟨Set⟩ → Maybe Key
-lookup k s = proj₁ <$> S.lookup k s
-
_∈?_ : Key → ⟨Set⟩ → Bool
_∈?_ = S._∈?_
diff --git a/src/Data/Bin.agda b/src/Data/Bin.agda
index 9ef0611..7d9605e 100644
--- a/src/Data/Bin.agda
+++ b/src/Data/Bin.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- A binary representation of natural numbers
------------------------------------------------------------------------
diff --git a/src/Data/Bool.agda b/src/Data/Bool.agda
index d1502a4..8c80424 100644
--- a/src/Data/Bool.agda
+++ b/src/Data/Bool.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Booleans
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Bool where
open import Function
diff --git a/src/Data/Bool/Properties.agda b/src/Data/Bool/Properties.agda
index 4debaa0..d154f37 100644
--- a/src/Data/Bool/Properties.agda
+++ b/src/Data/Bool/Properties.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- A bunch of properties
------------------------------------------------------------------------
@@ -9,15 +11,15 @@ open import Data.Fin
open import Function
open import Function.Equality using (_⟨$⟩_)
open import Function.Equivalence
- using (_⇔_; equivalent; module Equivalent)
+ using (_⇔_; equivalence; module Equivalence)
open import Algebra
open import Algebra.Structures
import Algebra.RingSolver.Simple as Solver
import Algebra.RingSolver.AlmostCommutativeRing as ACR
-open import Relation.Binary.PropositionalEquality
- hiding (proof-irrelevance)
-open ≡-Reasoning
-import Algebra.FunctionProperties as P; open P _≡_
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; _≢_; refl)
+open P.≡-Reasoning
+import Algebra.FunctionProperties as FP; open FP (_≡_ {A = Bool})
open import Data.Product
open import Data.Sum
open import Data.Empty
@@ -67,7 +69,7 @@ private
x ∧ (y ∨ z)
≡⟨ distˡ x y z ⟩
x ∧ y ∨ x ∧ z
- ≡⟨ cong₂ _∨_ (∧-comm x y) (∧-comm x z) ⟩
+ ≡⟨ P.cong₂ _∨_ (∧-comm x y) (∧-comm x z) ⟩
y ∧ x ∨ z ∧ x
@@ -76,18 +78,18 @@ isCommutativeSemiring-∨-∧
isCommutativeSemiring-∨-∧ = record
{ +-isCommutativeMonoid = record
{ isSemigroup = record
- { isEquivalence = isEquivalence
+ { isEquivalence = P.isEquivalence
; assoc = ∨-assoc
- ; ∙-cong = cong₂ _∨_
+ ; ∙-cong = P.cong₂ _∨_
}
; identityˡ = λ _ → refl
; comm = ∨-comm
}
; *-isCommutativeMonoid = record
{ isSemigroup = record
- { isEquivalence = isEquivalence
+ { isEquivalence = P.isEquivalence
; assoc = ∧-assoc
- ; ∙-cong = cong₂ _∧_
+ ; ∙-cong = P.cong₂ _∧_
}
; identityˡ = λ _ → refl
; comm = ∧-comm
@@ -128,7 +130,7 @@ private
x ∨ (y ∧ z)
≡⟨ distˡ x y z ⟩
(x ∨ y) ∧ (x ∨ z)
- ≡⟨ cong₂ _∧_ (∨-comm x y) (∨-comm x z) ⟩
+ ≡⟨ P.cong₂ _∧_ (∨-comm x y) (∨-comm x z) ⟩
(y ∨ x) ∧ (z ∨ x)
@@ -137,18 +139,18 @@ isCommutativeSemiring-∧-∨
isCommutativeSemiring-∧-∨ = record
{ +-isCommutativeMonoid = record
{ isSemigroup = record
- { isEquivalence = isEquivalence
+ { isEquivalence = P.isEquivalence
; assoc = ∧-assoc
- ; ∙-cong = cong₂ _∧_
+ ; ∙-cong = P.cong₂ _∧_
}
; identityˡ = λ _ → refl
; comm = ∧-comm
}
; *-isCommutativeMonoid = record
{ isSemigroup = record
- { isEquivalence = isEquivalence
+ { isEquivalence = P.isEquivalence
; assoc = ∨-assoc
- ; ∙-cong = cong₂ _∨_
+ ; ∙-cong = P.cong₂ _∨_
}
; identityˡ = λ _ → refl
; comm = ∨-comm
@@ -184,7 +186,7 @@ private
not-∧-inverse : Inverse false not _∧_
not-∧-inverse =
- ¬x∧x≡⊥ , (λ x → ∧-comm x (not x) ⟨ trans ⟩ ¬x∧x≡⊥ x)
+ ¬x∧x≡⊥ , (λ x → ∧-comm x (not x) ⟨ P.trans ⟩ ¬x∧x≡⊥ x)
where
¬x∧x≡⊥ : LeftInverse false not _∧_
¬x∧x≡⊥ false = refl
@@ -192,7 +194,7 @@ private
not-∨-inverse : Inverse true not _∨_
not-∨-inverse =
- ¬x∨x≡⊤ , (λ x → ∨-comm x (not x) ⟨ trans ⟩ ¬x∨x≡⊤ x)
+ ¬x∨x≡⊤ , (λ x → ∨-comm x (not x) ⟨ P.trans ⟩ ¬x∨x≡⊤ x)
where
¬x∨x≡⊤ : LeftInverse true not _∨_
¬x∨x≡⊤ false = refl
@@ -202,20 +204,20 @@ isBooleanAlgebra : IsBooleanAlgebra _≡_ _∨_ _∧_ not true false
isBooleanAlgebra = record
{ isDistributiveLattice = record
{ isLattice = record
- { isEquivalence = isEquivalence
+ { isEquivalence = P.isEquivalence
; ∨-comm = ∨-comm
; ∨-assoc = ∨-assoc
- ; ∨-cong = cong₂ _∨_
+ ; ∨-cong = P.cong₂ _∨_
; ∧-comm = ∧-comm
; ∧-assoc = ∧-assoc
- ; ∧-cong = cong₂ _∧_
+ ; ∧-cong = P.cong₂ _∧_
; absorptive = absorptive
}
; ∨-∧-distribʳ = proj₂ distrib-∨-∧
}
; ∨-complementʳ = proj₂ not-∨-inverse
; ∧-complementʳ = proj₂ not-∧-inverse
- ; ¬-cong = cong not
+ ; ¬-cong = P.cong not
}
booleanAlgebra : BooleanAlgebra _ _
@@ -235,7 +237,7 @@ 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₂ CS.*-identity _
+ xor-is-ok false y = P.sym $ proj₂ CS.*-identity _
where module CS = CommutativeSemiring commutativeSemiring-∨-∧
commutativeRing-xor-∧ : CommutativeRing _ _
@@ -267,25 +269,25 @@ not-¬ {false} refl ()
⇔→≡ : {b₁ b₂ b : Bool} → b₁ ≡ b ⇔ b₂ ≡ b → b₁ ≡ b₂
⇔→≡ {true } {true } 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)
+⇔→≡ {true } {false} {true } hyp = P.sym (Equivalence.to hyp ⟨$⟩ refl)
+⇔→≡ {true } {false} {false} hyp = Equivalence.from hyp ⟨$⟩ refl
+⇔→≡ {false} {true } {true } hyp = Equivalence.from hyp ⟨$⟩ refl
+⇔→≡ {false} {true } {false} hyp = P.sym (Equivalence.to hyp ⟨$⟩ refl)
⇔→≡ {false} {false} hyp = refl
T-≡ : ∀ {b} → T b ⇔ b ≡ true
-T-≡ {false} = equivalent (λ ()) (λ ())
-T-≡ {true} = equivalent (const refl) (const _)
+T-≡ {false} = equivalence (λ ()) (λ ())
+T-≡ {true} = equivalence (const refl) (const _)
T-∧ : ∀ {b₁ b₂} → T (b₁ ∧ b₂) ⇔ (T b₁ × T b₂)
-T-∧ {true} {true} = equivalent (const (_ , _)) (const _)
-T-∧ {true} {false} = equivalent (λ ()) proj₂
-T-∧ {false} {_} = equivalent (λ ()) proj₁
+T-∧ {true} {true} = equivalence (const (_ , _)) (const _)
+T-∧ {true} {false} = equivalence (λ ()) proj₂
+T-∧ {false} {_} = equivalence (λ ()) proj₁
T-∨ : ∀ {b₁ b₂} → T (b₁ ∨ b₂) ⇔ (T b₁ ⊎ T b₂)
-T-∨ {true} {b₂} = equivalent inj₁ (const _)
-T-∨ {false} {true} = equivalent inj₂ (const _)
-T-∨ {false} {false} = equivalent inj₁ [ id , id ]
+T-∨ {true} {b₂} = equivalence inj₁ (const _)
+T-∨ {false} {true} = equivalence inj₂ (const _)
+T-∨ {false} {false} = equivalence inj₁ [ id , id ]
proof-irrelevance : ∀ {b} (p q : T b) → p ≡ q
proof-irrelevance {true} _ _ = refl
diff --git a/src/Data/Bool/Show.agda b/src/Data/Bool/Show.agda
index 007b77c..658a8a4 100644
--- a/src/Data/Bool/Show.agda
+++ b/src/Data/Bool/Show.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Showing booleans
------------------------------------------------------------------------
diff --git a/src/Data/BoundedVec.agda b/src/Data/BoundedVec.agda
index b839607..8195616 100644
--- a/src/Data/BoundedVec.agda
+++ b/src/Data/BoundedVec.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Bounded vectors
------------------------------------------------------------------------
@@ -54,6 +56,7 @@ abstract
subst (BoundedVec a) lemma
(bVec {m = suc m} xs)
where
+ lemma : n + (1 + m) ≡ 1 + (n + m)
lemma = solve 2 (λ m n → n :+ (con 1 :+ m) := con 1 :+ (n :+ m))
refl m n
@@ -66,7 +69,9 @@ abstract
fromList {a = a} xs =
subst (BoundedVec a) lemma
(bVec {m = zero} (Vec.fromList xs))
- where lemma = solve 1 (λ m → m :+ con 0 := m) refl _
+ where
+ lemma : List.length xs + 0 ≡ List.length xs
+ lemma = solve 1 (λ m → m :+ con 0 := m) refl _
toList : ∀ {a n} → BoundedVec a n → List a
toList (bVec xs) = Vec.toList xs
diff --git a/src/Data/BoundedVec/Inefficient.agda b/src/Data/BoundedVec/Inefficient.agda
index 93efcc1..42eeba7 100644
--- a/src/Data/BoundedVec/Inefficient.agda
+++ b/src/Data/BoundedVec/Inefficient.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Bounded vectors (inefficient, concrete implementation)
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- Vectors of a specified maximum length.
module Data.BoundedVec.Inefficient where
diff --git a/src/Data/Char.agda b/src/Data/Char.agda
index cb44c72..4c768f4 100644
--- a/src/Data/Char.agda
+++ b/src/Data/Char.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Characters
------------------------------------------------------------------------
diff --git a/src/Data/Cofin.agda b/src/Data/Cofin.agda
index 8cbc2a1..f086244 100644
--- a/src/Data/Cofin.agda
+++ b/src/Data/Cofin.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- "Finite" sets indexed on coinductive "natural" numbers
------------------------------------------------------------------------
diff --git a/src/Data/Colist.agda b/src/Data/Colist.agda
index 62a2dcf..8363c5c 100644
--- a/src/Data/Colist.agda
+++ b/src/Data/Colist.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Coinductive lists
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Colist where
open import Category.Monad
@@ -21,13 +21,13 @@ open import Data.BoundedVec.Inefficient as BVec
open import Data.Product using (_,_)
open import Data.Sum using (_⊎_; inj₁; inj₂)
open import Function
-open import Level
+open import Level using (_⊔_)
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
+open RawMonad (¬¬-Monad {p = Level.zero})
------------------------------------------------------------------------
-- The type
@@ -38,6 +38,9 @@ data Colist {a} (A : Set a) : Set a where
[] : Colist A
_∷_ : (x : A) (xs : ∞ (Colist A)) → Colist A
+{-# IMPORT Data.FFI #-}
+{-# COMPILED_DATA Colist Data.FFI.AgdaList [] (:) #-}
+
data Any {a p} {A : Set a} (P : A → Set p) :
Colist A → Set (a ⊔ p) where
here : ∀ {x xs} (px : P x) → Any P (x ∷ xs)
diff --git a/src/Data/Conat.agda b/src/Data/Conat.agda
index 4a13191..185e892 100644
--- a/src/Data/Conat.agda
+++ b/src/Data/Conat.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Coinductive "natural" numbers
------------------------------------------------------------------------
diff --git a/src/Data/Container.agda b/src/Data/Container.agda
index bdc148c..4e9db90 100644
--- a/src/Data/Container.agda
+++ b/src/Data/Container.agda
@@ -1,17 +1,19 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- 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 Function.Inverse using (_↔_; module Inverse)
+import Function.Related as Related
open import Level
-open import Relation.Binary using (Setoid; module Setoid)
+open import Relation.Binary
+ using (Setoid; module Setoid; Preorder; module Preorder)
open import Relation.Binary.PropositionalEquality as P
using (_≡_; _≗_; refl)
open import Relation.Unary using (_⊆_)
@@ -22,10 +24,10 @@ open import Relation.Unary using (_⊆_)
-- A container is a set of shapes, and for every shape a set of
-- positions.
-infix 5 _◃_
+infix 5 _▷_
record Container (ℓ : Level) : Set (suc ℓ) where
- constructor _◃_
+ constructor _▷_
field
Shape : Set ℓ
Position : Shape → Set ℓ
@@ -51,14 +53,8 @@ private
Eq⇒≡ : ∀ {c} {C : Container c} {X : Set c} {xs ys : ⟦ C ⟧ X} →
P.Extensionality c c → Eq _≡_ xs ys → xs ≡ ys
- Eq⇒≡ {C = C} {X} ext (s≡s′ , f≈f′) = helper s≡s′ f≈f′
- where
- helper : {s s′ : Shape C} (eq : s ≡ s′) →
- {f : Position C s → X}
- {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′)
+ Eq⇒≡ {xs = s , f} {ys = .s , f′} ext (refl , f≈f′) =
+ P.cong (_,_ s) (ext f≈f′)
setoid : ∀ {ℓ} → Container ℓ → Setoid ℓ ℓ → Setoid ℓ ℓ
setoid C X = record
@@ -76,26 +72,11 @@ setoid C X = record
_≈_ = Eq X._≈_
sym : {xs ys : ⟦ C ⟧ X.Carrier} → xs ≈ ys → ys ≈ xs
- sym (eq , f) = helper eq f
- where
- helper : {s s′ : Shape C} (eq : s ≡ s′) →
- {f : Position C s → X.Carrier}
- {f′ : Position C s′ → X.Carrier} →
- (∀ p → X._≈_ (f p) (f′ $ P.subst (Position C) eq p)) →
- (s′ , f′) ≈ (s , f)
- helper refl eq = (refl , X.sym ⟨∘⟩ eq)
+ sym {_ , _} {._ , _} (refl , f) = (refl , X.sym ⟨∘⟩ f)
trans : ∀ {xs ys : ⟦ C ⟧ X.Carrier} zs → xs ≈ ys → ys ≈ zs → xs ≈ zs
- trans zs (eq₁ , f₁) (eq₂ , f₂) = helper eq₁ eq₂ (proj₂ zs) f₁ f₂
- where
- helper : {s s′ s″ : Shape C} (eq₁ : s ≡ s′) (eq₂ : s′ ≡ s″) →
- {f : Position C s → X.Carrier}
- {f′ : Position C s′ → X.Carrier} →
- (f″ : Position C s″ → X.Carrier) →
- (∀ p → X._≈_ (f p) (f′ $ P.subst (Position C) eq₁ p)) →
- (∀ p → X._≈_ (f′ p) (f″ $ P.subst (Position C) eq₂ p)) →
- (s , f) ≈ (s″ , f″)
- helper refl refl _ eq₁ eq₂ = (refl , λ p → X.trans (eq₁ p) (eq₂ p))
+ trans {_ , _} {._ , _} (._ , _) (refl , f₁) (refl , f₂) =
+ (refl , λ p → X.trans (f₁ p) (f₂ p))
------------------------------------------------------------------------
-- Functoriality
@@ -112,7 +93,7 @@ module Map where
(xs : ⟦ C ⟧ X.Carrier) → Eq X._≈_ (map ⟨id⟩ xs) xs
identity {C = C} X xs = Setoid.refl (setoid C X)
- composition : ∀ {c} {C : Container c} {X Y} Z →
+ composition : ∀ {c} {C : Container c} {X Y : Set c} Z →
let module Z = Setoid Z in
(f : Y → Z.Carrier) (g : X → Y) (xs : ⟦ C ⟧ X) →
Eq Z._≈_ (map f (map g xs)) (map (f ⟨∘⟩ g) xs)
@@ -206,7 +187,7 @@ module Morphism where
record _⊸_ {c} (C₁ C₂ : Container c) : Set c where
field
shape⊸ : Shape C₁ → Shape C₂
- position⊸ : ∀ {s} → Position C₂ (shape⊸ s) ⇿ Position C₁ s
+ position⊸ : ∀ {s} → Position C₂ (shape⊸ s) ↔ Position C₁ s
morphism : C₁ ⇒ C₂
morphism = record
@@ -250,21 +231,29 @@ _∈_ : ∀ {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.
+-- Bag and set equality and related preorders. 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.
+
+open Related public
+ using (Kind; Symmetric-kind)
+ renaming ( implication to subset
+ ; reverse-implication to superset
+ ; equivalence to set
+ ; injection to subbag
+ ; reverse-injection to superbag
+ ; bijection to bag
+ )
-open Inv public
- using (Kind) renaming (inverse to bag; equivalent to set)
+[_]-Order : ∀ {ℓ} → Kind → Container ℓ → Set ℓ → Preorder ℓ ℓ ℓ
+[ k ]-Order C X = Related.InducedPreorder₂ k (_∈_ {C = C} {X = X})
-[_]-Equality : ∀ {ℓ} → Kind → Container ℓ → Set ℓ → Setoid ℓ ℓ
-[ k ]-Equality C X = Inv.InducedEquivalence₂ k (_∈_ {C = C} {X = X})
+[_]-Equality : ∀ {ℓ} → Symmetric-kind → Container ℓ → Set ℓ → Setoid ℓ ℓ
+[ k ]-Equality C X = Related.InducedEquivalence₂ k (_∈_ {C = C} {X = X})
-infix 4 _≈[_]_
+infix 4 _∼[_]_
-_≈[_]_ : ∀ {c} {C : Container c} {X : Set c} →
+_∼[_]_ : ∀ {c} {C : Container c} {X : Set c} →
⟦ C ⟧ X → Kind → ⟦ C ⟧ X → Set c
-xs ≈[ k ] ys = Setoid._≈_ ([ k ]-Equality _ _) xs ys
+_∼[_]_ {C = C} {X} xs k ys = Preorder._∼_ ([ k ]-Order C X) xs ys
diff --git a/src/Data/Container/Any.agda b/src/Data/Container/Any.agda
index ecde2a9..64a6943 100644
--- a/src/Data/Container/Any.agda
+++ b/src/Data/Container/Any.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties related to ◇
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Container.Any where
open import Algebra
@@ -14,24 +14,25 @@ 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
+open import Function.Inverse as Inv using (_↔_; module Inverse)
+open import Function.Related as Related using (Related)
+open import Function.Related.TypeIsomorphisms
import Relation.Binary.HeterogeneousEquality as H
+open import Relation.Binary.Product.Pointwise
open import Relation.Binary.PropositionalEquality as P
using (_≡_; _≗_; refl)
import Relation.Binary.Sigma.Pointwise as Σ
-open Inv.EquationalReasoning
+open Related.EquationalReasoning
private
module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
-- ◇ can be expressed using _∈_.
-⇿∈ : ∀ {c} (C : Container c) {X : Set c}
+↔∈ : ∀ {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
+ ◇ P xs ↔ (∃ λ x → x ∈ xs × P x)
+↔∈ _ {P = P} {xs} = record
{ to = P.→-to-⟶ to
; from = P.→-to-⟶ from
; inverse-of = record
@@ -49,16 +50,16 @@ private
to∘from : to ∘ from ≗ id
to∘from (.(proj₂ xs p) , (p , refl) , Px) = refl
--- ◇ is a congruence for bag and set equality.
+-- ◇ is a congruence for bag and set equality and related preorders.
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 Inv.id (xs₁≈xs₂ ⟨ ×⊎.*-cong ⟩ P₁⇿P₂ _) ⟩
- (∃ λ x → x ∈ xs₂ × P₂ x) ⇿⟨ sym (⇿∈ C) ⟩
+ (∀ x → Related k (P₁ x) (P₂ x)) → xs₁ ∼[ k ] xs₂ →
+ Related 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 Inv.id (xs₁≈xs₂ ×-cong P₁↔P₂ _) ⟩
+ (∃ λ x → x ∈ xs₂ × P₂ x) ↔⟨ sym (↔∈ C) ⟩
◇ P₂ xs₂ ∎
-- Nested occurrences of ◇ can sometimes be swapped.
@@ -67,27 +68,27 @@ 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)
+ ◈ 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 Inv.id (λ {x} → Inv.id ⟨ ×⊎.*-cong {ℓ = c} ⟩ ⇿∈ C₂ {P = P x}) ⟩
- (∃ λ x → x ∈ xs × ∃ λ y → y ∈ ys × P x y) ⇿⟨ Σ.cong Inv.id (λ {x} → ∃∃⇿∃∃ {A = x ∈ xs} (λ _ y → y ∈ ys × P x y)) ⟩
- (∃₂ λ x y → x ∈ xs × y ∈ ys × P x y) ⇿⟨ ∃∃⇿∃∃ (λ x y → x ∈ xs × y ∈ ys × P x y) ⟩
- (∃₂ λ y x → x ∈ xs × y ∈ ys × P x y) ⇿⟨ Σ.cong Inv.id (λ {y} → Σ.cong Inv.id (λ {x} →
- (x ∈ xs × y ∈ ys × P x y) ⇿⟨ sym $ ×⊎.*-assoc _ _ _ ⟩
- ((x ∈ xs × y ∈ ys) × P x y) ⇿⟨ ×⊎.*-comm _ _ ⟨ ×⊎.*-cong {ℓ = c} ⟩ Inv.id ⟩
- ((y ∈ ys × x ∈ xs) × P x y) ⇿⟨ ×⊎.*-assoc _ _ _ ⟩
+ ◇ (λ x → ◇ (P x) ys) xs ↔⟨ ↔∈ C₁ ⟩
+ (∃ λ x → x ∈ xs × ◇ (P x) ys) ↔⟨ Σ.cong Inv.id (λ {x} → Inv.id ⟨ ×⊎.*-cong {ℓ = c} ⟩ ↔∈ C₂ {P = P x}) ⟩
+ (∃ λ x → x ∈ xs × ∃ λ y → y ∈ ys × P x y) ↔⟨ Σ.cong Inv.id (λ {x} → ∃∃↔∃∃ {A = x ∈ xs} (λ _ y → y ∈ ys × P x y)) ⟩
+ (∃₂ λ x y → x ∈ xs × y ∈ ys × P x y) ↔⟨ ∃∃↔∃∃ (λ x y → x ∈ xs × y ∈ ys × P x y) ⟩
+ (∃₂ λ y x → x ∈ xs × y ∈ ys × P x y) ↔⟨ Σ.cong Inv.id (λ {y} → Σ.cong Inv.id (λ {x} →
+ (x ∈ xs × y ∈ ys × P x y) ↔⟨ sym $ ×⊎.*-assoc _ _ _ ⟩
+ ((x ∈ xs × y ∈ ys) × P x y) ↔⟨ ×⊎.*-comm _ _ ⟨ ×⊎.*-cong {ℓ = c} ⟩ Inv.id ⟩
+ ((y ∈ ys × x ∈ xs) × P x y) ↔⟨ ×⊎.*-assoc _ _ _ ⟩
(y ∈ ys × x ∈ xs × P x y) ∎)) ⟩
- (∃₂ λ y x → y ∈ ys × x ∈ xs × P x y) ⇿⟨ Σ.cong Inv.id (λ {y} → ∃∃⇿∃∃ {B = y ∈ ys} (λ x _ → x ∈ xs × P x y)) ⟩
- (∃ λ y → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ⇿⟨ Σ.cong Inv.id (λ {y} → Inv.id ⟨ ×⊎.*-cong {ℓ = c} ⟩ sym (⇿∈ C₁ {P = flip P y})) ⟩
- (∃ λ y → y ∈ ys × ◇ (flip P y) xs) ⇿⟨ sym (⇿∈ C₂) ⟩
+ (∃₂ λ y x → y ∈ ys × x ∈ xs × P x y) ↔⟨ Σ.cong Inv.id (λ {y} → ∃∃↔∃∃ {B = y ∈ ys} (λ x _ → x ∈ xs × P x y)) ⟩
+ (∃ λ y → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ↔⟨ Σ.cong Inv.id (λ {y} → Inv.id ⟨ ×⊎.*-cong {ℓ = c} ⟩ sym (↔∈ C₁ {P = flip P y})) ⟩
+ (∃ λ y → y ∈ ys × ◇ (flip P y) xs) ↔⟨ sym (↔∈ C₂) ⟩
◇ (λ y → ◇ (flip P y) xs) ys ∎
-- Nested occurrences of ◇ can sometimes be flattened.
flatten : ∀ {c} {C₁ C₂ : Container c} {X}
P (xss : ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)) →
- ◇ (◇ P) xss ⇿
+ ◇ (◇ P) xss ↔
◇ P (Inverse.from (Composition.correct C₁ C₂) ⟨$⟩ xss)
flatten {C₁ = C₁} {C₂} {X} P xss = record
{ to = P.→-to-⟶ t
@@ -108,10 +109,10 @@ flatten {C₁ = C₁} {C₂} {X} P xss = record
-- Sums commute with ◇ (for a fixed instance of a given container).
-◇⊎⇿⊎◇ : ∀ {c} {C : Container c} {X : Set c} {xs : ⟦ C ⟧ X}
+◇⊎↔⊎◇ : ∀ {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
+ ◇ (λ 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
@@ -133,11 +134,11 @@ flatten {C₁ = C₁} {C₂} {X} P xss = record
-- Products "commute" with ◇.
-×◇⇿◇◇× : ∀ {c} {C₁ C₂ : Container c}
+×◇↔◇◇× : ∀ {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
+ ◇ (λ 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
@@ -154,35 +155,35 @@ flatten {C₁ = C₁} {C₂} {X} P xss = record
-- map can be absorbed by the predicate.
-map⇿∘ : ∀ {c} (C : Container c) {X Y : Set c}
+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
+ ◇ 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}
+∈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 ⟩
+ 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 is a congruence for bag and set equality and related preorders.
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₂
+ 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₁ ↔⟨ map↔∘ C (_≡_ x) f₁ ⟩
+ ◇ (λ y → x ≡ f₁ y) xs₁ ∼⟨ cong {xs₁ = xs₁} {xs₂ = xs₂} (Related.↔⇒ ∘ 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 → (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))
@@ -197,7 +198,7 @@ map-cong {c = c} {C} {f₁ = f₁} {f₂} {xs₁} {xs₂} f₁≗f₂ xs₁≈xs
remove-linear :
∀ {c} {C₁ C₂ : Container c} {X} {xs : ⟦ C₁ ⟧ X}
(P : X → Set c) (m : C₁ ⊸ C₂) →
- ◇ P (⟪ m ⟫⊸ xs) ⇿ ◇ P xs
+ ◇ P (⟪ m ⟫⊸ xs) ↔ ◇ P xs
remove-linear {xs = xs} P m = record
{ to = P.→-to-⟶ t
; from = P.→-to-⟶ f
@@ -239,22 +240,22 @@ remove-linear {xs = xs} P m = record
linear-identity :
∀ {c} {C : Container c} {X} {xs : ⟦ C ⟧ X} (m : C ⊸ C) →
- ⟪ m ⟫⊸ xs ≈[ bag ] xs
+ ⟪ m ⟫⊸ xs ∼[ bag ] xs
linear-identity {xs = xs} m {x} =
- x ∈ ⟪ m ⟫⊸ xs ⇿⟨ remove-linear (_≡_ x) m ⟩
+ 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}
+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 (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
index eadfef0..f3c7bf3 100644
--- a/src/Data/Container/Combinator.agda
+++ b/src/Data/Container/Combinator.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Container combinators
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Container.Combinator where
open import Data.Container
@@ -12,7 +12,7 @@ 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 Function.Inverse using (_↔_)
open import Level
open import Relation.Binary.PropositionalEquality as P
using (_≗_; refl)
@@ -23,19 +23,19 @@ open import Relation.Binary.PropositionalEquality as P
-- Identity.
id : ∀ {c} → Container c
-id = Lift ⊤ ◃ F.const (Lift ⊤)
+id = Lift ⊤ ▷ F.const (Lift ⊤)
-- Constant.
const : ∀ {c} → Set c → Container c
-const X = X ◃ F.const (Lift ⊥)
+const X = X ▷ F.const (Lift ⊥)
-- Composition.
infixr 9 _∘_
_∘_ : ∀ {c} → Container c → Container c → Container c
-C₁ ∘ C₂ = ⟦ C₁ ⟧ (Shape C₂) ◃ ◇ (Position C₂)
+C₁ ∘ C₂ = ⟦ C₁ ⟧ (Shape C₂) ▷ ◇ (Position C₂)
-- Product. (Note that, up to isomorphism, this is a special case of
-- indexed product.)
@@ -44,13 +44,13 @@ infixr 2 _×_
_×_ : ∀ {c} → Container c → Container c → Container c
C₁ × C₂ =
- (Shape C₁ ⟨×⟩ Shape 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)
+Π 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.)
@@ -58,12 +58,12 @@ C₁ × C₂ =
infixr 1 _⊎_
_⊎_ : ∀ {c} → Container c → Container c → Container c
-C₁ ⊎ C₂ = (Shape C₁ ⟨⊎⟩ Shape C₂) ◃ [ Position C₁ , Position 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)
+Σ C = (∃ λ i → Shape (C i)) ▷ λ s → Position (C (proj₁ s)) (proj₂ s)
-- Constant exponentiation. (Note that this is a special case of
-- indexed product.)
@@ -82,7 +82,7 @@ const[ X ]⟶ C = Π {I = X} (F.const C)
module Identity where
- correct : ∀ {c} {X : Set c} → ⟦ id ⟧ X ⇿ F.id X
+ 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)
@@ -94,7 +94,7 @@ module Identity where
module Constant (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
- correct : ∀ {ℓ} (X : Set ℓ) {Y} → ⟦ const X ⟧ Y ⇿ F.const X Y
+ correct : ∀ {ℓ} (X : Set ℓ) {Y} → ⟦ const X ⟧ Y ↔ F.const X Y
correct X {Y} = record
{ to = P.→-to-⟶ to
; from = P.→-to-⟶ from
@@ -114,7 +114,7 @@ module Constant (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
module Composition where
correct : ∀ {c} (C₁ C₂ : Container c) {X : Set c} →
- ⟦ C₁ ∘ C₂ ⟧ X ⇿ (⟦ C₁ ⟧ ⟨∘⟩ ⟦ C₂ ⟧) X
+ ⟦ C₁ ∘ C₂ ⟧ X ↔ (⟦ C₁ ⟧ ⟨∘⟩ ⟦ C₂ ⟧) X
correct C₁ C₂ {X} = record
{ to = P.→-to-⟶ to
; from = P.→-to-⟶ from
@@ -133,7 +133,7 @@ module Composition where
module Product (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
correct : ∀ {c} (C₁ C₂ : Container c) {X : Set c} →
- ⟦ C₁ × C₂ ⟧ X ⇿ (⟦ C₁ ⟧ X ⟨×⟩ ⟦ C₂ ⟧ X)
+ ⟦ C₁ × C₂ ⟧ X ↔ (⟦ C₁ ⟧ X ⟨×⟩ ⟦ C₂ ⟧ X)
correct {c} C₁ C₂ {X} = record
{ to = P.→-to-⟶ to
; from = P.→-to-⟶ from
@@ -156,7 +156,7 @@ module Product (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
module IndexedProduct where
correct : ∀ {c I} (C : I → Container c) {X : Set c} →
- ⟦ Π C ⟧ X ⇿ (∀ i → ⟦ C i ⟧ X)
+ ⟦ Π C ⟧ X ↔ (∀ i → ⟦ C i ⟧ X)
correct {I = I} C {X} = record
{ to = P.→-to-⟶ to
; from = P.→-to-⟶ from
@@ -175,7 +175,7 @@ module IndexedProduct where
module Sum where
correct : ∀ {c} (C₁ C₂ : Container c) {X : Set c} →
- ⟦ C₁ ⊎ C₂ ⟧ X ⇿ (⟦ C₁ ⟧ X ⟨⊎⟩ ⟦ C₂ ⟧ X)
+ ⟦ C₁ ⊎ C₂ ⟧ X ↔ (⟦ C₁ ⟧ X ⟨⊎⟩ ⟦ C₂ ⟧ X)
correct C₁ C₂ {X} = record
{ to = P.→-to-⟶ to
; from = P.→-to-⟶ from
@@ -199,7 +199,7 @@ module Sum where
module IndexedSum where
correct : ∀ {c I} (C : I → Container c) {X : Set c} →
- ⟦ Σ C ⟧ X ⇿ (∃ λ i → ⟦ C i ⟧ X)
+ ⟦ Σ C ⟧ X ↔ (∃ λ i → ⟦ C i ⟧ X)
correct {I = I} C {X} = record
{ to = P.→-to-⟶ to
; from = P.→-to-⟶ from
@@ -218,5 +218,5 @@ module IndexedSum where
module ConstantExponentiation where
correct : ∀ {c X} (C : Container c) {Y : Set c} →
- ⟦ const[ X ]⟶ C ⟧ Y ⇿ (X → ⟦ C ⟧ Y)
+ ⟦ const[ X ]⟶ C ⟧ Y ↔ (X → ⟦ C ⟧ Y)
correct C = IndexedProduct.correct (F.const C)
diff --git a/src/Data/Context.agda b/src/Data/Context.agda
deleted file mode 100644
index 3e924cd..0000000
--- a/src/Data/Context.agda
+++ /dev/null
@@ -1,188 +0,0 @@
-------------------------------------------------------------------------
--- Contexts, variables, substitutions, etc.
-------------------------------------------------------------------------
-
-{-# OPTIONS --universe-polymorphism #-}
-
--- Based on Conor McBride's "Outrageous but Meaningful Coincidences:
--- Dependent type-safe syntax and evaluation".
-
--- The contexts and variables are parametrised by a universe.
-
-open import Universe
-
-module Data.Context {u e} (Uni : Universe u e) where
-
-open Universe.Universe Uni
-
-open import Data.Product as Prod
-open import Data.Unit
-open import Function
-open import Level
-
-------------------------------------------------------------------------
--- Contexts and "types"
-
-mutual
-
- -- Contexts.
-
- infixl 5 _▻_
-
- data Ctxt : Set (u ⊔ e) where
- ε : Ctxt
- _▻_ : (Γ : Ctxt) (σ : Type Γ) → Ctxt
-
- -- Semantic types: maps from environments to universe codes.
-
- Type : Ctxt → Set (u ⊔ e)
- Type Γ = Env Γ → U
-
- -- Interpretation of contexts: environments.
-
- Env : Ctxt → Set e
- Env ε = Lift ⊤
- Env (Γ ▻ σ) = Σ (Env Γ) λ γ → El (σ γ)
-
--- Type signature for variables, terms, etc.
-
-Term-like : (ℓ : Level) → Set _
-Term-like ℓ = (Γ : Ctxt) → Type Γ → Set ℓ
-
--- Semantic values: maps from environments to universe values.
-
-Value : Term-like _
-Value Γ σ = (γ : Env Γ) → El (σ γ)
-
-------------------------------------------------------------------------
--- "Substitutions"
-
--- Semantic substitutions or context morphisms: maps from environments
--- to environments.
-
-infixr 4 _⇨_
-
-_⇨_ : Ctxt → Ctxt → Set _
-Γ ⇨ Δ = Env Δ → Env Γ
-
--- Maps between types.
-
-infixr 4 _⇨₁_
-
-_⇨₁_ : ∀ {Γ} → Type Γ → Type Γ → Set _
-σ ⇨₁ τ = ∀ {γ} → El (τ γ) → El (σ γ)
-
--- Application of substitutions to types.
---
--- Note that this operation is composition of functions. I have chosen
--- not to give a separate name to the identity substitution, which is
--- simply the identity function, nor to reverse composition of
--- substitutions, which is composition of functions.
-
-infixl 8 _/_
-
-_/_ : ∀ {Γ Δ} → Type Γ → Γ ⇨ Δ → Type Δ
-σ / ρ = σ ∘ ρ
-
--- Application of substitutions to values.
-
-infixl 8 _/v_
-
-_/v_ : ∀ {Γ Δ σ} → Value Γ σ → (ρ : Γ ⇨ Δ) → Value Δ (σ / ρ)
-v /v ρ = v ∘ ρ
-
--- Weakening.
-
-wk : ∀ {Γ σ} → Γ ⇨ Γ ▻ σ
-wk = proj₁
-
--- Extends a substitution with another value.
-
-infixl 5 _►_
-
-_►_ : ∀ {Γ Δ σ} (ρ : Γ ⇨ Δ) → Value Δ (σ / ρ) → Γ ▻ σ ⇨ Δ
-_►_ = <_,_>
-
--- A substitution which only replaces the first "variable".
-
-sub : ∀ {Γ σ} → Value Γ σ → Γ ▻ σ ⇨ Γ
-sub v = id ► v
-
--- One can construct a substitution between two non-empty contexts by
--- supplying two functions, one which handles the last element and one
--- which handles the rest.
-
-infixl 10 _↑_
-
-_↑_ : ∀ {Γ Δ σ τ} (ρ : Γ ⇨ Δ) → σ / ρ ⇨₁ τ → Γ ▻ σ ⇨ Δ ▻ τ
-_↑_ = Prod.map
-
--- Lifting.
-
-infix 10 _↑
-
-_↑ : ∀ {Γ Δ σ} (ρ : Γ ⇨ Δ) → Γ ▻ σ ⇨ Δ ▻ σ / ρ
-ρ ↑ = ρ ↑ id
-
-------------------------------------------------------------------------
--- Variables
-
--- Variables (de Bruijn indices).
-
-infix 4 _∋_
-
-data _∋_ : Term-like (u ⊔ e) where
- zero : ∀ {Γ σ} → Γ ▻ σ ∋ σ / wk
- suc : ∀ {Γ σ τ} (x : Γ ∋ σ) → Γ ▻ τ ∋ σ / wk
-
--- Interpretation of variables: a lookup function.
-
-lookup : ∀ {Γ σ} → Γ ∋ σ → Value Γ σ
-lookup zero (γ , v) = v
-lookup (suc x) (γ , v) = lookup x γ
-
--- Application of substitutions to variables.
-
-infixl 8 _/∋_
-
-_/∋_ : ∀ {Γ Δ σ} → Γ ∋ σ → (ρ : Γ ⇨ Δ) → Value Δ (σ / ρ)
-x /∋ ρ = lookup x /v ρ
-
-------------------------------------------------------------------------
--- Context extensions
-
-mutual
-
- -- Context extensions.
-
- infixl 5 _▻_
-
- data Ctxt⁺ (Γ : Ctxt) : Set (u ⊔ e) where
- ε : Ctxt⁺ Γ
- _▻_ : (Γ⁺ : Ctxt⁺ Γ) (σ : Type (Γ ++ Γ⁺)) → Ctxt⁺ Γ
-
- -- Appends a context extension to a context.
-
- infixl 5 _++_
-
- _++_ : (Γ : Ctxt) → Ctxt⁺ Γ → Ctxt
- Γ ++ ε = Γ
- Γ ++ (Γ⁺ ▻ σ) = (Γ ++ Γ⁺) ▻ σ
-
-mutual
-
- -- Application of substitutions to context extensions.
-
- infixl 8 _/⁺_
-
- _/⁺_ : ∀ {Γ Δ} → Ctxt⁺ Γ → Γ ⇨ Δ → Ctxt⁺ Δ
- ε /⁺ ρ = ε
- (Γ⁺ ▻ σ) /⁺ ρ = Γ⁺ /⁺ ρ ▻ σ / ρ ↑⋆ Γ⁺
-
- -- N-ary lifting of substitutions.
-
- infixl 10 _↑⋆_
-
- _↑⋆_ : ∀ {Γ Δ} (ρ : Γ ⇨ Δ) → ∀ Γ⁺ → Γ ++ Γ⁺ ⇨ Δ ++ Γ⁺ /⁺ ρ
- ρ ↑⋆ ε = ρ
- ρ ↑⋆ (Γ⁺ ▻ σ) = (ρ ↑⋆ Γ⁺) ↑
diff --git a/src/Data/Covec.agda b/src/Data/Covec.agda
index dc95a8f..aa7e176 100644
--- a/src/Data/Covec.agda
+++ b/src/Data/Covec.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Coinductive vectors
------------------------------------------------------------------------
diff --git a/src/Data/DifferenceList.agda b/src/Data/DifferenceList.agda
index 41692b4..7909ea0 100644
--- a/src/Data/DifferenceList.agda
+++ b/src/Data/DifferenceList.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Lists with fast append
------------------------------------------------------------------------
@@ -10,48 +12,50 @@ open import Data.Nat
infixr 5 _∷_ _++_
-DiffList : Set → Set
-DiffList a = List a → List a
+DiffList : ∀ {ℓ} → Set ℓ → Set ℓ
+DiffList A = List A → List A
-lift : ∀ {a} → (List a → List a) → (DiffList a → DiffList a)
+lift : ∀ {a} {A : Set a} →
+ (List A → List A) → (DiffList A → DiffList A)
lift f xs = λ k → f (xs k)
-[] : ∀ {a} → DiffList a
+[] : ∀ {a} {A : Set a} → DiffList A
[] = λ k → k
-_∷_ : ∀ {a} → a → DiffList a → DiffList a
+_∷_ : ∀ {a} {A : Set a} → A → DiffList A → DiffList A
_∷_ x = lift (L._∷_ x)
-[_] : ∀ {a} → a → DiffList a
+[_] : ∀ {a} {A : Set a} → A → DiffList A
[ x ] = x ∷ []
-_++_ : ∀ {a} → DiffList a → DiffList a → DiffList a
+_++_ : ∀ {a} {A : Set a} → DiffList A → DiffList A → DiffList A
xs ++ ys = λ k → xs (ys k)
-toList : ∀ {a} → DiffList a → List a
+toList : ∀ {a} {A : Set a} → DiffList A → List A
toList xs = xs L.[]
-- fromList xs is linear in the length of xs.
-fromList : ∀ {a} → List a → DiffList a
+fromList : ∀ {a} {A : Set a} → List A → DiffList A
fromList xs = λ k → xs ⟨ L._++_ ⟩ k
-- It is OK to use L._++_ here, since map is linear in the length of
-- the list anyway.
-map : ∀ {a b} → (a → b) → DiffList a → DiffList b
+map : ∀ {a b} {A : Set a} {B : Set b} →
+ (A → B) → DiffList A → DiffList B
map f xs = λ k → L.map f (toList xs) ⟨ L._++_ ⟩ k
-- concat is linear in the length of the outer list.
-concat : ∀ {a} → DiffList (DiffList a) → DiffList a
+concat : ∀ {a} {A : Set a} → DiffList (DiffList A) → DiffList A
concat xs = concat' (toList xs)
where
- concat' : ∀ {a} → List (DiffList a) → DiffList a
+ concat' : ∀ {a} {A : Set a} → List (DiffList A) → DiffList A
concat' = L.foldr _++_ []
-take : ∀ {a} → ℕ → DiffList a → DiffList a
+take : ∀ {a} {A : Set a} → ℕ → DiffList A → DiffList A
take n = lift (L.take n)
-drop : ∀ {a} → ℕ → DiffList a → DiffList a
+drop : ∀ {a} {A : Set a} → ℕ → DiffList A → DiffList A
drop n = lift (L.drop n)
diff --git a/src/Data/DifferenceNat.agda b/src/Data/DifferenceNat.agda
index 5e8f259..e60e38b 100644
--- a/src/Data/DifferenceNat.agda
+++ b/src/Data/DifferenceNat.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Natural numbers with fast addition (for use together with
-- DifferenceVec)
------------------------------------------------------------------------
diff --git a/src/Data/DifferenceVec.agda b/src/Data/DifferenceVec.agda
index 2718e54..7522cd0 100644
--- a/src/Data/DifferenceVec.agda
+++ b/src/Data/DifferenceVec.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Vectors with fast append
------------------------------------------------------------------------
@@ -11,39 +13,42 @@ import Data.Nat as N
infixr 5 _∷_ _++_
-DiffVec : Set → Diffℕ → Set
+DiffVec : ∀ {ℓ} → Set ℓ → Diffℕ → Set ℓ
DiffVec A m = ∀ {n} → Vec A n → Vec A (m n)
-[] : ∀ {A} → DiffVec A 0#
+[] : ∀ {a} {A : Set a} → DiffVec A 0#
[] = λ k → k
-_∷_ : ∀ {A n} → A → DiffVec A n → DiffVec A (suc n)
+_∷_ : ∀ {a} {A : Set a} {n} → A → DiffVec A n → DiffVec A (suc n)
x ∷ xs = λ k → V._∷_ x (xs k)
-[_] : ∀ {A} → A → DiffVec A 1#
+[_] : ∀ {a} {A : Set a} → A → DiffVec A 1#
[ x ] = x ∷ []
-_++_ : ∀ {A m n} → DiffVec A m → DiffVec A n → DiffVec A (m + n)
+_++_ : ∀ {a} {A : Set a} {m n} →
+ DiffVec A m → DiffVec A n → DiffVec A (m + n)
xs ++ ys = λ k → xs (ys k)
-toVec : ∀ {A n} → DiffVec A n → Vec A (toℕ n)
+toVec : ∀ {a} {A : Set a} {n} → DiffVec A n → Vec A (toℕ n)
toVec xs = xs V.[]
-- fromVec xs is linear in the length of xs.
-fromVec : ∀ {A n} → Vec A n → DiffVec A (fromℕ n)
+fromVec : ∀ {a} {A : Set a} {n} → Vec A n → DiffVec A (fromℕ n)
fromVec xs = λ k → xs ⟨ V._++_ ⟩ k
-head : ∀ {A n} → DiffVec A (suc n) → A
+head : ∀ {a} {A : Set a} {n} → DiffVec A (suc n) → A
head xs = V.head (toVec xs)
-tail : ∀ {A n} → DiffVec A (suc n) → DiffVec A n
+tail : ∀ {a} {A : Set a} {n} → DiffVec A (suc n) → DiffVec A n
tail xs = λ k → V.tail (xs k)
-take : ∀ {a} m {n} → DiffVec a (fromℕ m + n) → DiffVec a (fromℕ m)
+take : ∀ {a} {A : Set a} m {n} →
+ DiffVec A (fromℕ m + n) → DiffVec A (fromℕ m)
take N.zero xs = []
take (N.suc m) xs = head xs ∷ take m (tail xs)
-drop : ∀ {a} m {n} → DiffVec a (fromℕ m + n) → DiffVec a n
+drop : ∀ {a} {A : Set a} m {n} →
+ DiffVec A (fromℕ m + n) → DiffVec A n
drop N.zero xs = xs
drop (N.suc m) xs = drop m (tail xs)
diff --git a/src/Data/Digit.agda b/src/Data/Digit.agda
index 9a2576c..1b7adb8 100644
--- a/src/Data/Digit.agda
+++ b/src/Data/Digit.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Digits and digit expansions
------------------------------------------------------------------------
diff --git a/src/Data/Empty.agda b/src/Data/Empty.agda
index 83eb79e..b32d11b 100644
--- a/src/Data/Empty.agda
+++ b/src/Data/Empty.agda
@@ -1,14 +1,17 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Empty type
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Empty where
open import Level
data ⊥ : Set where
+{-# IMPORT Data.FFI #-}
+{-# COMPILED_DATA ⊥ Data.FFI.AgdaEmpty #-}
+
⊥-elim : ∀ {w} {Whatever : Set w} → ⊥ → Whatever
⊥-elim ()
diff --git a/src/Data/Fin.agda b/src/Data/Fin.agda
index fbdd2f8..7bfa320 100644
--- a/src/Data/Fin.agda
+++ b/src/Data/Fin.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Finite sets
------------------------------------------------------------------------
@@ -13,7 +15,7 @@ open import Data.Nat as Nat
renaming ( _+_ to _N+_; _∸_ to _N∸_
; _≤_ to _N≤_; _≥_ to _N≥_; _<_ to _N<_; _≤?_ to _N≤?_)
open import Function
-open import Level hiding (lift)
+import Level
open import Relation.Nullary.Decidable
open import Relation.Binary
@@ -162,10 +164,10 @@ pred (suc i) = inject₁ i
infix 4 _≤_ _<_
-_≤_ : ∀ {n} → Rel (Fin n) zero
+_≤_ : ∀ {n} → Rel (Fin n) Level.zero
_≤_ = _N≤_ on toℕ
-_<_ : ∀ {n} → Rel (Fin n) zero
+_<_ : ∀ {n} → Rel (Fin n) Level.zero
_<_ = _N<_ on toℕ
data _≺_ : ℕ → ℕ → Set where
diff --git a/src/Data/Fin/Dec.agda b/src/Data/Fin/Dec.agda
index d54ca0c..0bee913 100644
--- a/src/Data/Fin/Dec.agda
+++ b/src/Data/Fin/Dec.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Decision procedures for finite sets and subsets of finite sets
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Fin.Dec where
open import Function
@@ -11,7 +11,7 @@ import Data.Bool as Bool
open import Data.Nat hiding (_<_)
open import Data.Vec hiding (_∈_)
open import Data.Vec.Equality as VecEq
- using () renaming (module HeterogeneousEquality to HetVecEq)
+ using () renaming (module PropositionalEquality to PropVecEq)
open import Data.Fin
open import Data.Fin.Subset
open import Data.Fin.Subset.Props
@@ -19,7 +19,6 @@ open import Data.Product as Prod
open import Data.Empty
open import Function
import Function.Equivalence as Eq
-open import Level hiding (Lift)
open import Relation.Binary
import Relation.Binary.HeterogeneousEquality as H
open import Relation.Nullary
@@ -170,6 +169,5 @@ infix 4 _⊆?_
_⊆?_ : ∀ {n} → Decidable (_⊆_ {n = n})
p₁ ⊆? p₂ =
Dec.map (Eq.sym NaturalPoset.orders-equivalent) $
- Dec.map′ H.≅-to-≡ H.≡-to-≅ $
- Dec.map′ HetVecEq.to-≅ HetVecEq.from-≅ $
+ Dec.map′ PropVecEq.to-≡ PropVecEq.from-≡ $
VecEq.DecidableEquality._≟_ Bool.decSetoid p₁ (p₁ ∩ p₂)
diff --git a/src/Data/Fin/Props.agda b/src/Data/Fin/Props.agda
index 7422d90..6155a9a 100644
--- a/src/Data/Fin/Props.agda
+++ b/src/Data/Fin/Props.agda
@@ -1,10 +1,10 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties related to Fin, and operations making use of these
-- properties (or other properties not available in Data.Fin)
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Fin.Props where
open import Data.Fin
@@ -177,7 +177,7 @@ private
-- an equivalence relation with two equivalence classes, one with
-- all inhabited sets and the other with all uninhabited sets).
- sequence⁻¹ : ∀ {F A} {P : A → Set} → RawFunctor F →
+ sequence⁻¹ : ∀ {F}{A} {P : A → Set} → RawFunctor F →
F (∀ i → P i) → ∀ i → F (P i)
sequence⁻¹ RF F∀iPi i = (λ f → f i) <$> F∀iPi
where open RawFunctor RF
diff --git a/src/Data/Fin/Subset.agda b/src/Data/Fin/Subset.agda
index 484b275..42a61ea 100644
--- a/src/Data/Fin/Subset.agda
+++ b/src/Data/Fin/Subset.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Subsets of finite sets
------------------------------------------------------------------------
diff --git a/src/Data/Fin/Subset/Props.agda b/src/Data/Fin/Subset/Props.agda
index f6dc52b..698f3e2 100644
--- a/src/Data/Fin/Subset/Props.agda
+++ b/src/Data/Fin/Subset/Props.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some properties about subsets
------------------------------------------------------------------------
@@ -16,7 +18,7 @@ open import Data.Vec hiding (_∈_)
open import Function
open import Function.Equality using (_⟨$⟩_)
open import Function.Equivalence
- using (_⇔_; equivalent; module Equivalent)
+ using (_⇔_; equivalence; module Equivalence)
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as P using (_≡_)
@@ -65,7 +67,7 @@ Empty-unique {p = inside ∷ .⊥} ¬∃∈ | P.refl =
x∈⁅y⁆⇔x≡y : ∀ {n} {x y : Fin n} → x ∈ ⁅ y ⁆ ⇔ x ≡ y
x∈⁅y⁆⇔x≡y {x = x} {y} =
- equivalent (to y) (λ x≡y → P.subst (λ y → x ∈ ⁅ y ⁆) x≡y (x∈⁅x⁆ x))
+ equivalence (to y) (λ x≡y → P.subst (λ y → x ∈ ⁅ y ⁆) x≡y (x∈⁅x⁆ x))
where
to : ∀ {n x} (y : Fin n) → x ∈ ⁅ y ⁆ → x ≡ y
@@ -81,8 +83,8 @@ x∈⁅y⁆⇔x≡y {x = x} {y} =
------------------------------------------------------------------------
-- A property involving _∪_
-∪⇿⊎ : ∀ {n} {p₁ p₂ : Subset n} {x} → x ∈ p₁ ∪ p₂ ⇔ (x ∈ p₁ ⊎ x ∈ p₂)
-∪⇿⊎ = equivalent (to _ _) from
+∪⇔⊎ : ∀ {n} {p₁ p₂ : Subset n} {x} → x ∈ p₁ ∪ p₂ ⇔ (x ∈ p₁ ⊎ x ∈ p₂)
+∪⇔⊎ = equivalence (to _ _) from
where
to : ∀ {n} (p₁ p₂ : Subset n) {x} → x ∈ p₁ ∪ p₂ → x ∈ p₁ ⊎ x ∈ p₂
to [] [] ()
@@ -118,7 +120,7 @@ module NaturalPoset where
-- _⊆_ is equivalent to the natural lattice order.
orders-equivalent : ∀ {n} {p₁ p₂ : Subset n} → p₁ ⊆ p₂ ⇔ p₁ ≤ p₂
- orders-equivalent = equivalent (to _ _) (from _ _)
+ orders-equivalent = equivalence (to _ _) (from _ _)
where
to : ∀ {n} (p₁ p₂ : Subset n) → p₁ ⊆ p₂ → p₁ ≤ p₂
to [] [] p₁⊆p₂ = P.refl
@@ -151,4 +153,4 @@ poset n = record
where
open NaturalPoset
open module E {p₁ p₂} =
- Equivalent (orders-equivalent {n = n} {p₁ = p₁} {p₂ = p₂})
+ Equivalence (orders-equivalent {n = n} {p₁ = p₁} {p₂ = p₂})
diff --git a/src/Data/Fin/Substitution.agda b/src/Data/Fin/Substitution.agda
index be0ed71..20a9c4a 100644
--- a/src/Data/Fin/Substitution.agda
+++ b/src/Data/Fin/Substitution.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Substitutions
------------------------------------------------------------------------
diff --git a/src/Data/Fin/Substitution/Example.agda b/src/Data/Fin/Substitution/Example.agda
index 7be9cb3..737b9ad 100644
--- a/src/Data/Fin/Substitution/Example.agda
+++ b/src/Data/Fin/Substitution/Example.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- An example of how Data.Fin.Substitution can be used: a definition
-- of substitution for the untyped λ-calculus, along with some lemmas
------------------------------------------------------------------------
diff --git a/src/Data/Fin/Substitution/Lemmas.agda b/src/Data/Fin/Substitution/Lemmas.agda
index 3f5baa0..225dbb4 100644
--- a/src/Data/Fin/Substitution/Lemmas.agda
+++ b/src/Data/Fin/Substitution/Lemmas.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Substitution lemmas
------------------------------------------------------------------------
diff --git a/src/Data/Fin/Substitution/List.agda b/src/Data/Fin/Substitution/List.agda
index 0afa5c4..b8f153f 100644
--- a/src/Data/Fin/Substitution/List.agda
+++ b/src/Data/Fin/Substitution/List.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Application of substitutions to lists, along with various lemmas
------------------------------------------------------------------------
diff --git a/src/Data/Graph/Acyclic.agda b/src/Data/Graph/Acyclic.agda
index 68e0a70..0bdd5fd 100644
--- a/src/Data/Graph/Acyclic.agda
+++ b/src/Data/Graph/Acyclic.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Directed acyclic multigraphs
------------------------------------------------------------------------
@@ -15,14 +17,14 @@ 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 Function
open import Data.Empty
open import Data.Unit using (⊤; tt)
open import Data.Vec as Vec using (Vec; []; _∷_)
open import Data.List as List using (List; []; _∷_)
-open import Relation.Nullary
-open import Relation.Binary.PropositionalEquality
+open import Function
open import Induction.Nat
+open import Relation.Nullary
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
------------------------------------------------------------------------
-- A lemma
@@ -172,7 +174,7 @@ private
test-nodes : nodes example ≡ (# 0 , 0) ∷ (# 1 , 1) ∷ (# 2 , 2) ∷
(# 3 , 3) ∷ (# 4 , 4) ∷ []
- test-nodes = refl
+ test-nodes = P.refl
-- Topological sort. Gives a vector where earlier nodes are never
-- successors of later nodes.
@@ -196,7 +198,7 @@ private
test-edges : edges example ≡ (# 1 , 10 , # 1) ∷ (# 1 , 11 , # 1) ∷
(# 2 , 12 , # 0) ∷ []
- test-edges = refl
+ test-edges = P.refl
-- The successors of a given node i (edge label × node number relative
-- to i).
@@ -208,7 +210,7 @@ sucs g i = successors $ head $ g [ i ]
private
test-sucs : sucs example (# 1) ≡ (10 , # 1) ∷ (11 , # 1) ∷ []
- test-sucs = refl
+ test-sucs = P.refl
-- The predecessors of a given node i (node number relative to i ×
-- edge label).
@@ -221,14 +223,14 @@ preds (c & g) (suc i) =
where
p : ∀ {E : Set} {n} (i : Fin n) → E × Fin n → Maybe (Fin′ (suc i) × E)
p i (e , j) with i ≟ j
- p i (e , .i) | yes refl = just (zero , e)
- p i (e , j) | no _ = nothing
+ p i (e , .i) | yes P.refl = just (zero , e)
+ p i (e , j) | no _ = nothing
private
test-preds : preds example (# 3) ≡
(# 1 , 10) ∷ (# 1 , 11) ∷ (# 2 , 12) ∷ []
- test-preds = refl
+ test-preds = P.refl
------------------------------------------------------------------------
-- Operations
@@ -255,7 +257,7 @@ private
context (# 3 , 3) [] &
context (# 4 , 4) [] &
∅)
- test-number = refl
+ test-number = P.refl
-- Reverses all the edges in the graph.
@@ -272,7 +274,7 @@ reverse {N} {E} g =
private
test-reverse : reverse (reverse example) ≡ example
- test-reverse = refl
+ test-reverse = P.refl
------------------------------------------------------------------------
-- Views
@@ -280,7 +282,7 @@ private
-- Expands the subgraph induced by a given node into a tree (thus
-- losing all sharing).
-data Tree N E : Set where
+data Tree (N E : Set) : Set where
node : (label : N) (successors : List (E × Tree N E)) → Tree N E
toTree : ∀ {N E n} → Graph N E n → Fin n → Tree N E
@@ -310,4 +312,4 @@ private
node 3 [] ∷
node 4 [] ∷
[]
- test-toForest = refl
+ test-toForest = P.refl
diff --git a/src/Data/Integer.agda b/src/Data/Integer.agda
index daaade4..f33484f 100644
--- a/src/Data/Integer.agda
+++ b/src/Data/Integer.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Integers
------------------------------------------------------------------------
diff --git a/src/Data/Integer/Divisibility.agda b/src/Data/Integer/Divisibility.agda
index 7f033dc..bb3b23c 100644
--- a/src/Data/Integer/Divisibility.agda
+++ b/src/Data/Integer/Divisibility.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Divisibility and coprimality
------------------------------------------------------------------------
diff --git a/src/Data/Integer/Properties.agda b/src/Data/Integer/Properties.agda
index af023eb..076dad1 100644
--- a/src/Data/Integer/Properties.agda
+++ b/src/Data/Integer/Properties.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some properties about integers
------------------------------------------------------------------------
diff --git a/src/Data/List.agda b/src/Data/List.agda
index c76b482..fb6a100 100644
--- a/src/Data/List.agda
+++ b/src/Data/List.agda
@@ -1,7 +1,8 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Lists
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
module Data.List where
@@ -28,6 +29,9 @@ data List {a} (A : Set a) : Set a where
{-# BUILTIN NIL [] #-}
{-# BUILTIN CONS _∷_ #-}
+{-# IMPORT Data.FFI #-}
+{-# COMPILED_DATA List Data.FFI.AgdaList [] (:) #-}
+
------------------------------------------------------------------------
-- Some operations
diff --git a/src/Data/List/All.agda b/src/Data/List/All.agda
index a48c608..bc4debf 100644
--- a/src/Data/List/All.agda
+++ b/src/Data/List/All.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Lists where all elements satisfy a given property
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.List.All where
open import Data.List as List hiding (map; all)
diff --git a/src/Data/List/All/Properties.agda b/src/Data/List/All/Properties.agda
index 2ed7a24..6837a46 100644
--- a/src/Data/List/All/Properties.agda
+++ b/src/Data/List/All/Properties.agda
@@ -1,16 +1,16 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties relating All to various list functions
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.List.All.Properties where
open import Data.Bool
open import Data.Bool.Properties
open import Function
open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence using (module Equivalent)
+open import Function.Equivalence using (module Equivalence)
open import Data.List as List
import Data.List.Any as Any
open Any.Membership-≡
@@ -45,12 +45,12 @@ gmap g = All-map ∘ All.map g
All-all : ∀ {a} {A : Set a} (p : A → Bool) {xs} →
All (T ∘ p) xs → T (all p xs)
All-all p [] = _
-All-all p (px ∷ pxs) = Equivalent.from T-∧ ⟨$⟩ (px , All-all p pxs)
+All-all p (px ∷ pxs) = Equivalence.from T-∧ ⟨$⟩ (px , All-all p pxs)
all-All : ∀ {a} {A : Set a} (p : A → Bool) xs →
T (all p xs) → All (T ∘ p) xs
all-All p [] _ = []
-all-All p (x ∷ xs) px∷xs with Equivalent.to (T-∧ {p x}) ⟨$⟩ px∷xs
+all-All p (x ∷ xs) px∷xs with Equivalence.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 da577fb..7a19b17 100644
--- a/src/Data/List/Any.agda
+++ b/src/Data/List/Any.agda
@@ -1,21 +1,20 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Lists where at least one element satisfies a given property
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.List.Any where
open import Data.Empty
open import Data.Fin
open import Function
open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence as Equiv using (module Equivalent)
-open import Function.Inverse as Inv
- using (module Inverse)
+open import Function.Equivalence as Equiv using (module Equivalence)
+open import Function.Related as Related hiding (_∼[_]_)
open import Data.List as List using (List; []; _∷_)
open import Data.Product as Prod using (∃; _×_; _,_)
-open import Level
+open import Level using (Level; _⊔_)
open import Relation.Nullary
import Relation.Nullary.Decidable as Dec
open import Relation.Unary using () renaming (_⊆_ to _⋐_)
@@ -165,24 +164,34 @@ module Membership-≡ where
⊆-preorder A = Ind.InducedPreorder₂ (PropEq.setoid (List A)) _∈_
(PropEq.subst (_∈_ _))
- -- Set and bag equality.
+ -- Set and bag equality and related preorders.
+
+ open Related public
+ using (Kind; Symmetric-kind)
+ renaming ( implication to subset
+ ; reverse-implication to superset
+ ; equivalence to set
+ ; injection to subbag
+ ; reverse-injection to superbag
+ ; bijection to bag
+ )
- open Inv public
- using (Kind) renaming (equivalent to set; inverse to bag)
+ [_]-Order : Kind → ∀ {a} → Set a → Preorder _ _ _
+ [ k ]-Order A = Related.InducedPreorder₂ k (_∈_ {A = A})
- [_]-Equality : Kind → ∀ {a} → Set a → Setoid _ _
- [ k ]-Equality A = Inv.InducedEquivalence₂ k (_∈_ {A = A})
+ [_]-Equality : Symmetric-kind → ∀ {a} → Set a → Setoid _ _
+ [ k ]-Equality A = Related.InducedEquivalence₂ k (_∈_ {A = A})
- infix 4 _≈[_]_
+ infix 4 _∼[_]_
- _≈[_]_ : ∀ {a} {A : Set a} → List A → Kind → List A → Set _
- _≈[_]_ {A = A} xs k ys = Setoid._≈_ ([ k ]-Equality A) xs ys
+ _∼[_]_ : ∀ {a} {A : Set a} → List A → Kind → List A → Set _
+ _∼[_]_ {A = A} xs k ys = Preorder._∼_ ([ k ]-Order A) xs ys
- -- Bag equality implies set equality.
+ -- Bag equality implies the other relations.
- bag-=⇒set-= : ∀ {a} {A : Set a} {xs ys : List A} →
- xs ≈[ bag ] ys → xs ≈[ set ] ys
- bag-=⇒set-= xs≈ys = Inv.⇿⇒ xs≈ys
+ bag-=⇒ : ∀ {k a} {A : Set a} {xs ys : List A} →
+ xs ∼[ bag ] ys → xs ∼[ k ] ys
+ bag-=⇒ xs≈ys = ↔⇒ xs≈ys
-- "Equational" reasoning for _⊆_.
@@ -192,17 +201,16 @@ module Membership-≡ where
open module PR {a} {A : Set a} = PreR (⊆-preorder A) public
renaming (_∼⟨_⟩_ to _⊆⟨_⟩_; _≈⟨_⟩_ to _≡⟨_⟩_)
- infixr 2 _≈⟨_⟩_
+ infixr 2 _∼⟨_⟩_
infix 1 _∈⟨_⟩_
_∈⟨_⟩_ : ∀ {a} {A : Set a} x {xs ys : List A} →
x ∈ xs → xs IsRelatedTo ys → x ∈ ys
x ∈⟨ x∈xs ⟩ xs⊆ys = (begin xs⊆ys) x∈xs
- _≈⟨_⟩_ : ∀ {k a} {A : Set a} xs {ys zs : List A} →
- xs ≈[ k ] ys → ys IsRelatedTo zs → xs IsRelatedTo zs
- xs ≈⟨ xs≈ys ⟩ ys≈zs =
- xs ⊆⟨ _⟨$⟩_ (Equivalent.to (Inv.⇒⇔ xs≈ys)) ⟩ ys≈zs
+ _∼⟨_⟩_ : ∀ {k a} {A : Set a} xs {ys zs : List A} →
+ xs ∼[ ⌊ k ⌋→ ] ys → ys IsRelatedTo zs → xs IsRelatedTo zs
+ xs ∼⟨ xs≈ys ⟩ ys≈zs = xs ⊆⟨ ⇒→ xs≈ys ⟩ ys≈zs
------------------------------------------------------------------------
-- Another function
diff --git a/src/Data/List/Any/BagAndSetEquality.agda b/src/Data/List/Any/BagAndSetEquality.agda
index 00c0e34..bbda405 100644
--- a/src/Data/List/Any/BagAndSetEquality.agda
+++ b/src/Data/List/Any/BagAndSetEquality.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties related to bag and set equality
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- Bag and set equality are defined in Data.List.Any.
module Data.List.Any.BagAndSetEquality where
@@ -11,96 +11,61 @@ module Data.List.Any.BagAndSetEquality where
open import Algebra
open import Algebra.FunctionProperties
open import Category.Monad
-open import Data.Empty
open import Data.List as List
import Data.List.Properties as LP
open import Data.List.Any as Any using (Any); open Any.Any
open import Data.List.Any.Properties
open import Data.Product
open import Data.Sum
-open import Data.Unit
open import Function
open import Function.Equality using (_⟨$⟩_)
import Function.Equivalence as FE
-open import Function.Inverse as Inv using (_⇿_; module Inverse)
-open import Function.Inverse.TypeIsomorphisms
-open import Level
+open import Function.Inverse as Inv using (_↔_; module Inverse)
+open import Function.Related as Related using (↔⇒; ⌊_⌋; ⌊_⌋→; ⇒→)
+open import Function.Related.TypeIsomorphisms
open import Relation.Binary
import Relation.Binary.EqReasoning as EqR
-import Relation.Binary.Sigma.Pointwise as Σ
open import Relation.Binary.PropositionalEquality as P
- using (_≡_; _≗_; _with-≡_)
+ using (_≡_; _≗_)
+open import Relation.Binary.Sum
open import Relation.Nullary
open Any.Membership-≡
private
- module Eq {k a} {A : Set a} = Setoid ([ k ]-Equality A)
+ module Eq {k a} {A : Set a} = Setoid ([ k ]-Equality A)
+ module Ord {k a} {A : Set a} = Preorder ([ k ]-Order A)
module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
open module ListMonad {ℓ} = RawMonad (List.monad {ℓ = ℓ})
module ListMonoid {a} {A : Set a} = Monoid (List.monoid A)
--- _++_ and [] form a commutative monoid, with either bag or set
--- equality as the underlying equality.
-
-commutativeMonoid : ∀ {a} → Kind → Set a → CommutativeMonoid _ _
-commutativeMonoid {a} 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 $ ++⇿ {a = a} {p = a} ⟩
- (x ∈ xs₁ ⊎ x ∈ xs₃) ≈⟨ xs₁≈xs₂ ⟨ ×⊎.+-cong ⟩ xs₃≈xs₄ ⟩
- (x ∈ xs₂ ⊎ x ∈ xs₄) ⇿⟨ ++⇿ {a = a} {p = a} ⟩
- x ∈ xs₂ ++ xs₄ ∎
- }
- ; identityˡ = λ xs {x} → x ∈ xs ∎
- ; comm = λ xs ys {x} →
- x ∈ xs ++ ys ⇿⟨ ++⇿++ {a = a} {p = a} xs ys ⟩
- x ∈ ys ++ xs ∎
- }
- }
- where open Inv.EquationalReasoning
-
--- The only list which is bag or set equal to the empty list is the
--- empty list itself.
-
-empty-unique : ∀ {k a} {A : Set a} {xs : List A} →
- xs ≈[ k ] [] → xs ≡ []
-empty-unique {xs = []} _ = P.refl
-empty-unique {xs = _ ∷ _} ∷≈[]
- with FE.Equivalent.to (Inv.⇒⇔ ∷≈[]) ⟨$⟩ here P.refl
-... | ()
+------------------------------------------------------------------------
+-- Congruence lemmas
--- _++_ is idempotent (under set equality).
+-- _∷_ is a congruence.
-++-idempotent : ∀ {a} {A : Set a} →
- Idempotent (λ (xs ys : List A) → xs ≈[ set ] ys) _++_
-++-idempotent {a} xs {x} =
- x ∈ xs ++ xs ≈⟨ FE.equivalent ([ id , id ]′ ∘ _⟨$⟩_ (Inverse.from $ ++⇿ {a = a} {p = a}))
- (_⟨$⟩_ (Inverse.to $ ++⇿ {a = a} {p = a}) ∘ inj₁) ⟩
- x ∈ xs ∎
- where open Inv.EquationalReasoning
+∷-cong : ∀ {a k} {A : Set a} {x₁ x₂ : A} {xs₁ xs₂} →
+ x₁ ≡ x₂ → xs₁ ∼[ k ] xs₂ → x₁ ∷ xs₁ ∼[ k ] x₂ ∷ xs₂
+∷-cong {x₂ = x} {xs₁} {xs₂} P.refl xs₁≈xs₂ {y} =
+ y ∈ x ∷ xs₁ ↔⟨ sym $ ∷↔ (_≡_ y) ⟩
+ (y ≡ x ⊎ y ∈ xs₁) ∼⟨ (y ≡ x ∎) ⊎-cong xs₁≈xs₂ ⟩
+ (y ≡ x ⊎ y ∈ xs₂) ↔⟨ ∷↔ (_≡_ y) ⟩
+ y ∈ x ∷ xs₂ ∎
+ where open Related.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₂
+ 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⇿ {a = ℓ} {b = ℓ} {p = ℓ} ⟩
- Any (λ y → x ≡ f₁ y) xs₁ ≈⟨ Any-cong (Inv.⇿⇒ ∘ helper) xs₁≈xs₂ ⟩
- Any (λ y → x ≡ f₂ y) xs₂ ⇿⟨ map⇿ {a = ℓ} {b = ℓ} {p = ℓ} ⟩
+ x ∈ List.map f₁ xs₁ ↔⟨ sym $ map↔ {a = ℓ} {b = ℓ} {p = ℓ} ⟩
+ Any (λ y → x ≡ f₁ y) xs₁ ∼⟨ Any-cong (↔⇒ ∘ helper) xs₁≈xs₂ ⟩
+ Any (λ y → x ≡ f₂ y) xs₂ ↔⟨ map↔ {a = ℓ} {b = ℓ} {p = ℓ} ⟩
x ∈ List.map f₂ xs₂ ∎
where
- open Inv.EquationalReasoning
+ open Related.EquationalReasoning
- helper : ∀ y → x ≡ f₁ y ⇿ x ≡ f₂ y
+ 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))
@@ -110,68 +75,127 @@ map-cong {ℓ} {f₁ = f₁} {f₂} {xs₁} {xs₂} f₁≗f₂ xs₁≈xs₂ {x
}
}
+-- _++_ is a congruence.
+
+++-cong : ∀ {a k} {A : Set a} {xs₁ xs₂ ys₁ ys₂ : List A} →
+ xs₁ ∼[ k ] xs₂ → ys₁ ∼[ k ] ys₂ →
+ xs₁ ++ ys₁ ∼[ k ] xs₂ ++ ys₂
+++-cong {a} {xs₁ = xs₁} {xs₂} {ys₁} {ys₂} xs₁≈xs₂ ys₁≈ys₂ {x} =
+ x ∈ xs₁ ++ ys₁ ↔⟨ sym $ ++↔ {a = a} {p = a} ⟩
+ (x ∈ xs₁ ⊎ x ∈ ys₁) ∼⟨ xs₁≈xs₂ ⊎-cong ys₁≈ys₂ ⟩
+ (x ∈ xs₂ ⊎ x ∈ ys₂) ↔⟨ ++↔ {a = a} {p = a} ⟩
+ x ∈ xs₂ ++ ys₂ ∎
+ where open Related.EquationalReasoning
+
-- concat is a congruence.
concat-cong : ∀ {a k} {A : Set a} {xss₁ xss₂ : List (List A)} →
- xss₁ ≈[ k ] xss₂ → concat xss₁ ≈[ k ] concat xss₂
+ xss₁ ∼[ k ] xss₂ → concat xss₁ ∼[ k ] concat xss₂
concat-cong {a} {xss₁ = xss₁} {xss₂} xss₁≈xss₂ {x} =
- x ∈ concat xss₁ ⇿⟨ sym $ concat⇿ {a = a} {p = a} ⟩
- Any (Any (_≡_ x)) xss₁ ≈⟨ Any-cong (λ _ → _ ∎) xss₁≈xss₂ ⟩
- Any (Any (_≡_ x)) xss₂ ⇿⟨ concat⇿ {a = a} {p = a} ⟩
+ x ∈ concat xss₁ ↔⟨ sym $ concat↔ {a = a} {p = a} ⟩
+ Any (Any (_≡_ x)) xss₁ ∼⟨ Any-cong (λ _ → _ ∎) xss₁≈xss₂ ⟩
+ Any (Any (_≡_ x)) xss₂ ↔⟨ concat↔ {a = a} {p = a} ⟩
x ∈ concat xss₂ ∎
- where open Inv.EquationalReasoning
+ where open Related.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₂)
+ 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 $ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟩
- Any (λ y → x ∈ f₁ y) xs₁ ≈⟨ Any-cong (λ x → f₁≈f₂ x) xs₁≈xs₂ ⟩
- Any (λ y → x ∈ f₂ y) xs₂ ⇿⟨ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟩
+ x ∈ (xs₁ >>= f₁) ↔⟨ sym $ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
+ Any (λ y → x ∈ f₁ y) xs₁ ∼⟨ Any-cong (λ x → f₁≈f₂ x) xs₁≈xs₂ ⟩
+ Any (λ y → x ∈ f₂ y) xs₂ ↔⟨ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
x ∈ (xs₂ >>= f₂) ∎
- where open Inv.EquationalReasoning
+ where open Related.EquationalReasoning
-- _⊛_ is a congruence.
⊛-cong : ∀ {ℓ k} {A B : Set ℓ} {fs₁ fs₂ : List (A → B)} {xs₁ xs₂} →
- fs₁ ≈[ k ] fs₂ → xs₁ ≈[ k ] xs₂ → fs₁ ⊛ xs₁ ≈[ k ] fs₂ ⊛ 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
+ where open Related.EquationalReasoning
-- _⊗_ is a congruence.
⊗-cong : ∀ {ℓ k} {A B : Set ℓ} {xs₁ xs₂ : List A} {ys₁ ys₂ : List B} →
- xs₁ ≈[ k ] xs₂ → ys₁ ≈[ k ] ys₂ →
- (xs₁ ⊗ ys₁) ≈[ k ] (xs₂ ⊗ ys₂)
+ xs₁ ∼[ k ] xs₂ → ys₁ ∼[ k ] ys₂ →
+ (xs₁ ⊗ ys₁) ∼[ k ] (xs₂ ⊗ ys₂)
⊗-cong {ℓ} xs₁≈xs₂ ys₁≈ys₂ =
- ⊛-cong (⊛-cong (Eq.refl {x = [ _,_ {a = ℓ} {b = ℓ} ]})
+ ⊛-cong (⊛-cong (Ord.refl {x = [ _,_ {a = ℓ} {b = ℓ} ]})
xs₁≈xs₂)
ys₁≈ys₂
+------------------------------------------------------------------------
+-- Other properties
+
+-- _++_ and [] form a commutative monoid, with either bag or set
+-- equality as the underlying equality.
+
+commutativeMonoid : ∀ {a} → Symmetric-kind → Set a →
+ CommutativeMonoid _ _
+commutativeMonoid {a} 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 = ++-cong
+ }
+ ; identityˡ = λ xs {x} → x ∈ xs ∎
+ ; comm = λ xs ys {x} →
+ x ∈ xs ++ ys ↔⟨ ++↔++ {a = a} {p = a} xs ys ⟩
+ x ∈ ys ++ xs ∎
+ }
+ }
+ where open Related.EquationalReasoning
+
+-- The only list which is bag or set equal to the empty list (or a
+-- subset or subbag of the list) is the empty list itself.
+
+empty-unique : ∀ {k a} {A : Set a} {xs : List A} →
+ xs ∼[ ⌊ k ⌋→ ] [] → xs ≡ []
+empty-unique {xs = []} _ = P.refl
+empty-unique {xs = _ ∷ _} ∷∼[] with ⇒→ ∷∼[] (here P.refl)
+... | ()
+
+-- _++_ is idempotent (under set equality).
+
+++-idempotent : ∀ {a} {A : Set a} →
+ Idempotent (λ (xs ys : List A) → xs ∼[ set ] ys) _++_
+++-idempotent {a} xs {x} =
+ x ∈ xs ++ xs ∼⟨ FE.equivalence ([ id , id ]′ ∘ _⟨$⟩_ (Inverse.from $ ++↔ {a = a} {p = a}))
+ (_⟨$⟩_ (Inverse.to $ ++↔ {a = a} {p = a}) ∘ inj₁) ⟩
+ x ∈ xs ∎
+ where open Related.EquationalReasoning
+
-- 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)
+ (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 $ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟩
- Any (λ x → y ∈ f x ++ g x) xs ⇿⟨ sym (Any-cong (λ _ → ++⇿ {a = ℓ} {p = ℓ}) (_ ∎)) ⟩
- Any (λ x → y ∈ f x ⊎ y ∈ g x) xs ⇿⟨ sym $ ⊎⇿ {a = ℓ} {p = ℓ} {q = ℓ} ⟩
- (Any (λ x → y ∈ f x) xs ⊎ Any (λ x → y ∈ g x) xs) ⇿⟨ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟨ ×⊎.+-cong {ℓ = ℓ} ⟩ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟩
- (y ∈ (xs >>= f) ⊎ y ∈ (xs >>= g)) ⇿⟨ ++⇿ {a = ℓ} {p = ℓ} ⟩
+ y ∈ (xs >>= λ x → f x ++ g x) ↔⟨ sym $ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
+ Any (λ x → y ∈ f x ++ g x) xs ↔⟨ sym (Any-cong (λ _ → ++↔ {a = ℓ} {p = ℓ}) (_ ∎)) ⟩
+ Any (λ x → y ∈ f x ⊎ y ∈ g x) xs ↔⟨ sym $ ⊎↔ {a = ℓ} {p = ℓ} {q = ℓ} ⟩
+ (Any (λ x → y ∈ f x) xs ⊎ Any (λ x → y ∈ g x) xs) ↔⟨ >>=↔ {ℓ = ℓ} {p = ℓ} ⟨ ×⊎.+-cong {ℓ = ℓ} ⟩ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
+ (y ∈ (xs >>= f) ⊎ y ∈ (xs >>= g)) ↔⟨ ++↔ {a = ℓ} {p = ℓ} ⟩
y ∈ (xs >>= f) ++ (xs >>= g) ∎
- where open Inv.EquationalReasoning
+ where open Related.EquationalReasoning
-- The same applies to _⊛_.
⊛-left-distributive :
∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) xs₁ xs₂ →
- fs ⊛ (xs₁ ++ xs₂) ≈[ bag ] (fs ⊛ xs₁) ++ (fs ⊛ xs₂)
+ fs ⊛ (xs₁ ++ xs₂) ∼[ bag ] (fs ⊛ xs₁) ++ (fs ⊛ xs₂)
⊛-left-distributive {B = B} fs xs₁ xs₂ = begin
fs ⊛ (xs₁ ++ xs₂) ≡⟨ P.refl ⟩
(fs >>= λ f → xs₁ ++ xs₂ >>= return ∘ f) ≡⟨ (LP.Monad.cong (P.refl {x = fs}) λ f →
@@ -192,133 +216,57 @@ private
¬-drop-cons :
∀ {a} {A : Set a} {x : A} →
- ¬ (∀ {xs ys} → x ∷ xs ≈[ set ] x ∷ ys → xs ≈[ set ] ys)
+ ¬ (∀ {xs ys} → x ∷ xs ∼[ set ] x ∷ ys → xs ∼[ set ] ys)
¬-drop-cons {x = x} drop-cons
- with FE.Equivalent.to x≈[] ⟨$⟩ here P.refl
+ with FE.Equivalence.to x∼[] ⟨$⟩ here P.refl
where
- x,x≈x : (x ∷ x ∷ []) ≈[ set ] [ x ]
+ x,x≈x : (x ∷ x ∷ []) ∼[ set ] [ x ]
x,x≈x = ++-idempotent [ x ]
- x≈[] : [ x ] ≈[ set ] []
- x≈[] = drop-cons x,x≈x
+ x∼[] : [ x ] ∼[ set ] []
+ x∼[] = drop-cons x,x≈x
... | ()
-- However, the corresponding property does hold for bag equality.
drop-cons : ∀ {a} {A : Set a} {x : A} {xs ys} →
- x ∷ xs ≈[ bag ] x ∷ ys → xs ≈[ bag ] ys
-drop-cons {A = A} {x} {xs} {ys} x∷xs≈x∷ys {z} =
- z ∈ xs ⇿⟨ lemma xs ⟩
- (∃ λ (z∈x∷xs : z ∈ x ∷ xs) → There z∈x∷xs) ⇿⟨ Σ.⇿ x∷xs≈x∷ys′ x∷xs≈x∷ys′-preserves-There ⟩
- (∃ λ (z∈x∷ys : z ∈ x ∷ ys) → There z∈x∷ys) ⇿⟨ sym $ lemma ys ⟩
- z ∈ ys ∎
- where
- open Inv.EquationalReasoning
-
- -- Inhabited for there but not here.
-
- There : ∀ {z : A} {xs} → z ∈ xs → Set
- There (here _) = ⊥
- There (there _) = ⊤
-
- -- There is proof irrelevant.
-
- proof-irrelevance : ∀ {z xs} {z∈xs : z ∈ xs}
- (p q : There z∈xs) → p ≡ q
- proof-irrelevance {z∈xs = here _} () ()
- proof-irrelevance {z∈xs = there _} _ _ = P.refl
-
- -- An isomorphism.
-
- lemma : ∀ xs → z ∈ xs ⇿ (∃ λ (z∈x∷xs : z ∈ x ∷ xs) → There z∈x∷xs)
- lemma xs = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = λ _ → P.refl
- ; right-inverse-of = to∘from
- }
+ x ∷ xs ∼[ bag ] x ∷ ys → xs ∼[ bag ] ys
+drop-cons {A = A} {x} {xs} {ys} x∷xs≈x∷ys {z} = record
+ { to = P.→-to-⟶ $ f x∷xs≈x∷ys
+ ; from = P.→-to-⟶ $ f $ Inv.sym x∷xs≈x∷ys
+ ; inverse-of = record
+ { left-inverse-of = f∘f x∷xs≈x∷ys
+ ; right-inverse-of = f∘f $ Inv.sym x∷xs≈x∷ys
}
- where
- to : z ∈ xs → (∃ λ (z∈x∷xs : z ∈ x ∷ xs) → There z∈x∷xs)
- to z∈xs = (there z∈xs , _)
-
- from : (∃ λ (z∈x∷xs : z ∈ x ∷ xs) → There z∈x∷xs) → z ∈ xs
- from (here z≡x , ())
- from (there z∈xs , _) = z∈xs
-
- to∘from : ∀ p → to (from p) ≡ p
- to∘from (here z≡x , ())
- to∘from (there z∈xs , _) = P.refl
-
- -- The isomorphisms x∷xs≈x∷ys may not preserve There, because they
- -- could map here P.refl to something other than here P.refl.
- -- However, we can construct isomorphisms which do preserve There:
+ }
+ where
+ open Inverse
+ open P.≡-Reasoning
+
+ f : ∀ {xs ys z} → (z ∈ x ∷ xs) ↔ (z ∈ x ∷ ys) → z ∈ xs → z ∈ ys
+ f inv z∈xs with to inv ⟨$⟩ there z∈xs | left-inverse-of inv (there z∈xs)
+ f inv z∈xs | there z∈ys | left⁺ = z∈ys
+ f inv z∈xs | here z≡x | left⁺ with to inv ⟨$⟩ here z≡x | left-inverse-of inv (here z≡x)
+ f inv z∈xs | here z≡x | left⁺ | there z∈ys | left⁰ = z∈ys
+ f inv z∈xs | here P.refl | left⁺ | here P.refl | left⁰ with begin
+ here P.refl ≡⟨ P.sym left⁰ ⟩
+ from inv ⟨$⟩ here P.refl ≡⟨ left⁺ ⟩
+ there z∈xs ∎
+ ... | ()
- x∷xs≈x∷ys′ : x ∷ xs ≈[ bag ] x ∷ ys
- x∷xs≈x∷ys′ = record
- { to = P.→-to-⟶ $ f x∷xs≈x∷ys
- ; from = P.→-to-⟶ $ f $ Inv.sym x∷xs≈x∷ys
- ; inverse-of = record
- { left-inverse-of = f∘f x∷xs≈x∷ys
- ; right-inverse-of = f∘f $ Inv.sym x∷xs≈x∷ys
- }
- }
- where
- f : ∀ {xs ys} → x ∷ xs ≈[ bag ] x ∷ ys →
- ∀ {z} → z ∈ x ∷ xs → z ∈ x ∷ ys
- f x∷xs≈x∷ys (here P.refl) = here P.refl
- f x∷xs≈x∷ys (there z∈xs) with Inverse.to x∷xs≈x∷ys ⟨$⟩ there z∈xs
- ... | there z∈ys = there z∈ys
- ... | here P.refl = Inverse.to x∷xs≈x∷ys ⟨$⟩ here P.refl
-
- f∘f : ∀ {xs ys}
- (x∷xs≈x∷ys : x ∷ xs ≈[ bag ] x ∷ ys) →
- ∀ {z} (p : z ∈ x ∷ xs) →
- f (Inv.sym x∷xs≈x∷ys) (f x∷xs≈x∷ys p) ≡ p
- f∘f x∷xs≈x∷ys (here P.refl) = P.refl
- f∘f x∷xs≈x∷ys (there z∈xs)
- with Inverse.to x∷xs≈x∷ys ⟨$⟩ there z∈xs
- | Inverse.left-inverse-of x∷xs≈x∷ys (there z∈xs)
- ... | there z∈ys | left rewrite left = P.refl
- ... | here P.refl | left
- with Inverse.to x∷xs≈x∷ys ⟨$⟩ here P.refl
- | Inverse.left-inverse-of x∷xs≈x∷ys (here P.refl)
- ... | there z∈xs′ | left′ rewrite left′ = left
- ... | here P.refl | left′ rewrite left′ = left
-
- x∷xs≈x∷ys′-preserves-There :
- ∀ {z} {z∈x∷xs : z ∈ x ∷ xs} →
- There z∈x∷xs ⇿ There (Inverse.to x∷xs≈x∷ys′ ⟨$⟩ z∈x∷xs)
- x∷xs≈x∷ys′-preserves-There {z∈x∷xs = z∈x∷xs} = record
- { to = P.→-to-⟶ $ to z∈x∷xs
- ; from = P.→-to-⟶ $ from z∈x∷xs
- ; inverse-of = record
- { left-inverse-of = λ _ → proof-irrelevance _ _
- ; right-inverse-of = λ _ → proof-irrelevance _ _
- }
- }
- where
- open P.≡-Reasoning renaming (_∎ to _□)
-
- to : ∀ {z} (z∈x∷xs : z ∈ x ∷ xs) →
- There z∈x∷xs → There (Inverse.to x∷xs≈x∷ys′ ⟨$⟩ z∈x∷xs)
- to (here _) ()
- to (there z∈) _ with P.inspect (Inverse.to x∷xs≈x∷ys ⟨$⟩ there z∈)
- ... | there _ with-≡ eq rewrite eq = _
- ... | here P.refl with-≡ eq rewrite eq
- with P.inspect (Inverse.to x∷xs≈x∷ys ⟨$⟩ here P.refl)
- ... | there _ with-≡ eq′ rewrite eq′ = _
- ... | here P.refl with-≡ eq′ rewrite eq′
- with begin
- there z∈ ≡⟨ P.sym $ Inverse.left-inverse-of x∷xs≈x∷ys _ ⟩
- Inverse.from x∷xs≈x∷ys ⟨$⟩ (Inverse.to x∷xs≈x∷ys ⟨$⟩ there z∈) ≡⟨ P.cong (_⟨$⟩_ (Inverse.from x∷xs≈x∷ys)) eq ⟩
- Inverse.from x∷xs≈x∷ys ⟨$⟩ here P.refl ≡⟨ P.cong (_⟨$⟩_ (Inverse.from x∷xs≈x∷ys)) (P.sym eq′) ⟩
- Inverse.from x∷xs≈x∷ys ⟨$⟩ (Inverse.to x∷xs≈x∷ys ⟨$⟩ here P.refl) ≡⟨ Inverse.left-inverse-of x∷xs≈x∷ys _ ⟩
- Any.here {xs = xs} (P.refl {x = x}) □
- ... | ()
-
- from : ∀ {z} (z∈x∷xs : z ∈ x ∷ xs) →
- There (Inverse.to x∷xs≈x∷ys′ ⟨$⟩ z∈x∷xs) → There z∈x∷xs
- from (here P.refl) ()
- from (there z∈xs) _ = _
+ f∘f : ∀ {xs ys z}
+ (inv : (z ∈ x ∷ xs) ↔ (z ∈ x ∷ ys)) (p : z ∈ xs) →
+ f (Inv.sym inv) (f inv p) ≡ p
+ f∘f inv z∈xs with to inv ⟨$⟩ there z∈xs | left-inverse-of inv (there z∈xs)
+ f∘f inv z∈xs | there z∈ys | left⁺ with from inv ⟨$⟩ there z∈ys | right-inverse-of inv (there z∈ys)
+ f∘f inv z∈xs | there z∈ys | P.refl | .(there z∈xs) | _ = P.refl
+ f∘f inv z∈xs | here z≡x | left⁺ with to inv ⟨$⟩ here z≡x | left-inverse-of inv (here z≡x)
+ f∘f inv z∈xs | here z≡x | left⁺ | there z∈ys | left⁰ with from inv ⟨$⟩ there z∈ys | right-inverse-of inv (there z∈ys)
+ f∘f inv z∈xs | here z≡x | left⁺ | there z∈ys | P.refl | .(here z≡x) | _ with from inv ⟨$⟩ here z≡x
+ | right-inverse-of inv (here z≡x)
+ f∘f inv z∈xs | here z≡x | P.refl | there z∈ys | P.refl | .(here z≡x) | _ | .(there z∈xs) | _ = P.refl
+ f∘f inv z∈xs | here P.refl | left⁺ | here P.refl | left⁰ with begin
+ here P.refl ≡⟨ P.sym left⁰ ⟩
+ from inv ⟨$⟩ here P.refl ≡⟨ left⁺ ⟩
+ there z∈xs ∎
+ ... | ()
diff --git a/src/Data/List/Any/Membership.agda b/src/Data/List/Any/Membership.agda
index c98235d..2cc230e 100644
--- a/src/Data/List/Any/Membership.agda
+++ b/src/Data/List/Any/Membership.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties related to list membership
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- List membership is defined in Data.List.Any. This module does not
-- treat the general variant of list membership, parametrised on a
-- setoid, only the variant where the equality is fixed to be
@@ -17,10 +17,11 @@ open import Data.Bool
open import Data.Empty
open import Function
open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence using (module Equivalent)
+open import Function.Equivalence using (module Equivalence)
import Function.Injection as Inj
-open import Function.Inverse as Inv using (_⇿_; module Inverse)
-open import Function.Inverse.TypeIsomorphisms
+open import Function.Inverse as Inv using (_↔_; module Inverse)
+import Function.Related as Related
+open import Function.Related.TypeIsomorphisms
open import Data.List as List
open import Data.List.Any as Any using (Any; here; there)
open import Data.List.Any.Properties
@@ -28,7 +29,6 @@ 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; _≗_)
@@ -46,52 +46,52 @@ private
------------------------------------------------------------------------
-- Properties relating _∈_ to various list functions
-map-∈⇿ : ∀ {a b} {A : Set a} {B : Set b} {f : A → B} {y xs} →
- (∃ λ x → x ∈ xs × y ≡ f x) ⇿ y ∈ List.map f xs
-map-∈⇿ {a} {b} {f = f} {y} {xs} =
- (∃ λ x → x ∈ xs × y ≡ f x) ⇿⟨ Any⇿ {a = a} {p = b} ⟩
- Any (λ x → y ≡ f x) xs ⇿⟨ map⇿ {a = a} {b = b} {p = b} ⟩
+map-∈↔ : ∀ {a b} {A : Set a} {B : Set b} {f : A → B} {y xs} →
+ (∃ λ x → x ∈ xs × y ≡ f x) ↔ y ∈ List.map f xs
+map-∈↔ {a} {b} {f = f} {y} {xs} =
+ (∃ λ x → x ∈ xs × y ≡ f x) ↔⟨ Any↔ {a = a} {p = b} ⟩
+ Any (λ x → y ≡ f x) xs ↔⟨ map↔ {a = a} {b = b} {p = b} ⟩
y ∈ List.map f xs ∎
- where open Inv.EquationalReasoning
-
-concat-∈⇿ : ∀ {a} {A : Set a} {x : A} {xss} →
- (∃ λ xs → x ∈ xs × xs ∈ xss) ⇿ x ∈ concat xss
-concat-∈⇿ {a} {x = x} {xss} =
- (∃ λ xs → x ∈ xs × xs ∈ xss) ⇿⟨ Σ.cong {a₁ = a} {b₁ = a} {b₂ = a} Inv.id $ ×⊎.*-comm _ _ ⟩
- (∃ λ xs → xs ∈ xss × x ∈ xs) ⇿⟨ Any⇿ {a = a} {p = a} ⟩
- Any (Any (_≡_ x)) xss ⇿⟨ concat⇿ {a = a} {p = a} ⟩
+ where open Related.EquationalReasoning
+
+concat-∈↔ : ∀ {a} {A : Set a} {x : A} {xss} →
+ (∃ λ xs → x ∈ xs × xs ∈ xss) ↔ x ∈ concat xss
+concat-∈↔ {a} {x = x} {xss} =
+ (∃ λ xs → x ∈ xs × xs ∈ xss) ↔⟨ Σ.cong {a₁ = a} {b₁ = a} {b₂ = a} Inv.id $ ×⊎.*-comm _ _ ⟩
+ (∃ λ xs → xs ∈ xss × x ∈ xs) ↔⟨ Any↔ {a = a} {p = a} ⟩
+ Any (Any (_≡_ x)) xss ↔⟨ concat↔ {a = a} {p = a} ⟩
x ∈ concat xss ∎
- where open Inv.EquationalReasoning
+ where open Related.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⇿ {a = ℓ} {p = ℓ} ⟩
- Any (Any (_≡_ y) ∘ f) xs ⇿⟨ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟩
+>>=-∈↔ : ∀ {ℓ} {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↔ {a = ℓ} {p = ℓ} ⟩
+ Any (Any (_≡_ y) ∘ f) xs ↔⟨ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
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 {a₁ = ℓ} {b₁ = ℓ} {b₂ = ℓ} Inv.id (∃∃⇿∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
- (∃ λ f → f ∈ fs × ∃ λ x → x ∈ xs × y ≡ f x) ⇿⟨ Σ.cong {a₁ = ℓ} {b₁ = ℓ} {b₂ = ℓ}
- Inv.id ((_ ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any⇿ {a = ℓ} {p = ℓ}) ⟩
- (∃ λ f → f ∈ fs × Any (_≡_ y ∘ f) xs) ⇿⟨ Any⇿ {a = ℓ} {p = ℓ} ⟩
- Any (λ f → Any (_≡_ y ∘ f) xs) fs ⇿⟨ ⊛⇿ ⟩
+ where open Related.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 {a₁ = ℓ} {b₁ = ℓ} {b₂ = ℓ} Inv.id (∃∃↔∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
+ (∃ λ f → f ∈ fs × ∃ λ x → x ∈ xs × y ≡ f x) ↔⟨ Σ.cong {a₁ = ℓ} {b₁ = ℓ} {b₂ = ℓ}
+ Inv.id ((_ ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any↔ {a = ℓ} {p = ℓ}) ⟩
+ (∃ λ f → f ∈ fs × Any (_≡_ y ∘ f) xs) ↔⟨ Any↔ {a = ℓ} {p = ℓ} ⟩
+ Any (λ f → Any (_≡_ y ∘ f) xs) fs ↔⟨ ⊛↔ ⟩
y ∈ fs ⊛ xs ∎
- where open Inv.EquationalReasoning
+ where open Related.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 (_ ∎) ⟩
+⊗-∈↔ : ∀ {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
+ open Related.EquationalReasoning
- helper : (p : A × B) → (x ≡ proj₁ p × y ≡ proj₂ p) ⇿ (x , y) ≡ p
+ 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₂ >
@@ -108,60 +108,60 @@ concat-∈⇿ {a} {x = x} {xss} =
mono : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
xs ⊆ ys → Any P xs → Any P ys
mono xs⊆ys =
- _⟨$⟩_ (Inverse.to Any⇿) ∘′
+ _⟨$⟩_ (Inverse.to Any↔) ∘′
Prod.map id (Prod.map xs⊆ys id) ∘
- _⟨$⟩_ (Inverse.from Any⇿)
+ _⟨$⟩_ (Inverse.from Any↔)
map-mono : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) {xs ys} →
xs ⊆ ys → List.map f xs ⊆ List.map f ys
map-mono f xs⊆ys =
- _⟨$⟩_ (Inverse.to map-∈⇿) ∘
+ _⟨$⟩_ (Inverse.to map-∈↔) ∘
Prod.map id (Prod.map xs⊆ys id) ∘
- _⟨$⟩_ (Inverse.from map-∈⇿)
+ _⟨$⟩_ (Inverse.from map-∈↔)
_++-mono_ : ∀ {a} {A : Set a} {xs₁ xs₂ ys₁ ys₂ : List A} →
xs₁ ⊆ ys₁ → xs₂ ⊆ ys₂ → xs₁ ++ xs₂ ⊆ ys₁ ++ ys₂
_++-mono_ xs₁⊆ys₁ xs₂⊆ys₂ =
- _⟨$⟩_ (Inverse.to ++⇿) ∘
+ _⟨$⟩_ (Inverse.to ++↔) ∘
Sum.map xs₁⊆ys₁ xs₂⊆ys₂ ∘
- _⟨$⟩_ (Inverse.from ++⇿)
+ _⟨$⟩_ (Inverse.from ++↔)
concat-mono : ∀ {a} {A : Set a} {xss yss : List (List A)} →
xss ⊆ yss → concat xss ⊆ concat yss
concat-mono {a} xss⊆yss =
- _⟨$⟩_ (Inverse.to $ concat-∈⇿ {a = a}) ∘
+ _⟨$⟩_ (Inverse.to $ concat-∈↔ {a = a}) ∘
Prod.map id (Prod.map id xss⊆yss) ∘
- _⟨$⟩_ (Inverse.from $ concat-∈⇿ {a = a})
+ _⟨$⟩_ (Inverse.from $ concat-∈↔ {a = a})
>>=-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 $ >>=-∈⇿ {ℓ = ℓ}) ∘
+ _⟨$⟩_ (Inverse.to $ >>=-∈↔ {ℓ = ℓ}) ∘
Prod.map id (Prod.map xs⊆ys f⊆g) ∘
- _⟨$⟩_ (Inverse.from $ >>=-∈⇿ {ℓ = ℓ})
+ _⟨$⟩_ (Inverse.from $ >>=-∈↔ {ℓ = ℓ})
_⊛-mono_ : ∀ {ℓ} {A B : Set ℓ}
{fs gs : List (A → B)} {xs ys : List A} →
fs ⊆ gs → xs ⊆ ys → (fs ⊛ xs) ⊆ (gs ⊛ ys)
_⊛-mono_ {fs = fs} {gs} fs⊆gs xs⊆ys =
- _⟨$⟩_ (Inverse.to $ ⊛-∈⇿ gs) ∘
+ _⟨$⟩_ (Inverse.to $ ⊛-∈↔ gs) ∘
Prod.map id (Prod.map id (Prod.map fs⊆gs (Prod.map xs⊆ys id))) ∘
- _⟨$⟩_ (Inverse.from $ ⊛-∈⇿ fs)
+ _⟨$⟩_ (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 ⊗-∈⇿) ∘
+ _⟨$⟩_ (Inverse.to ⊗-∈↔) ∘
Prod.map xs₁⊆ys₁ xs₂⊆ys₂ ∘
- _⟨$⟩_ (Inverse.from ⊗-∈⇿)
+ _⟨$⟩_ (Inverse.from ⊗-∈↔)
any-mono : ∀ {a} {A : Set a} (p : A → Bool) →
∀ {xs ys} → xs ⊆ ys → T (any p xs) → T (any p ys)
any-mono {a} p xs⊆ys =
- _⟨$⟩_ (Equivalent.to $ any⇔ {a = a}) ∘
+ _⟨$⟩_ (Equivalence.to $ any⇔ {a = a}) ∘
mono xs⊆ys ∘
- _⟨$⟩_ (Equivalent.from $ any⇔ {a = a})
+ _⟨$⟩_ (Equivalence.from $ any⇔ {a = a})
map-with-∈-mono :
∀ {a b} {A : Set a} {B : Set b}
@@ -170,12 +170,12 @@ map-with-∈-mono :
(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-∈⇿) ∘
+ _⟨$⟩_ (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-∈⇿)
+ _⟨$⟩_ (Inverse.from map-with-∈↔)
where open P.≡-Reasoning
------------------------------------------------------------------------
diff --git a/src/Data/List/Any/Properties.agda b/src/Data/List/Any/Properties.agda
index 8ff99e6..7d9c9df 100644
--- a/src/Data/List/Any/Properties.agda
+++ b/src/Data/List/Any/Properties.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties related to Any
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- The other modules under Data.List.Any also contain properties
-- related to Any.
@@ -21,21 +21,22 @@ open import Data.Product as Prod hiding (swap)
open import Data.Sum as Sum using (_⊎_; inj₁; inj₂; [_,_]′)
open import Function
open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence
- using (_⇔_; equivalent; module Equivalent)
-open import Function.Inverse as Inv
- using (Isomorphism; _⇿_; module Inverse)
-open import Function.Inverse.TypeIsomorphisms
+open import Function.Equivalence as Eq using (_⇔_; module Equivalence)
+open import Function.Inverse as Inv using (_↔_; module Inverse)
+open import Function.Related as Related using (Related)
+open import Function.Related.TypeIsomorphisms
open import Level
open import Relation.Binary
import Relation.Binary.HeterogeneousEquality as H
+open import Relation.Binary.Product.Pointwise
open import Relation.Binary.PropositionalEquality as P
- using (_≡_; refl; inspect; _with-≡_)
+ using (_≡_; refl; inspect) renaming ([_] to P[_])
open import Relation.Unary using (_⟨×⟩_; _⟨→⟩_) renaming (_⊆_ to _⋐_)
import Relation.Binary.Sigma.Pointwise as Σ
+open import Relation.Binary.Sum
open Any.Membership-≡
-open Inv.EquationalReasoning
+open Related.EquationalReasoning
private
module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
open module ListMonad {ℓ} = RawMonad (List.monad {ℓ = ℓ})
@@ -102,9 +103,9 @@ private
-- Any can be expressed using _∈_.
-Any⇿ : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
- (∃ λ x → x ∈ xs × P x) ⇿ Any P xs
-Any⇿ {P = P} {xs} = record
+Any↔ : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
+ (∃ λ x → x ∈ xs × P x) ↔ Any P xs
+Any↔ {P = P} {xs} = record
{ to = P.→-to-⟶ to
; from = P.→-to-⟶ (find {P = P})
; inverse-of = record
@@ -121,41 +122,41 @@ Any⇿ {P = P} {xs} = record
-- Any is a congruence
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⇿ {P = P₁} ⟩
- (∃ λ x → x ∈ xs₁ × P₁ x) ≈⟨ Σ.cong Inv.id (xs₁≈xs₂ ⟨ ×⊎.*-cong ⟩ P₁⇿P₂ _) ⟩
- (∃ λ x → x ∈ xs₂ × P₂ x) ⇿⟨ Any⇿ {P = P₂} ⟩
+ (∀ x → Related k (P₁ x) (P₂ x)) → xs₁ ∼[ k ] xs₂ →
+ Related k (Any P₁ xs₁) (Any P₂ xs₂)
+Any-cong {P₁ = P₁} {P₂} {xs₁} {xs₂} P₁↔P₂ xs₁≈xs₂ =
+ Any P₁ xs₁ ↔⟨ sym $ Any↔ {P = P₁} ⟩
+ (∃ λ x → x ∈ xs₁ × P₁ x) ∼⟨ Σ.cong Inv.id (xs₁≈xs₂ ×-cong P₁↔P₂ _) ⟩
+ (∃ λ x → x ∈ xs₂ × P₂ x) ↔⟨ Any↔ {P = P₂} ⟩
Any P₂ xs₂ ∎
------------------------------------------------------------------------
-- Swapping
--- Nested occurrences of Any can sometimes be swapped. See also ×⇿.
+-- 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
+ 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⇿ {a = ℓ} {p = ℓ} ⟩
- (∃ λ x → x ∈ xs × Any (P x) ys) ⇿⟨ sym $ Σ.cong Inv.id (λ {x} → (x ∈ xs ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any⇿ {a = ℓ} {p = ℓ}) ⟩
- (∃ λ x → x ∈ xs × ∃ λ y → y ∈ ys × P x y) ⇿⟨ Σ.cong {a₁ = ℓ} Inv.id (∃∃⇿∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
- (∃₂ λ x y → x ∈ xs × y ∈ ys × P x y) ⇿⟨ ∃∃⇿∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _ ⟩
- (∃₂ λ y x → x ∈ xs × y ∈ ys × P x y) ⇿⟨ Σ.cong Inv.id (λ {y} → Σ.cong Inv.id (λ {x} →
- (x ∈ xs × y ∈ ys × P x y) ⇿⟨ sym $ ×⊎.*-assoc _ _ _ ⟩
- ((x ∈ xs × y ∈ ys) × P x y) ⇿⟨ ×⊎.*-comm (x ∈ xs) (y ∈ ys) ⟨ ×⊎.*-cong ⟩ (P x y ∎) ⟩
- ((y ∈ ys × x ∈ xs) × P x y) ⇿⟨ ×⊎.*-assoc _ _ _ ⟩
+ Any (λ x → Any (P x) ys) xs ↔⟨ sym $ Any↔ {a = ℓ} {p = ℓ} ⟩
+ (∃ λ x → x ∈ xs × Any (P x) ys) ↔⟨ sym $ Σ.cong Inv.id (λ {x} → (x ∈ xs ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any↔ {a = ℓ} {p = ℓ}) ⟩
+ (∃ λ x → x ∈ xs × ∃ λ y → y ∈ ys × P x y) ↔⟨ Σ.cong {a₁ = ℓ} Inv.id (∃∃↔∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
+ (∃₂ λ x y → x ∈ xs × y ∈ ys × P x y) ↔⟨ ∃∃↔∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _ ⟩
+ (∃₂ λ y x → x ∈ xs × y ∈ ys × P x y) ↔⟨ Σ.cong Inv.id (λ {y} → Σ.cong Inv.id (λ {x} →
+ (x ∈ xs × y ∈ ys × P x y) ↔⟨ sym $ ×⊎.*-assoc _ _ _ ⟩
+ ((x ∈ xs × y ∈ ys) × P x y) ↔⟨ ×⊎.*-comm (x ∈ xs) (y ∈ ys) ⟨ ×⊎.*-cong ⟩ (P x y ∎) ⟩
+ ((y ∈ ys × x ∈ xs) × P x y) ↔⟨ ×⊎.*-assoc _ _ _ ⟩
(y ∈ ys × x ∈ xs × P x y) ∎)) ⟩
- (∃₂ λ y x → y ∈ ys × x ∈ xs × P x y) ⇿⟨ Σ.cong {a₁ = ℓ} Inv.id (∃∃⇿∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
- (∃ λ y → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ⇿⟨ Σ.cong Inv.id (λ {y} → (y ∈ ys ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any⇿ {a = ℓ} {p = ℓ}) ⟩
- (∃ λ y → y ∈ ys × Any (flip P y) xs) ⇿⟨ Any⇿ {a = ℓ} {p = ℓ} ⟩
+ (∃₂ λ y x → y ∈ ys × x ∈ xs × P x y) ↔⟨ Σ.cong {a₁ = ℓ} Inv.id (∃∃↔∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
+ (∃ λ y → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ↔⟨ Σ.cong Inv.id (λ {y} → (y ∈ ys ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any↔ {a = ℓ} {p = ℓ}) ⟩
+ (∃ λ y → y ∈ ys × Any (flip P y) xs) ↔⟨ Any↔ {a = ℓ} {p = ℓ} ⟩
Any (λ y → Any (flip P y) xs) ys ∎
------------------------------------------------------------------------
-- Lemmas relating Any to ⊥
-⊥⇿Any⊥ : ∀ {a} {A : Set a} {xs : List A} → ⊥ ⇿ Any (const ⊥) xs
-⊥⇿Any⊥ {A = A} = record
+⊥↔Any⊥ : ∀ {a} {A : Set a} {xs : List A} → ⊥ ↔ Any (const ⊥) xs
+⊥↔Any⊥ {A = A} = record
{ to = P.→-to-⟶ (λ ())
; from = P.→-to-⟶ (λ p → from p)
; inverse-of = record
@@ -168,8 +169,8 @@ swap {ℓ} {P = P} {xs} {ys} =
from (here ())
from (there p) = from p
-⊥⇿Any[] : ∀ {a} {A : Set a} {P : A → Set} → ⊥ ⇿ Any P []
-⊥⇿Any[] = record
+⊥↔Any[] : ∀ {a} {A : Set a} {P : A → Set} → ⊥ ↔ Any P []
+⊥↔Any[] = record
{ to = P.→-to-⟶ (λ ())
; from = P.→-to-⟶ (λ ())
; inverse-of = record
@@ -183,9 +184,9 @@ swap {ℓ} {P = P} {xs} {ys} =
-- Sums commute with Any (for a fixed list).
-⊎⇿ : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} {xs} →
- (Any P xs ⊎ Any Q xs) ⇿ Any (λ x → P x ⊎ Q x) xs
-⊎⇿ {P = P} {Q} = record
+⊎↔ : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} {xs} →
+ (Any P xs ⊎ Any Q xs) ↔ Any (λ x → P x ⊎ Q x) xs
+⊎↔ {P = P} {Q} = record
{ to = P.→-to-⟶ to
; from = P.→-to-⟶ from
; inverse-of = record
@@ -218,10 +219,10 @@ swap {ℓ} {P = P} {xs} {ys} =
-- Products "commute" with Any.
-×⇿ : {A B : Set} {P : A → Set} {Q : B → Set}
+×↔ : {A B : Set} {P : A → Set} {Q : B → Set}
{xs : List A} {ys : List B} →
- (Any P xs × Any Q ys) ⇿ Any (λ x → Any (λ y → P x × Q y) ys) xs
-×⇿ {P = P} {Q} {xs} {ys} = record
+ (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
@@ -304,10 +305,10 @@ private
map⁻∘map⁺ P (here p) = refl
map⁻∘map⁺ P (there p) = P.cong there (map⁻∘map⁺ P p)
-map⇿ : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
+map↔ : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
{f : A → B} {xs} →
- Any (P ∘ f) xs ⇿ Any P (List.map f xs)
-map⇿ {P = P} {f = f} = record
+ Any (P ∘ f) xs ↔ Any P (List.map f xs)
+map↔ {P = P} {f = f} = record
{ to = P.→-to-⟶ $ map⁺ {P = P} {f = f}
; from = P.→-to-⟶ $ map⁻ {P = P} {f = f}
; inverse-of = record
@@ -354,9 +355,9 @@ private
++⁻∘++⁺ (x ∷ xs) {ys} (inj₁ (there p)) rewrite ++⁻∘++⁺ xs {ys} (inj₁ p) = refl
++⁻∘++⁺ (x ∷ xs) (inj₂ p) rewrite ++⁻∘++⁺ xs (inj₂ p) = refl
-++⇿ : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
- (Any P xs ⊎ Any P ys) ⇿ Any P (xs ++ ys)
-++⇿ {P = P} {xs = xs} = record
+++↔ : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
+ (Any P xs ⊎ Any P ys) ↔ Any P (xs ++ ys)
+++↔ {P = P} {xs = xs} = record
{ to = P.→-to-⟶ [ ++⁺ˡ {P = P}, ++⁺ʳ {P = P} xs ]′
; from = P.→-to-⟶ $ ++⁻ {P = P} xs
; inverse-of = record
@@ -388,9 +389,9 @@ private
return⁻ {P = P} (return⁺ p) ≡ p
return⁻∘return⁺ P p = refl
-return⇿ : ∀ {a p} {A : Set a} {P : A → Set p} {x} →
- P x ⇿ Any P (return x)
-return⇿ {P = P} = record
+return↔ : ∀ {a p} {A : Set a} {P : A → Set p} {x} →
+ P x ↔ Any P (return x)
+return↔ {P = P} = record
{ to = P.→-to-⟶ $ return⁺ {P = P}
; from = P.→-to-⟶ $ return⁻ {P = P}
; inverse-of = record
@@ -399,6 +400,15 @@ return⇿ {P = P} = record
}
}
+-- _∷_.
+
+∷↔ : ∀ {a p} {A : Set a} (P : A → Set p) {x xs} →
+ (P x ⊎ Any P xs) ↔ Any P (x ∷ xs)
+∷↔ P {x} {xs} =
+ (P x ⊎ Any P xs) ↔⟨ return↔ {P = P} ⊎-cong (Any P xs ∎) ⟩
+ (Any P [ x ] ⊎ Any P xs) ↔⟨ ++↔ {P = P} {xs = [ x ]} ⟩
+ Any P (x ∷ xs) ∎
+
-- concat.
private
@@ -445,9 +455,9 @@ private
rewrite concat⁻∘++⁺ʳ xs xss (concat⁺ p) =
P.cong there $ concat⁻∘concat⁺ p
-concat⇿ : ∀ {a p} {A : Set a} {P : A → Set p} {xss} →
- Any (Any P) xss ⇿ Any P (concat xss)
-concat⇿ {P = P} {xss = xss} = record
+concat↔ : ∀ {a p} {A : Set a} {P : A → Set p} {xss} →
+ Any (Any P) xss ↔ Any P (concat xss)
+concat↔ {P = P} {xss = xss} = record
{ to = P.→-to-⟶ $ concat⁺ {P = P}
; from = P.→-to-⟶ $ concat⁻ {P = P} xss
; inverse-of = record
@@ -458,22 +468,22 @@ concat⇿ {P = P} {xss = xss} = record
-- _>>=_.
->>=⇿ : ∀ {ℓ p} {A B : Set ℓ} {P : B → Set p} {xs} {f : A → List B} →
- Any (Any P ∘ f) xs ⇿ Any P (xs >>= f)
->>=⇿ {P = P} {xs} {f} =
- Any (Any P ∘ f) xs ⇿⟨ map⇿ {P = Any P} {f = f} ⟩
- Any (Any P) (List.map f xs) ⇿⟨ concat⇿ {P = P} ⟩
+>>=↔ : ∀ {ℓ p} {A B : Set ℓ} {P : B → Set p} {xs} {f : A → List B} →
+ Any (Any P ∘ f) xs ↔ Any P (xs >>= f)
+>>=↔ {P = P} {xs} {f} =
+ Any (Any P ∘ f) xs ↔⟨ map↔ {P = Any P} {f = f} ⟩
+ Any (Any P) (List.map f xs) ↔⟨ concat↔ {P = P} ⟩
Any P (xs >>= f) ∎
-- _⊛_.
-⊛⇿ : ∀ {ℓ} {A B : Set ℓ} {P : B → Set ℓ}
+⊛↔ : ∀ {ℓ} {A B : Set ℓ} {P : B → Set ℓ}
{fs : List (A → B)} {xs : List A} →
- Any (λ f → Any (P ∘ f) xs) fs ⇿ Any P (fs ⊛ xs)
-⊛⇿ {ℓ} {P = P} {fs} {xs} =
- Any (λ f → Any (P ∘ f) xs) fs ⇿⟨ Any-cong (λ _ → Any-cong (λ _ → return⇿ {a = ℓ} {p = ℓ}) (_ ∎)) (_ ∎) ⟩
- Any (λ f → Any (Any P ∘ return ∘ f) xs) fs ⇿⟨ Any-cong (λ _ → >>=⇿ {ℓ = ℓ} {p = ℓ}) (_ ∎) ⟩
- Any (λ f → Any P (xs >>= return ∘ f)) fs ⇿⟨ >>=⇿ {ℓ = ℓ} {p = ℓ} ⟩
+ 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↔ {a = ℓ} {p = ℓ}) (_ ∎)) (_ ∎) ⟩
+ Any (λ f → Any (Any P ∘ return ∘ f) xs) fs ↔⟨ Any-cong (λ _ → >>=↔ {ℓ = ℓ} {p = ℓ}) (_ ∎) ⟩
+ Any (λ f → Any P (xs >>= return ∘ f)) fs ↔⟨ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
Any P (fs ⊛ xs) ∎
-- An alternative introduction rule for _⊛_.
@@ -482,35 +492,35 @@ concat⇿ {P = P} {xss = xss} = record
{fs : List (A → B)} {xs} →
Any (P ⟨→⟩ Q) fs → Any P xs → Any Q (fs ⊛ xs)
⊛⁺′ {ℓ} pq p =
- Inverse.to (⊛⇿ {ℓ = ℓ}) ⟨$⟩
+ Inverse.to (⊛↔ {ℓ = ℓ}) ⟨$⟩
Any.map (λ pq → Any.map (λ {x} → pq {x}) p) pq
-- _⊗_.
-⊗⇿ : ∀ {ℓ} {A B : Set ℓ} {P : A × B → Set ℓ}
+⊗↔ : ∀ {ℓ} {A B : Set ℓ} {P : A × B → Set ℓ}
{xs : List A} {ys : List B} →
- Any (λ x → Any (λ y → P (x , y)) ys) xs ⇿ Any P (xs ⊗ ys)
-⊗⇿ {ℓ} {P = P} {xs} {ys} =
- Any (λ x → Any (λ y → P (x , y)) ys) xs ⇿⟨ return⇿ {a = ℓ} {p = ℓ} ⟩
- Any (λ _,_ → Any (λ x → Any (λ y → P (x , y)) ys) xs) (return _,_) ⇿⟨ ⊛⇿ ⟩
- Any (λ x, → Any (P ∘ x,) ys) (_,_ <$> xs) ⇿⟨ ⊛⇿ ⟩
+ 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↔ {a = ℓ} {p = ℓ} ⟩
+ Any (λ _,_ → Any (λ x → Any (λ y → P (x , y)) ys) xs) (return _,_) ↔⟨ ⊛↔ ⟩
+ Any (λ x, → Any (P ∘ x,) ys) (_,_ <$> xs) ↔⟨ ⊛↔ ⟩
Any P (xs ⊗ ys) ∎
-⊗⇿′ : {A B : Set} {P : A → Set} {Q : B → Set}
+⊗↔′ : {A B : Set} {P : A → Set} {Q : B → Set}
{xs : List A} {ys : List B} →
- (Any P xs × Any Q ys) ⇿ Any (P ⟨×⟩ Q) (xs ⊗ ys)
-⊗⇿′ {P = P} {Q} {xs} {ys} =
- (Any P xs × Any Q ys) ⇿⟨ ×⇿ ⟩
- Any (λ x → Any (λ y → P x × Q y) ys) xs ⇿⟨ ⊗⇿ ⟩
+ (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-∈⇿ :
+map-with-∈↔ :
∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p} {xs : List A}
{f : ∀ {x} → x ∈ xs → B} →
- (∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)) ⇿ Any P (map-with-∈ xs f)
-map-with-∈⇿ {A = A} {B} {P} = record
+ (∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)) ↔ Any P (map-with-∈ xs f)
+map-with-∈↔ {A = A} {B} {P} = record
{ to = P.→-to-⟶ (map-with-∈⁺ _)
; from = P.→-to-⟶ (map-with-∈⁻ _ _)
; inverse-of = record
@@ -560,7 +570,7 @@ private
any⁺ : ∀ {a} {A : Set a} (p : A → Bool) {xs} →
Any (T ∘ p) xs → T (any p xs)
- any⁺ p (here px) = Equivalent.from T-∨ ⟨$⟩ inj₁ px
+ any⁺ p (here px) = Equivalence.from T-∨ ⟨$⟩ inj₁ px
any⁺ p (there {x = x} pxs) with p x
... | true = _
... | false = any⁺ p pxs
@@ -568,15 +578,13 @@ private
any⁻ : ∀ {a} {A : Set a} (p : A → Bool) xs →
T (any p xs) → Any (T ∘ p) xs
any⁻ p [] ()
- any⁻ p (x ∷ xs) px∷xs with inspect (p x)
- 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⁻ p (x ∷ xs) px∷xs with p x | inspect p x
+ any⁻ p (x ∷ xs) px∷xs | true | P[ eq ] = here (Equivalence.from T-≡ ⟨$⟩ eq)
+ any⁻ p (x ∷ xs) px∷xs | false | _ = there (any⁻ p xs px∷xs)
any⇔ : ∀ {a} {A : Set a} {p : A → Bool} {xs} →
Any (T ∘ p) xs ⇔ T (any p xs)
-any⇔ = equivalent (any⁺ _) (any⁻ _ _)
+any⇔ = Eq.equivalence (any⁺ _) (any⁻ _ _)
------------------------------------------------------------------------
-- _++_ is commutative
@@ -604,9 +612,9 @@ private
rewrite ++⁻∘++⁺ ys {ys = xs} (inj₁ p)
| ++⁻∘++⁺ ys {ys = x ∷ xs} (inj₁ p) = P.refl
-++⇿++ : ∀ {a p} {A : Set a} {P : A → Set p} xs ys →
- Any P (xs ++ ys) ⇿ Any P (ys ++ xs)
-++⇿++ {P = P} xs ys = record
+++↔++ : ∀ {a p} {A : Set a} {P : A → Set p} xs ys →
+ Any P (xs ++ ys) ↔ Any P (ys ++ xs)
+++↔++ {P = P} xs ys = record
{ to = P.→-to-⟶ $ ++-comm {P = P} xs ys
; from = P.→-to-⟶ $ ++-comm {P = P} ys xs
; inverse-of = record
diff --git a/src/Data/List/Countdown.agda b/src/Data/List/Countdown.agda
index 3bbb50c..33266c5 100644
--- a/src/Data/List/Countdown.agda
+++ b/src/Data/List/Countdown.agda
@@ -1,12 +1,14 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- A data structure which keeps track of an upper bound on the number
-- of elements /not/ in a given list
------------------------------------------------------------------------
-open import Level
+import Level
open import Relation.Binary
-module Data.List.Countdown (D : DecSetoid zero zero) where
+module Data.List.Countdown (D : DecSetoid Level.zero Level.zero) where
open import Data.Empty
open import Data.Fin using (Fin; zero; suc)
@@ -129,7 +131,7 @@ private
------------------------------------------------------------------------
-- The countdown data structure
--- If counted ⊕_n is inhabited then there are at most n values of type
+-- If counted ⊕ n is inhabited then there are at most n values of type
-- Elem which are not members of counted (up to _≈_). You can read the
-- symbol _⊕_ as partitioning Elem into two parts: counted and
-- uncounted.
diff --git a/src/Data/List/NonEmpty.agda b/src/Data/List/NonEmpty.agda
index d21827c..6668ed8 100644
--- a/src/Data/List/NonEmpty.agda
+++ b/src/Data/List/NonEmpty.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Non-empty lists
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.List.NonEmpty where
open import Data.Product hiding (map)
diff --git a/src/Data/List/NonEmpty/Properties.agda b/src/Data/List/NonEmpty/Properties.agda
index 78d04cc..812018f 100644
--- a/src/Data/List/NonEmpty/Properties.agda
+++ b/src/Data/List/NonEmpty/Properties.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties of non-empty lists
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.List.NonEmpty.Properties where
open import Algebra
diff --git a/src/Data/List/Properties.agda b/src/Data/List/Properties.agda
index d17428c..8b7c72c 100644
--- a/src/Data/List/Properties.agda
+++ b/src/Data/List/Properties.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- List-related properties
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- Note that the lemmas below could be generalised to work with other
-- equalities than _≡_.
diff --git a/src/Data/List/Reverse.agda b/src/Data/List/Reverse.agda
index ef65ce0..1977ce1 100644
--- a/src/Data/List/Reverse.agda
+++ b/src/Data/List/Reverse.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Reverse view
------------------------------------------------------------------------
diff --git a/src/Data/Maybe.agda b/src/Data/Maybe.agda
index 5260a4e..4d590f3 100644
--- a/src/Data/Maybe.agda
+++ b/src/Data/Maybe.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- The Maybe type
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Maybe where
open import Level
@@ -39,6 +39,17 @@ maybe j n nothing = n
maybe′ : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → B → Maybe A → B
maybe′ = maybe
+-- A safe variant of "fromJust". If the value is nothing, then the
+-- return type is the unit type.
+
+From-just : ∀ {a} (A : Set a) → Maybe A → Set a
+From-just A (just _) = A
+From-just A nothing = Lift ⊤
+
+from-just : ∀ {a} {A : Set a} (x : Maybe A) → From-just A x
+from-just (just x) = x
+from-just nothing = _
+
------------------------------------------------------------------------
-- Maybe monad
diff --git a/src/Data/Maybe/Core.agda b/src/Data/Maybe/Core.agda
index 4837089..6110109 100644
--- a/src/Data/Maybe/Core.agda
+++ b/src/Data/Maybe/Core.agda
@@ -1,5 +1,6 @@
-{-# OPTIONS --universe-polymorphism #-}
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- The Maybe type
------------------------------------------------------------------------
@@ -12,3 +13,6 @@ open import Level
data Maybe {a} (A : Set a) : Set a where
just : (x : A) → Maybe A
nothing : Maybe A
+
+{-# IMPORT Data.FFI #-}
+{-# COMPILED_DATA Maybe Data.FFI.AgdaMaybe Just Nothing #-}
diff --git a/src/Data/Nat.agda b/src/Data/Nat.agda
index cfcd1e7..95eed21 100644
--- a/src/Data/Nat.agda
+++ b/src/Data/Nat.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Natural numbers
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Nat where
open import Function
@@ -12,7 +12,7 @@ open import Function.Injection
using (Injection; module Injection)
open import Data.Sum
open import Data.Empty
-open import Level using (zero)
+import Level
open import Relation.Nullary
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as PropEq
@@ -34,17 +34,17 @@ data ℕ : Set where
infix 4 _≤_ _<_ _≥_ _>_
-data _≤_ : Rel ℕ zero where
+data _≤_ : Rel ℕ Level.zero where
z≤n : ∀ {n} → zero ≤ n
s≤s : ∀ {m n} (m≤n : m ≤ n) → suc m ≤ suc n
-_<_ : Rel ℕ zero
+_<_ : Rel ℕ Level.zero
m < n = suc m ≤ n
-_≥_ : Rel ℕ zero
+_≥_ : Rel ℕ Level.zero
m ≥ n = n ≤ m
-_>_ : Rel ℕ zero
+_>_ : Rel ℕ Level.zero
m > n = n < m
-- The following, alternative definition of _≤_ is more suitable for
@@ -56,13 +56,13 @@ data _≤′_ (m : ℕ) : ℕ → Set where
≤′-refl : m ≤′ m
≤′-step : ∀ {n} (m≤′n : m ≤′ n) → m ≤′ suc n
-_<′_ : Rel ℕ zero
+_<′_ : Rel ℕ Level.zero
m <′ n = suc m ≤′ n
-_≥′_ : Rel ℕ zero
+_≥′_ : Rel ℕ Level.zero
m ≥′ n = n ≤′ m
-_>′_ : Rel ℕ zero
+_>′_ : Rel ℕ Level.zero
m >′ n = n <′ m
------------------------------------------------------------------------
@@ -166,7 +166,7 @@ suc m ≤? suc n with m ≤? n
-- A comparison view. Taken from "View from the left"
-- (McBride/McKinna); details may differ.
-data Ordering : Rel ℕ zero where
+data Ordering : Rel ℕ Level.zero where
less : ∀ m k → Ordering m (suc (m + k))
equal : ∀ m → Ordering m m
greater : ∀ m k → Ordering (suc (m + k)) m
diff --git a/src/Data/Nat/Coprimality.agda b/src/Data/Nat/Coprimality.agda
index 45853cf..6420079 100644
--- a/src/Data/Nat/Coprimality.agda
+++ b/src/Data/Nat/Coprimality.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Coprimality
------------------------------------------------------------------------
@@ -65,7 +67,7 @@ sym c = c ∘ swap
-- Everything is coprime to 1.
1-coprimeTo : ∀ m → Coprime 1 m
-1-coprimeTo m = 1∣1 ∘ proj₁
+1-coprimeTo m = ∣1⇒≡1 ∘ proj₁
-- Nothing except for 1 is coprime to 0.
diff --git a/src/Data/Nat/DivMod.agda b/src/Data/Nat/DivMod.agda
index 76d6d9a..6342ac1 100644
--- a/src/Data/Nat/DivMod.agda
+++ b/src/Data/Nat/DivMod.agda
@@ -1,10 +1,12 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Integer division
------------------------------------------------------------------------
module Data.Nat.DivMod where
-open import Data.Nat
+open import Data.Nat as Nat
open import Data.Nat.Properties
open SemiringSolver
open import Data.Fin as Fin using (Fin; zero; suc; toℕ; fromℕ)
@@ -20,7 +22,8 @@ open import Function
private
- lem₁ : ∀ m k → _
+ lem₁ : (m k : ℕ) →
+ Nat.suc m ≡ suc (toℕ (Fin.inject+ k (fromℕ m)) + 0)
lem₁ m k = cong suc $ begin
m
≡⟨ sym $ Fin.to-from m ⟩
@@ -34,7 +37,7 @@ private
lem₂ : ∀ n → _
lem₂ = solve 1 (λ n → con 1 :+ n := con 1 :+ (n :+ con 0)) refl
- lem₃ : ∀ n k q r eq → _
+ lem₃ : ∀ n k q (r : Fin n) eq → suc n + k ≡ toℕ r + suc q * n
lem₃ n k q r eq = begin
suc n + k
≡⟨ solve 2 (λ n k → con 1 :+ n :+ k := n :+ (con 1 :+ k))
diff --git a/src/Data/Nat/Divisibility.agda b/src/Data/Nat/Divisibility.agda
index 262e920..c2b1c82 100644
--- a/src/Data/Nat/Divisibility.agda
+++ b/src/Data/Nat/Divisibility.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Divisibility
------------------------------------------------------------------------
@@ -103,13 +105,13 @@ n ∣0 = divides 0 refl
-- 0 only divides 0.
-0∣0 : ∀ {n} → 0 ∣ n → n ≡ 0
-0∣0 {n} 0∣n = P.antisym (n ∣0) 0∣n
+0∣⇒≡0 : ∀ {n} → 0 ∣ n → n ≡ 0
+0∣⇒≡0 {n} 0∣n = P.antisym (n ∣0) 0∣n
-- Only 1 divides 1.
-1∣1 : ∀ {n} → n ∣ 1 → n ≡ 1
-1∣1 {n} n∣1 = P.antisym n∣1 (1∣ n)
+∣1⇒≡1 : ∀ {n} → n ∣ 1 → n ≡ 1
+∣1⇒≡1 {n} n∣1 = P.antisym n∣1 (1∣ n)
-- If i divides m and n, then i divides their sum.
@@ -196,7 +198,7 @@ nonZeroDivisor-lemma m (suc q) r r≢zero d =
_∣?_ : Decidable _∣_
zero ∣? zero = yes (0 ∣0)
-zero ∣? suc n = no ((λ ()) ∘′ 0∣0)
+zero ∣? suc n = no ((λ ()) ∘′ 0∣⇒≡0)
suc m ∣? n with n divMod suc m
suc m ∣? .(q * suc m) | result q zero =
yes $ divides q refl
diff --git a/src/Data/Nat/GCD.agda b/src/Data/Nat/GCD.agda
index 29d8bad..de7382c 100644
--- a/src/Data/Nat/GCD.agda
+++ b/src/Data/Nat/GCD.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Greatest common divisor
------------------------------------------------------------------------
diff --git a/src/Data/Nat/GCD/Lemmas.agda b/src/Data/Nat/GCD/Lemmas.agda
index ddef011..8b5a6c0 100644
--- a/src/Data/Nat/GCD/Lemmas.agda
+++ b/src/Data/Nat/GCD/Lemmas.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Boring lemmas used in Data.Nat.GCD and Data.Nat.Coprimality
------------------------------------------------------------------------
diff --git a/src/Data/Nat/InfinitelyOften.agda b/src/Data/Nat/InfinitelyOften.agda
index efafabd..f5c552a 100644
--- a/src/Data/Nat/InfinitelyOften.agda
+++ b/src/Data/Nat/InfinitelyOften.agda
@@ -1,9 +1,12 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Definition of and lemmas related to "true infinitely often"
------------------------------------------------------------------------
module Data.Nat.InfinitelyOften where
+import Level
open import Algebra
open import Category.Monad
open import Data.Empty
@@ -16,7 +19,7 @@ open import Relation.Binary.PropositionalEquality
open import Relation.Nullary
open import Relation.Nullary.Negation
open import Relation.Unary using (_∪_; _⊆_)
-open RawMonad ¬¬-Monad
+open RawMonad (¬¬-Monad {p = Level.zero})
private
module NatLattice = DistributiveLattice NatProp.distributiveLattice
diff --git a/src/Data/Nat/LCM.agda b/src/Data/Nat/LCM.agda
index 4e41e8a..36abcc6 100644
--- a/src/Data/Nat/LCM.agda
+++ b/src/Data/Nat/LCM.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Least common multiple
------------------------------------------------------------------------
@@ -13,7 +15,7 @@ open import Data.Nat.Coprimality as Coprime
open import Data.Product
open import Function
open import Relation.Binary.PropositionalEquality as PropEq
- using (_≡_; refl; inspect; _with-≡_)
+ using (_≡_; refl)
open import Algebra
open import Relation.Binary
private
diff --git a/src/Data/Nat/Primality.agda b/src/Data/Nat/Primality.agda
new file mode 100644
index 0000000..3ea1faf
--- /dev/null
+++ b/src/Data/Nat/Primality.agda
@@ -0,0 +1,37 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Primality
+------------------------------------------------------------------------
+
+module Data.Nat.Primality where
+
+open import Data.Empty
+open import Data.Fin as Fin hiding (_+_)
+open import Data.Fin.Dec
+open import Data.Nat
+open import Data.Nat.Divisibility
+open import Relation.Nullary
+open import Relation.Nullary.Decidable
+open import Relation.Nullary.Negation
+
+-- Definition of primality.
+
+Prime : ℕ → Set
+Prime 0 = ⊥
+Prime 1 = ⊥
+Prime (suc (suc n)) = (i : Fin n) → ¬ (2 + Fin.toℕ i ∣ 2 + n)
+
+-- Decision procedure for primality.
+
+prime? : ∀ n → Dec (Prime n)
+prime? 0 = no λ()
+prime? 1 = no λ()
+prime? (suc (suc n)) = all? λ _ → ¬? (_ ∣? _)
+
+private
+
+ -- Example: 2 is prime.
+
+ 2-is-prime : Prime 2
+ 2-is-prime = from-yes (prime? 2)
diff --git a/src/Data/Nat/Properties.agda b/src/Data/Nat/Properties.agda
index 4e22dd7..4682ca0 100644
--- a/src/Data/Nat/Properties.agda
+++ b/src/Data/Nat/Properties.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- A bunch of properties about natural number operations
------------------------------------------------------------------------
@@ -19,7 +21,7 @@ open import Relation.Nullary
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; _≢_; refl; sym; cong; cong₂)
open PropEq.≡-Reasoning
-import Algebra.FunctionProperties as P; open P _≡_
+import Algebra.FunctionProperties as P; open P (_≡_ {A = ℕ})
open import Data.Product
------------------------------------------------------------------------
diff --git a/src/Data/Nat/Show.agda b/src/Data/Nat/Show.agda
index 65a31a1..7ecb490 100644
--- a/src/Data/Nat/Show.agda
+++ b/src/Data/Nat/Show.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Showing natural numbers
------------------------------------------------------------------------
@@ -18,11 +20,11 @@ showInBase : (base : ℕ)
{base≥2 : True (2 ≤? base)}
{base≤16 : True (base ≤? 16)} →
ℕ → String
-showInBase base {base≥2} {base≤16} =
- String.fromList ∘
- map (showDigit {base≤16 = base≤16}) ∘
- reverse ∘
- theDigits base {base≥2 = base≥2}
+showInBase base {base≥2} {base≤16} n =
+ String.fromList $
+ map (showDigit {base≤16 = base≤16}) $
+ reverse $
+ theDigits base {base≥2 = base≥2} n
-- show n is a string containing the decimal expansion of n (starting
-- with the most significant digit).
diff --git a/src/Data/Plus.agda b/src/Data/Plus.agda
index c2b1b95..b70d648 100644
--- a/src/Data/Plus.agda
+++ b/src/Data/Plus.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Transitive closures
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Plus where
open import Function
@@ -73,7 +73,7 @@ _++_ : ∀ {a ℓ} {A : Set a} {_∼_ : Rel A ℓ} {x y z} →
equivalent : ∀ {a ℓ} {A : Set a} {_∼_ : Rel A ℓ} {x y} →
Plus _∼_ x y ⇔ Plus′ _∼_ x y
-equivalent {_∼_ = _∼_} = Equiv.equivalent complete sound
+equivalent {_∼_ = _∼_} = Equiv.equivalence complete sound
where
complete : Plus _∼_ ⇒ Plus′ _∼_
complete [ x∼y ] = [ x∼y ]
diff --git a/src/Data/Product.agda b/src/Data/Product.agda
index 0ce4716..8164500 100644
--- a/src/Data/Product.agda
+++ b/src/Data/Product.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Products
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Product where
open import Function
diff --git a/src/Data/Product/N-ary.agda b/src/Data/Product/N-ary.agda
index 83c4dd7..b8b4e28 100644
--- a/src/Data/Product/N-ary.agda
+++ b/src/Data/Product/N-ary.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- N-ary products
------------------------------------------------------------------------
@@ -8,8 +10,6 @@
-- avoided: pairs are represented as pairs (x , y), not as triples
-- (x , y , unit).
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Product.N-ary where
open import Data.Nat
@@ -17,7 +17,7 @@ open import Data.Product
open import Data.Unit
open import Data.Vec
open import Function.Inverse
-open import Level
+open import Level using (Lift; lift)
open import Relation.Binary.PropositionalEquality as P using (_≡_)
-- N-ary product.
@@ -31,8 +31,8 @@ A ^ suc (suc n) = A × A ^ suc n
-- Conversions.
-⇿Vec : ∀ {a} {A : Set a} n → A ^ n ⇿ Vec A n
-⇿Vec n = record
+↔Vec : ∀ {a} {A : Set a} n → A ^ n ↔ Vec A n
+↔Vec n = record
{ to = P.→-to-⟶ (toVec n)
; from = P.→-to-⟶ fromVec
; inverse-of = record
diff --git a/src/Data/Rational.agda b/src/Data/Rational.agda
index 0f11fb3..0d49196 100644
--- a/src/Data/Rational.agda
+++ b/src/Data/Rational.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Rational numbers
------------------------------------------------------------------------
@@ -12,7 +14,7 @@ import Data.Integer.Properties as ℤ
open import Data.Nat.Divisibility as ℕDiv using (_∣_)
import Data.Nat.Coprimality as C
open import Data.Nat as ℕ renaming (_*_ to _ℕ*_)
-open import Level
+import Level
open import Relation.Nullary.Decidable
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as PropEq
@@ -69,7 +71,7 @@ private
infix 4 _≃_
-_≃_ : Rel ℚ zero
+_≃_ : Rel ℚ Level.zero
p ≃ q = P.numerator ℤ* Q.denominator ≡
Q.numerator ℤ* P.denominator
where module P = ℚ p; module Q = ℚ q
diff --git a/src/Data/ReflexiveClosure.agda b/src/Data/ReflexiveClosure.agda
index 1fabe8f..cf12175 100644
--- a/src/Data/ReflexiveClosure.agda
+++ b/src/Data/ReflexiveClosure.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Reflexive closures
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.ReflexiveClosure where
open import Data.Unit
diff --git a/src/Data/Sign.agda b/src/Data/Sign.agda
index 33c234c..65ef14f 100644
--- a/src/Data/Sign.agda
+++ b/src/Data/Sign.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Signs
------------------------------------------------------------------------
diff --git a/src/Data/Sign/Properties.agda b/src/Data/Sign/Properties.agda
index b143bc3..55f6278 100644
--- a/src/Data/Sign/Properties.agda
+++ b/src/Data/Sign/Properties.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some properties about signs
------------------------------------------------------------------------
diff --git a/src/Data/Star.agda b/src/Data/Star.agda
index a3b3deb..b85e442 100644
--- a/src/Data/Star.agda
+++ b/src/Data/Star.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- The reflexive transitive closures of McBride, Norell and Jansson
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- This module could be placed under Relation.Binary. However, since
-- its primary purpose is to be used for _data_ it has been placed
-- under Data instead.
diff --git a/src/Data/Star/BoundedVec.agda b/src/Data/Star/BoundedVec.agda
index 454007f..0d78b53 100644
--- a/src/Data/Star/BoundedVec.agda
+++ b/src/Data/Star/BoundedVec.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Bounded vectors (inefficient implementation)
------------------------------------------------------------------------
diff --git a/src/Data/Star/Decoration.agda b/src/Data/Star/Decoration.agda
index 8a831da..0e954cc 100644
--- a/src/Data/Star/Decoration.agda
+++ b/src/Data/Star/Decoration.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Decorated star-lists
------------------------------------------------------------------------
@@ -12,7 +14,7 @@ open import Level
-- A predicate on relation "edges" (think of the relation as a graph).
-EdgePred : ∀ {I} → Rel I zero → Set₁
+EdgePred : {I : Set} → Rel I zero → Set₁
EdgePred T = ∀ {i j} → T i j → Set
data NonEmptyEdgePred {I : Set}
diff --git a/src/Data/Star/Environment.agda b/src/Data/Star/Environment.agda
index 9537348..826d8e7 100644
--- a/src/Data/Star/Environment.agda
+++ b/src/Data/Star/Environment.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Environments (heterogeneous collections)
------------------------------------------------------------------------
diff --git a/src/Data/Star/Fin.agda b/src/Data/Star/Fin.agda
index 401d88b..add8e50 100644
--- a/src/Data/Star/Fin.agda
+++ b/src/Data/Star/Fin.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Finite sets defined in terms of Data.Star
------------------------------------------------------------------------
diff --git a/src/Data/Star/List.agda b/src/Data/Star/List.agda
index 3cbff46..54aeb1a 100644
--- a/src/Data/Star/List.agda
+++ b/src/Data/Star/List.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Lists defined in terms of Data.Star
------------------------------------------------------------------------
diff --git a/src/Data/Star/Nat.agda b/src/Data/Star/Nat.agda
index b012ce7..a54a465 100644
--- a/src/Data/Star/Nat.agda
+++ b/src/Data/Star/Nat.agda
@@ -1,16 +1,16 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Natural numbers defined in terms of Data.Star
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Star.Nat where
open import Data.Star
open import Data.Unit
+open import Function
open import Relation.Binary
open import Relation.Binary.Simple
-open import Function
-- Natural numbers.
@@ -23,12 +23,12 @@ zero : ℕ
zero = ε
suc : ℕ → ℕ
-suc = _◅_ tt
+suc = _◅_ _
-- The length of a star-list.
length : ∀ {i t} {I : Set i} {T : Rel I t} {i j} → Star T i j → ℕ
-length = gmap (const tt) (const tt)
+length = gmap (const _) (const _)
-- Arithmetic.
diff --git a/src/Data/Star/Pointer.agda b/src/Data/Star/Pointer.agda
index ccddcca..610b155 100644
--- a/src/Data/Star/Pointer.agda
+++ b/src/Data/Star/Pointer.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Pointers into star-lists
------------------------------------------------------------------------
diff --git a/src/Data/Star/Properties.agda b/src/Data/Star/Properties.agda
index 2294a53..ccc3cb2 100644
--- a/src/Data/Star/Properties.agda
+++ b/src/Data/Star/Properties.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some properties related to Data.Star
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Star.Properties where
open import Data.Star
diff --git a/src/Data/Star/Vec.agda b/src/Data/Star/Vec.agda
index 7f03db9..e7e0a53 100644
--- a/src/Data/Star/Vec.agda
+++ b/src/Data/Star/Vec.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Vectors defined in terms of Data.Star
------------------------------------------------------------------------
diff --git a/src/Data/Stream.agda b/src/Data/Stream.agda
index ea6a270..bf76626 100644
--- a/src/Data/Stream.agda
+++ b/src/Data/Stream.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Streams
------------------------------------------------------------------------
@@ -18,23 +20,26 @@ infixr 5 _∷_
data Stream (A : Set) : Set where
_∷_ : (x : A) (xs : ∞ (Stream A)) → Stream A
+{-# IMPORT Data.FFI #-}
+{-# COMPILED_DATA Stream Data.FFI.AgdaStream Data.FFI.Cons #-}
+
------------------------------------------------------------------------
-- Some operations
-head : forall {A} → Stream A → A
+head : ∀ {A} → Stream A → A
head (x ∷ xs) = x
-tail : forall {A} → Stream A → Stream A
+tail : ∀ {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} →
+zipWith : ∀ {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 : ∀ {A} n → Stream A → Vec A n
take zero xs = []
take (suc n) (x ∷ xs) = x ∷ take n (♭ xs)
@@ -42,7 +47,7 @@ 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 : ∀ {A} → A → Stream A
repeat x = x ∷ ♯ repeat x
iterate : ∀ {A} → (A → A) → A → Stream A
@@ -60,13 +65,13 @@ mutual
-- Takes every other element from the stream, starting with the
-- first one.
- evens : {A : Set} → Stream A → Stream A
+ evens : ∀ {A} → 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 : ∀ {A} → Stream A → Stream A
odds (x ∷ xs) = evens (♭ xs)
toColist : ∀ {A} → Stream A → Colist A
@@ -89,14 +94,14 @@ _++_ : ∀ {A} → Colist A → Stream A → Stream A
infix 4 _≈_
-data _≈_ {A} : (xs ys : Stream A) → Set where
+data _≈_ {A} : Stream A → Stream A → Set where
_∷_ : ∀ x {xs ys} (xs≈ : ∞ (♭ xs ≈ ♭ ys)) → x ∷ xs ≈ x ∷ ys
-- x ∈ xs means that x is a member of xs.
infix 4 _∈_
-data _∈_ {A : Set} : A → Stream A → Set where
+data _∈_ {A} : A → Stream A → Set where
here : ∀ {x xs} → x ∈ x ∷ xs
there : ∀ {x y xs} (x∈xs : x ∈ ♭ xs) → x ∈ y ∷ xs
@@ -104,7 +109,7 @@ data _∈_ {A : Set} : A → Stream A → Set where
infix 4 _⊑_
-data _⊑_ {A : Set} : Colist A → Stream A → Set where
+data _⊑_ {A} : Colist A → Stream A → Set where
[] : ∀ {ys} → [] ⊑ ys
_∷_ : ∀ x {xs ys} (p : ∞ (♭ xs ⊑ ♭ ys)) → x ∷ xs ⊑ x ∷ ys
@@ -123,7 +128,7 @@ setoid A = record
}
where
refl : Reflexive _≈_
- refl {x ∷ xs} = x ∷ ♯ refl
+ refl {x ∷ _} = x ∷ ♯ refl
sym : Symmetric _≈_
sym (x ∷ xs≈) = x ∷ ♯ sym (♭ xs≈)
@@ -131,7 +136,7 @@ setoid A = record
trans : Transitive _≈_
trans (x ∷ xs≈) (.x ∷ ys≈) = x ∷ ♯ trans (♭ xs≈) (♭ ys≈)
-map-cong : ∀ {A B} (f : A → B) {xs ys : Stream A} →
+map-cong : ∀ {A B} (f : A → B) {xs ys} →
xs ≈ ys → map f xs ≈ map f ys
map-cong f (x ∷ xs≈) = f x ∷ ♯ map-cong f (♭ xs≈)
diff --git a/src/Data/String.agda b/src/Data/String.agda
index 41ba5e4..acf9696 100644
--- a/src/Data/String.agda
+++ b/src/Data/String.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Strings
------------------------------------------------------------------------
diff --git a/src/Data/Sum.agda b/src/Data/Sum.agda
index 6c54002..0313256 100644
--- a/src/Data/Sum.agda
+++ b/src/Data/Sum.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Sums (disjoint unions)
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Sum where
open import Function
@@ -19,6 +19,9 @@ data _⊎_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where
inj₁ : (x : A) → A ⊎ B
inj₂ : (y : B) → A ⊎ B
+{-# IMPORT Data.FFI #-}
+{-# COMPILED_DATA _⊎_ Data.FFI.AgdaEither Left Right #-}
+
------------------------------------------------------------------------
-- Functions
diff --git a/src/Data/Unit.agda b/src/Data/Unit.agda
index 92fc219..7e690b5 100644
--- a/src/Data/Unit.agda
+++ b/src/Data/Unit.agda
@@ -1,5 +1,7 @@
------------------------------------------------------------------------
--- The unit type
+-- The Agda standard library
+--
+-- Some unit types
------------------------------------------------------------------------
module Data.Unit where
@@ -10,8 +12,12 @@ open import Relation.Binary
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; refl)
+-- Some types and operations are defined in Data.Unit.Core.
+
+open import Data.Unit.Core public
+
------------------------------------------------------------------------
--- Types
+-- A unit type defined as a record type
-- Note that the name of this type is "\top", not T.
diff --git a/src/Data/Unit/Core.agda b/src/Data/Unit/Core.agda
new file mode 100644
index 0000000..a4ddcd5
--- /dev/null
+++ b/src/Data/Unit/Core.agda
@@ -0,0 +1,42 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Some unit types
+------------------------------------------------------------------------
+
+-- The definitions in this file are reexported by Data.Unit.
+
+module Data.Unit.Core where
+
+open import Level
+
+------------------------------------------------------------------------
+-- A unit type defined as a data-type
+
+-- The ⊤ type (see Data.Unit) comes with η-equality, which is often
+-- nice to have, but sometimes it is convenient to be able to stop
+-- unfolding (see "Hidden types" below).
+
+data Unit : Set where
+ unit : Unit
+
+------------------------------------------------------------------------
+-- Hidden types
+
+-- "Hidden" values.
+
+Hidden : ∀ {a} → Set a → Set a
+Hidden A = Unit → A
+
+-- The hide function can be used to hide function applications. Note
+-- that the type-checker doesn't see that "hide f x" contains the
+-- application "f x".
+
+hide : ∀ {a b} {A : Set a} {B : A → Set b} →
+ ((x : A) → B x) → ((x : A) → Hidden (B x))
+hide f x unit = f x
+
+-- Reveals a hidden value.
+
+reveal : ∀ {a} {A : Set a} → Hidden A → A
+reveal f = f unit
diff --git a/src/Data/Vec.agda b/src/Data/Vec.agda
index 3181cdb..c2a55cb 100644
--- a/src/Data/Vec.agda
+++ b/src/Data/Vec.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Vectors
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Vec where
open import Category.Applicative
@@ -12,8 +12,7 @@ open import Data.Fin using (Fin; zero; suc)
open import Data.List as List using (List)
open import Data.Product using (∃; ∃₂; _×_; _,_)
open import Function
-open import Level
-open import Relation.Binary.PropositionalEquality
+open import Relation.Binary.PropositionalEquality using (_≡_; refl)
------------------------------------------------------------------------
-- Types
diff --git a/src/Data/Vec/Equality.agda b/src/Data/Vec/Equality.agda
index 0892021..19544d3 100644
--- a/src/Data/Vec/Equality.agda
+++ b/src/Data/Vec/Equality.agda
@@ -1,17 +1,17 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Semi-heterogeneous vector equality
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Vec.Equality where
open import Data.Vec
open import Data.Nat using (suc)
open import Function
-open import Level
+open import Level using (_⊔_)
open import Relation.Binary
-open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
module Equality {s₁ s₂} (S : Setoid s₁ s₂) where
@@ -32,8 +32,8 @@ module Equality {s₁ s₂} (S : Setoid s₁ s₂) where
length-equal : ∀ {n¹} {xs¹ : Vec A n¹}
{n²} {xs² : Vec A n²} →
xs¹ ≈ xs² → n¹ ≡ n²
- length-equal []-cong = PropEq.refl
- length-equal (_ ∷-cong eq₂) = PropEq.cong suc $ length-equal eq₂
+ length-equal []-cong = P.refl
+ length-equal (_ ∷-cong eq₂) = P.cong suc $ length-equal eq₂
refl : ∀ {n} (xs : Vec A n) → xs ≈ xs
refl [] = []-cong
@@ -79,21 +79,13 @@ module DecidableEquality {d₁ d₂} (D : DecSetoid d₁ d₂) where
helper : ¬ (x ∷ xs ≈ y ∷ ys)
helper (x≊y ∷-cong _) = ¬x≊y x≊y
-module HeterogeneousEquality {a} {A : Set a} where
+module PropositionalEquality {a} {A : Set a} where
- open import Relation.Binary.HeterogeneousEquality as HetEq
- using (_≅_)
- open Equality (PropEq.setoid A)
+ open Equality (P.setoid A)
- to-≅ : ∀ {n m} {xs : Vec A n} {ys : Vec A m} →
- xs ≈ ys → xs ≅ ys
- to-≅ []-cong = HetEq.refl
- to-≅ (PropEq.refl ∷-cong xs¹≈xs²) with length-equal xs¹≈xs²
- ... | PropEq.refl = HetEq.cong (_∷_ _) $ to-≅ xs¹≈xs²
+ to-≡ : ∀ {n} {xs ys : Vec A n} → xs ≈ ys → xs ≡ ys
+ to-≡ []-cong = P.refl
+ to-≡ (P.refl ∷-cong xs¹≈xs²) = P.cong (_∷_ _) $ to-≡ xs¹≈xs²
- from-≅ : ∀ {n m} {xs : Vec A n} {ys : Vec A m} →
- xs ≅ ys → xs ≈ ys
- from-≅ {xs = []} {ys = []} _ = refl _
- from-≅ {xs = x ∷ xs} {ys = .x ∷ .xs} HetEq.refl = refl _
- from-≅ {xs = _ ∷ _} {ys = []} ()
- from-≅ {xs = []} {ys = _ ∷ _} ()
+ from-≡ : ∀ {n} {xs ys : Vec A n} → xs ≡ ys → xs ≈ ys
+ from-≡ P.refl = refl _
diff --git a/src/Data/Vec/N-ary.agda b/src/Data/Vec/N-ary.agda
index 0139ccf..827a9a8 100644
--- a/src/Data/Vec/N-ary.agda
+++ b/src/Data/Vec/N-ary.agda
@@ -1,17 +1,17 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Code for converting Vec A n → B to and from n-ary functions
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Vec.N-ary where
open import Data.Nat hiding (_⊔_)
open import Data.Product as Prod
open import Data.Vec
open import Function
-open import Function.Equivalence using (_⇔_; equivalent)
-open import Level
+open import Function.Equivalence using (_⇔_; equivalence)
+open import Level using (Level; _⊔_)
open import Relation.Binary
open import Relation.Binary.PropositionalEquality
open import Relation.Nullary.Decidable
@@ -95,7 +95,7 @@ right-inverse (suc n) f = λ x → right-inverse n (f x)
uncurry-∀ⁿ : ∀ n {a ℓ} {A : Set a} {P : N-ary n A (Set ℓ)} →
∀ⁿ n P ⇔ (∀ (xs : Vec A n) → P $ⁿ xs)
-uncurry-∀ⁿ n {a} {ℓ} {A} = equivalent (⇒ n) (⇐ n)
+uncurry-∀ⁿ n {a} {ℓ} {A} = equivalence (⇒ n) (⇐ n)
where
⇒ : ∀ n {P : N-ary n A (Set ℓ)} →
∀ⁿ n P → (∀ (xs : Vec A n) → P $ⁿ xs)
@@ -111,7 +111,7 @@ uncurry-∀ⁿ n {a} {ℓ} {A} = equivalent (⇒ n) (⇐ n)
uncurry-∃ⁿ : ∀ n {a ℓ} {A : Set a} {P : N-ary n A (Set ℓ)} →
∃ⁿ n P ⇔ (∃ λ (xs : Vec A n) → P $ⁿ xs)
-uncurry-∃ⁿ n {a} {ℓ} {A} = equivalent (⇒ n) (⇐ n)
+uncurry-∃ⁿ n {a} {ℓ} {A} = equivalence (⇒ n) (⇐ n)
where
⇒ : ∀ n {P : N-ary n A (Set ℓ)} →
∃ⁿ n P → (∃ λ (xs : Vec A n) → P $ⁿ xs)
diff --git a/src/Data/Vec/Properties.agda b/src/Data/Vec/Properties.agda
index 1fa6d08..674b973 100644
--- a/src/Data/Vec/Properties.agda
+++ b/src/Data/Vec/Properties.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some Vec-related properties
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.Vec.Properties where
open import Algebra
@@ -16,8 +16,7 @@ import Data.Nat.Properties as Nat
open import Data.Fin as Fin using (Fin; zero; suc; toℕ; fromℕ)
open import Data.Fin.Props using (_+′_)
open import Function
-open import Function.Inverse using (_⇿_)
-open import Level
+open import Function.Inverse using (_↔_)
open import Relation.Binary
module UsingVectorEquality {s₁ s₂} (S : Setoid s₁ s₂) where
@@ -257,9 +256,9 @@ proof-irrelevance-[]= (there xs[i]=x) (there xs[i]=x') =
-- _[_]=_ can be expressed using lookup and _≡_.
-[]=⇿lookup : ∀ {a n i} {A : Set a} {x} {xs : Vec A n} →
- xs [ i ]= x ⇿ lookup i xs ≡ x
-[]=⇿lookup {i = i} {x = x} {xs} = record
+[]=↔lookup : ∀ {a n i} {A : Set a} {x} {xs : Vec A n} →
+ xs [ i ]= x ↔ lookup i xs ≡ x
+[]=↔lookup {i = i} {x = x} {xs} = record
{ to = PropEq.→-to-⟶ to
; from = PropEq.→-to-⟶ (from i xs)
; inverse-of = record
diff --git a/src/Data/W.agda b/src/Data/W.agda
index 5e68a9d..4d47f7b 100644
--- a/src/Data/W.agda
+++ b/src/Data/W.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- W-types
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Data.W where
open import Level
diff --git a/src/Foreign/Haskell.agda b/src/Foreign/Haskell.agda
index d3ebae4..e0c462a 100644
--- a/src/Foreign/Haskell.agda
+++ b/src/Foreign/Haskell.agda
@@ -1,42 +1,14 @@
------------------------------------------------------------------------
--- Types used (only) when calling out to Haskell via the FFI
+-- The Agda standard library
+--
+-- Type(s) used (only) when calling out to Haskell via the FFI
------------------------------------------------------------------------
module Foreign.Haskell where
-open import Coinduction
-open import Data.Char using (Char)
-open import Data.Colist as C using ([]; _∷_)
-open import Function using (_∘_)
-open import Data.String as String using (String)
-
-------------------------------------------------------------------------
--- Simple types
-
-- A unit type.
data Unit : Set where
unit : Unit
{-# COMPILED_DATA Unit () () #-}
-
--- Potentially infinite lists.
-
-infixr 5 _∷_
-
-data Colist (A : Set) : Set where
- [] : 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)
-
-toColist : ∀ {A} → Colist A → C.Colist A
-toColist [] = []
-toColist (x ∷ xs) = x ∷ ♯ toColist (♭ xs)
-
-fromString : String → Colist Char
-fromString = fromColist ∘ String.toCostring
diff --git a/src/Function.agda b/src/Function.agda
index 0f1187e..5bdf646 100644
--- a/src/Function.agda
+++ b/src/Function.agda
@@ -1,14 +1,15 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Simple combinators working solely on and with functions
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Function where
open import Level
infixr 9 _∘_ _∘′_
+infixl 8 _ˢ_
infixl 1 _on_
infixl 1 _⟨_⟩_
infixr 0 _-[_]-_ _$_
@@ -46,9 +47,20 @@ id x = x
const : ∀ {a b} {A : Set a} {B : Set b} → A → B → A
const x = λ _ → x
+-- The S combinator. (Written infix as in Conor McBride's paper
+-- "Outrageous but Meaningful Coincidences: Dependent type-safe syntax
+-- and evaluation".)
+
+_ˢ_ : ∀ {a b c}
+ {A : Set a} {B : A → Set b} {C : (x : A) → B x → Set c} →
+ ((x : A) (y : B x) → C x y) →
+ (g : (x : A) → B x) →
+ ((x : A) → C x (g x))
+f ˢ g = λ x → f x (g x)
+
flip : ∀ {a b c} {A : Set a} {B : Set b} {C : A → B → Set c} →
((x : A) (y : B) → C x y) → ((y : B) (x : A) → C x y)
-flip f = λ x y → f y x
+flip f = λ y x → f x y
-- Note that _$_ is right associative, like in Haskell. If you want a
-- left associative infix application operator, use
@@ -79,3 +91,16 @@ type-signature : ∀ {a} (A : Set a) → A → A
type-signature A x = x
syntax type-signature A x = x ∶ A
+
+-- Case expressions (to be used with pattern-matching lambdas, see
+-- README.Case).
+
+infix 0 case_return_of_ case_of_
+
+case_return_of_ :
+ ∀ {a b} {A : Set a}
+ (x : A) (B : A → Set b) → ((x : A) → B x) → B x
+case x return B of f = f x
+
+case_of_ : ∀ {a b} {A : Set a} {B : Set b} → A → (A → B) → B
+case x of f = case x return _ of f
diff --git a/src/Function/Bijection.agda b/src/Function/Bijection.agda
index ad66dc9..6f84ae9 100644
--- a/src/Function/Bijection.agda
+++ b/src/Function/Bijection.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Bijections
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Function.Bijection where
open import Data.Product
@@ -53,7 +53,7 @@ record Bijection {f₁ f₂ t₁ t₂}
; surjective = surjective
}
- open Surjection surjection public using (equivalent; right-inverse)
+ open Surjection surjection public using (equivalence; right-inverse)
left-inverse : LeftInverse From To
left-inverse = record
diff --git a/src/Function/Equality.agda b/src/Function/Equality.agda
index 23011ed..85351f8 100644
--- a/src/Function/Equality.agda
+++ b/src/Function/Equality.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Function setoids and related constructions
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Function.Equality where
open import Function as Fun using (_on_)
diff --git a/src/Function/Equivalence.agda b/src/Function/Equivalence.agda
index 39f563c..f32f8f5 100644
--- a/src/Function/Equivalence.agda
+++ b/src/Function/Equivalence.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Equivalence (coinhabitance)
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Function.Equivalence where
open import Function using (flip)
@@ -15,9 +15,9 @@ 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
+record Equivalence {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
@@ -27,11 +27,11 @@ record Equivalent {f₁ f₂ t₁ t₂}
infix 3 _⇔_
_⇔_ : ∀ {f t} → Set f → Set t → Set _
-From ⇔ To = Equivalent (P.setoid From) (P.setoid To)
+From ⇔ To = Equivalence (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 }
+equivalence : ∀ {f t} {From : Set f} {To : Set t} →
+ (From → To) → (To → From) → From ⇔ To
+equivalence to from = record { to = P.→-to-⟶ to; from = P.→-to-⟶ from }
------------------------------------------------------------------------
-- Map and zip
@@ -41,9 +41,9 @@ map : ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂
{From′ : Setoid f₁′ f₂′} {To′ : Setoid t₁′ t₂′} →
((From ⟶ To) → (From′ ⟶ To′)) →
((To ⟶ From) → (To′ ⟶ From′)) →
- Equivalent From To → Equivalent From′ To′
+ Equivalence From To → Equivalence From′ To′
map t f eq = record { to = t to; from = f from }
- where open Equivalent eq
+ where open Equivalence eq
zip : ∀ {f₁₁ f₂₁ t₁₁ t₂₁}
{From₁ : Setoid f₁₁ f₂₁} {To₁ : Setoid t₁₁ t₂₁}
@@ -52,17 +52,18 @@ zip : ∀ {f₁₁ f₂₁ 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
+ Equivalence From₁ To₁ → Equivalence From₂ To₂ →
+ Equivalence From To
zip t f eq₁ eq₂ =
record { to = t (to eq₁) (to eq₂); from = f (from eq₁) (from eq₂) }
- where open Equivalent
+ where open Equivalence
------------------------------------------------------------------------
--- Equivalent is an equivalence relation
+-- Equivalence is an equivalence relation
-- Identity and composition (reflexivity and transitivity).
-id : ∀ {s₁ s₂} → Reflexive (Equivalent {s₁} {s₂})
+id : ∀ {s₁ s₂} → Reflexive (Equivalence {s₁} {s₂})
id {x = S} = record
{ to = F.id
; from = F.id
@@ -71,29 +72,30 @@ id {x = S} = record
infixr 9 _∘_
_∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂} →
- TransFlip (Equivalent {f₁} {f₂} {m₁} {m₂})
- (Equivalent {m₁} {m₂} {t₁} {t₂})
- (Equivalent {f₁} {f₂} {t₁} {t₂})
+ TransFlip (Equivalence {f₁} {f₂} {m₁} {m₂})
+ (Equivalence {m₁} {m₂} {t₁} {t₂})
+ (Equivalence {f₁} {f₂} {t₁} {t₂})
f ∘ g = record
{ to = to f ⟪∘⟫ to g
; from = from g ⟪∘⟫ from f
- } where open Equivalent
+ } where open Equivalence
-- Symmetry.
sym : ∀ {f₁ f₂ t₁ t₂} →
- Sym (Equivalent {f₁} {f₂} {t₁} {t₂}) (Equivalent {t₁} {t₂} {f₁} {f₂})
+ Sym (Equivalence {f₁} {f₂} {t₁} {t₂})
+ (Equivalence {t₁} {t₂} {f₁} {f₂})
sym eq = record
{ from = to
; to = from
- } where open Equivalent eq
+ } where open Equivalence 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
+ ; _≈_ = Equivalence
; isEquivalence = record {refl = id; sym = sym; trans = flip _∘_}
}
diff --git a/src/Function/Injection.agda b/src/Function/Injection.agda
index dc2e4b2..dfc7743 100644
--- a/src/Function/Injection.agda
+++ b/src/Function/Injection.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Injections
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Function.Injection where
open import Function as Fun using () renaming (_∘_ to _⟨∘⟩_)
@@ -11,6 +11,7 @@ open import Level
open import Relation.Binary
open import Function.Equality as F
using (_⟶_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_)
+import Relation.Binary.PropositionalEquality as P
-- Injective functions.
@@ -30,6 +31,13 @@ record Injection {f₁ f₂ t₁ t₂}
to : From ⟶ To
injective : Injective to
+-- The set of all injections from one set to another.
+
+infix 3 _↣_
+
+_↣_ : ∀ {f t} → Set f → Set t → Set _
+From ↣ To = Injection (P.setoid From) (P.setoid To)
+
-- Identity and composition.
infixr 9 _∘_
diff --git a/src/Function/Inverse.agda b/src/Function/Inverse.agda
index d8a8e76..a3a3f31 100644
--- a/src/Function/Inverse.agda
+++ b/src/Function/Inverse.agda
@@ -1,19 +1,17 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Inverses
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Function.Inverse where
open import Level
-open import Function as Fun using (flip)
+open import Function 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
@@ -62,14 +60,14 @@ record Inverse {f₁ f₂ t₁ t₂}
}
open Bijection bijection public
- using (equivalent; surjective; surjection; right-inverse)
+ using (equivalence; surjective; surjection; right-inverse)
-- The set of all inverses between two sets.
-infix 3 _⇿_
+infix 3 _↔_
-_⇿_ : ∀ {f t} → Set f → Set t → Set _
-From ⇿ To = Inverse (P.setoid From) (P.setoid To)
+_↔_ : ∀ {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.
@@ -151,123 +149,13 @@ f ∘ g = record
-- 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
+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 Setoid (Isomorphism-setoid _ _)
-
-open Dummy public
+ } where open Inverse inv
diff --git a/src/Function/LeftInverse.agda b/src/Function/LeftInverse.agda
index 36dae63..04587c3 100644
--- a/src/Function/LeftInverse.agda
+++ b/src/Function/LeftInverse.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Left inverses
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Function.LeftInverse where
open import Data.Product
@@ -12,8 +12,9 @@ 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.Equivalence using (Equivalence)
open import Function.Injection using (Injective; Injection)
+import Relation.Binary.PropositionalEquality as P
-- Left and right inverses.
@@ -51,8 +52,8 @@ record LeftInverse {f₁ f₂ t₁ t₂}
injection : Injection From To
injection = record { to = to; injective = injective }
- equivalent : Equivalent From To
- equivalent = record
+ equivalence : Equivalence From To
+ equivalence = record
{ to = to
; from = from
}
@@ -63,6 +64,14 @@ RightInverse : ∀ {f₁ f₂ t₁ t₂}
(From : Setoid f₁ f₂) (To : Setoid t₁ t₂) → Set _
RightInverse From To = LeftInverse To From
+-- The set of all left inverses from one set to another. (Read A ↞ B
+-- as "surjection from B to A".)
+
+infix 3 _↞_
+
+_↞_ : ∀ {f t} → Set f → Set t → Set _
+From ↞ To = LeftInverse (P.setoid From) (P.setoid To)
+
-- Identity and composition.
id : ∀ {s₁ s₂} {S : Setoid s₁ s₂} → LeftInverse S S
diff --git a/src/Function/Related.agda b/src/Function/Related.agda
new file mode 100644
index 0000000..6eb2430
--- /dev/null
+++ b/src/Function/Related.agda
@@ -0,0 +1,382 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- A universe which includes several kinds of "relatedness" for sets,
+-- such as equivalences, surjections and bijections
+------------------------------------------------------------------------
+
+module Function.Related where
+
+open import Level
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence as Eq using (Equivalence)
+open import Function.Injection as Inj using (Injection; _↣_)
+open import Function.Inverse as Inv using (Inverse; _↔_)
+open import Function.LeftInverse as LeftInv using (LeftInverse)
+open import Function.Surjection as Surj using (Surjection)
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+
+------------------------------------------------------------------------
+-- Wrapper types
+
+-- Synonyms which are used to make _∼[_]_ below "constructor-headed"
+-- (which implies that Agda can deduce the universe code from an
+-- expression matching any of the right-hand sides).
+
+record _←_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where
+ constructor lam
+ field app-← : B → A
+
+open _←_ public
+
+record _↢_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where
+ constructor lam
+ field app-↢ : B ↣ A
+
+open _↢_ public
+
+------------------------------------------------------------------------
+-- Relatedness
+
+-- There are several kinds of "relatedness".
+
+-- The idea to include kinds other than equivalence and bijection came
+-- from Simon Thompson and Bengt Nordström. /NAD
+
+data Kind : Set where
+ implication reverse-implication
+ equivalence
+ injection reverse-injection
+ left-inverse surjection
+ bijection
+ : Kind
+
+-- Interpretation of the codes above. The code "bijection" is
+-- interpreted as Inverse rather than Bijection; the two types are
+-- equivalent.
+
+infix 4 _∼[_]_
+
+_∼[_]_ : ∀ {ℓ₁ ℓ₂} → Set ℓ₁ → Kind → Set ℓ₂ → Set _
+A ∼[ implication ] B = A → B
+A ∼[ reverse-implication ] B = A ← B
+A ∼[ equivalence ] B = Equivalence (P.setoid A) (P.setoid B)
+A ∼[ injection ] B = Injection (P.setoid A) (P.setoid B)
+A ∼[ reverse-injection ] B = A ↢ B
+A ∼[ left-inverse ] B = LeftInverse (P.setoid A) (P.setoid B)
+A ∼[ surjection ] B = Surjection (P.setoid A) (P.setoid B)
+A ∼[ bijection ] B = Inverse (P.setoid A) (P.setoid B)
+
+-- A non-infix synonym.
+
+Related : Kind → ∀ {ℓ₁ ℓ₂} → Set ℓ₁ → Set ℓ₂ → Set _
+Related k A B = A ∼[ k ] B
+
+-- The bijective equality implies any kind of relatedness.
+
+↔⇒ : ∀ {k x y} {X : Set x} {Y : Set y} →
+ X ∼[ bijection ] Y → X ∼[ k ] Y
+↔⇒ {implication} = _⟨$⟩_ ∘ Inverse.to
+↔⇒ {reverse-implication} = lam ∘′ _⟨$⟩_ ∘ Inverse.from
+↔⇒ {equivalence} = Inverse.equivalence
+↔⇒ {injection} = Inverse.injection
+↔⇒ {reverse-injection} = lam ∘′ Inverse.injection ∘ Inv.sym
+↔⇒ {left-inverse} = Inverse.left-inverse
+↔⇒ {surjection} = Inverse.surjection
+↔⇒ {bijection} = id
+
+-- Actual equality also implies any kind of relatedness.
+
+≡⇒ : ∀ {k ℓ} {X Y : Set ℓ} → X ≡ Y → X ∼[ k ] Y
+≡⇒ P.refl = ↔⇒ Inv.id
+
+------------------------------------------------------------------------
+-- Special kinds of kinds
+
+-- Kinds whose interpretation is symmetric.
+
+data Symmetric-kind : Set where
+ equivalence bijection : Symmetric-kind
+
+-- Forgetful map.
+
+⌊_⌋ : Symmetric-kind → Kind
+⌊ equivalence ⌋ = equivalence
+⌊ bijection ⌋ = bijection
+
+-- The proof of symmetry can be found below.
+
+-- Kinds whose interpretation include a function which "goes in the
+-- forward direction".
+
+data Forward-kind : Set where
+ implication equivalence injection
+ left-inverse surjection bijection : Forward-kind
+
+-- Forgetful map.
+
+⌊_⌋→ : Forward-kind → Kind
+⌊ implication ⌋→ = implication
+⌊ equivalence ⌋→ = equivalence
+⌊ injection ⌋→ = injection
+⌊ left-inverse ⌋→ = left-inverse
+⌊ surjection ⌋→ = surjection
+⌊ bijection ⌋→ = bijection
+
+-- The function.
+
+⇒→ : ∀ {k x y} {X : Set x} {Y : Set y} → X ∼[ ⌊ k ⌋→ ] Y → X → Y
+⇒→ {implication} = id
+⇒→ {equivalence} = _⟨$⟩_ ∘ Equivalence.to
+⇒→ {injection} = _⟨$⟩_ ∘ Injection.to
+⇒→ {left-inverse} = _⟨$⟩_ ∘ LeftInverse.to
+⇒→ {surjection} = _⟨$⟩_ ∘ Surjection.to
+⇒→ {bijection} = _⟨$⟩_ ∘ Inverse.to
+
+-- Kinds whose interpretation include a function which "goes backwards".
+
+data Backward-kind : Set where
+ reverse-implication equivalence reverse-injection
+ left-inverse surjection bijection : Backward-kind
+
+-- Forgetful map.
+
+⌊_⌋← : Backward-kind → Kind
+⌊ reverse-implication ⌋← = reverse-implication
+⌊ equivalence ⌋← = equivalence
+⌊ reverse-injection ⌋← = reverse-injection
+⌊ left-inverse ⌋← = left-inverse
+⌊ surjection ⌋← = surjection
+⌊ bijection ⌋← = bijection
+
+-- The function.
+
+⇒← : ∀ {k x y} {X : Set x} {Y : Set y} → X ∼[ ⌊ k ⌋← ] Y → Y → X
+⇒← {reverse-implication} = app-←
+⇒← {equivalence} = _⟨$⟩_ ∘ Equivalence.from
+⇒← {reverse-injection} = _⟨$⟩_ ∘ Injection.to ∘ app-↢
+⇒← {left-inverse} = _⟨$⟩_ ∘ LeftInverse.from
+⇒← {surjection} = _⟨$⟩_ ∘ Surjection.from
+⇒← {bijection} = _⟨$⟩_ ∘ Inverse.from
+
+-- Kinds whose interpretation include functions going in both
+-- directions.
+
+data Equivalence-kind : Set where
+ equivalence left-inverse surjection bijection : Equivalence-kind
+
+-- Forgetful map.
+
+⌊_⌋⇔ : Equivalence-kind → Kind
+⌊ equivalence ⌋⇔ = equivalence
+⌊ left-inverse ⌋⇔ = left-inverse
+⌊ surjection ⌋⇔ = surjection
+⌊ bijection ⌋⇔ = bijection
+
+-- The functions.
+
+⇒⇔ : ∀ {k x y} {X : Set x} {Y : Set y} →
+ X ∼[ ⌊ k ⌋⇔ ] Y → X ∼[ equivalence ] Y
+⇒⇔ {equivalence} = id
+⇒⇔ {left-inverse} = LeftInverse.equivalence
+⇒⇔ {surjection} = Surjection.equivalence
+⇒⇔ {bijection} = Inverse.equivalence
+
+-- Conversions between special kinds.
+
+⇔⌊_⌋ : Symmetric-kind → Equivalence-kind
+⇔⌊ equivalence ⌋ = equivalence
+⇔⌊ bijection ⌋ = bijection
+
+→⌊_⌋ : Equivalence-kind → Forward-kind
+→⌊ equivalence ⌋ = equivalence
+→⌊ left-inverse ⌋ = left-inverse
+→⌊ surjection ⌋ = surjection
+→⌊ bijection ⌋ = bijection
+
+←⌊_⌋ : Equivalence-kind → Backward-kind
+←⌊ equivalence ⌋ = equivalence
+←⌊ left-inverse ⌋ = left-inverse
+←⌊ surjection ⌋ = surjection
+←⌊ bijection ⌋ = bijection
+
+------------------------------------------------------------------------
+-- Opposites
+
+-- For every kind there is an opposite kind.
+
+_op : Kind → Kind
+implication op = reverse-implication
+reverse-implication op = implication
+equivalence op = equivalence
+injection op = reverse-injection
+reverse-injection op = injection
+left-inverse op = surjection
+surjection op = left-inverse
+bijection op = bijection
+
+-- For every morphism there is a corresponding reverse morphism of the
+-- opposite kind.
+
+reverse : ∀ {k a b} {A : Set a} {B : Set b} →
+ A ∼[ k ] B → B ∼[ k op ] A
+reverse {implication} = lam
+reverse {reverse-implication} = app-←
+reverse {equivalence} = Eq.sym
+reverse {injection} = lam
+reverse {reverse-injection} = app-↢
+reverse {left-inverse} = Surj.fromRightInverse
+reverse {surjection} = Surjection.right-inverse
+reverse {bijection} = Inv.sym
+
+------------------------------------------------------------------------
+-- Equational reasoning
+
+-- Equational reasoning for related things.
+
+module EquationalReasoning where
+
+ private
+
+ refl : ∀ {k ℓ} → Reflexive (Related k {ℓ})
+ refl {implication} = id
+ refl {reverse-implication} = lam id
+ refl {equivalence} = Eq.id
+ refl {injection} = Inj.id
+ refl {reverse-injection} = lam Inj.id
+ refl {left-inverse} = LeftInv.id
+ refl {surjection} = Surj.id
+ refl {bijection} = Inv.id
+
+ trans : ∀ {k ℓ₁ ℓ₂ ℓ₃} →
+ Trans (Related k {ℓ₁} {ℓ₂})
+ (Related k {ℓ₂} {ℓ₃})
+ (Related k {ℓ₁} {ℓ₃})
+ trans {implication} = flip _∘′_
+ trans {reverse-implication} = λ f g → lam (app-← f ∘ app-← g)
+ trans {equivalence} = flip Eq._∘_
+ trans {injection} = flip Inj._∘_
+ trans {reverse-injection} = λ f g → lam (Inj._∘_ (app-↢ f) (app-↢ g))
+ trans {left-inverse} = flip LeftInv._∘_
+ trans {surjection} = flip Surj._∘_
+ trans {bijection} = flip Inv._∘_
+
+ sym : ∀ {k ℓ₁ ℓ₂} →
+ Sym (Related ⌊ k ⌋ {ℓ₁} {ℓ₂})
+ (Related ⌊ k ⌋ {ℓ₂} {ℓ₁})
+ sym {equivalence} = Eq.sym
+ sym {bijection} = Inv.sym
+
+ infix 2 _∎
+ infixr 2 _∼⟨_⟩_ _↔⟨_⟩_ _≡⟨_⟩_
+
+ _∼⟨_⟩_ : ∀ {k x y z} (X : Set x) {Y : Set y} {Z : Set z} →
+ X ∼[ k ] Y → Y ∼[ k ] Z → X ∼[ k ] Z
+ _ ∼⟨ X↝Y ⟩ Y↝Z = trans X↝Y Y↝Z
+
+ -- Isomorphisms can be combined with any other kind of relatedness.
+
+ _↔⟨_⟩_ : ∀ {k x y z} (X : Set x) {Y : Set y} {Z : Set z} →
+ X ↔ Y → Y ∼[ k ] Z → X ∼[ k ] Z
+ X ↔⟨ X↔Y ⟩ Y⇔Z = X ∼⟨ ↔⇒ X↔Y ⟩ Y⇔Z
+
+ _≡⟨_⟩_ : ∀ {k ℓ z} (X : Set ℓ) {Y : Set ℓ} {Z : Set z} →
+ X ≡ Y → Y ∼[ k ] Z → X ∼[ k ] Z
+ X ≡⟨ X≡Y ⟩ Y⇔Z = X ∼⟨ ≡⇒ X≡Y ⟩ Y⇔Z
+
+ _∎ : ∀ {k x} (X : Set x) → X ∼[ k ] X
+ X ∎ = refl
+
+-- For a symmetric kind and a fixed universe level we can construct a
+-- setoid.
+
+setoid : Symmetric-kind → (ℓ : Level) → Setoid _ _
+setoid k ℓ = record
+ { Carrier = Set ℓ
+ ; _≈_ = Related ⌊ k ⌋
+ ; isEquivalence =
+ record {refl = _ ∎; sym = sym; trans = _∼⟨_⟩_ _}
+ } where open EquationalReasoning
+
+-- For an arbitrary kind and a fixed universe level we can construct a
+-- preorder.
+
+preorder : Kind → (ℓ : Level) → Preorder _ _ _
+preorder k ℓ = record
+ { Carrier = Set ℓ
+ ; _≈_ = _↔_
+ ; _∼_ = Related k
+ ; isPreorder = record
+ { isEquivalence = Setoid.isEquivalence (setoid bijection ℓ)
+ ; reflexive = ↔⇒
+ ; trans = _∼⟨_⟩_ _
+ }
+ } where open EquationalReasoning
+
+------------------------------------------------------------------------
+-- Some induced relations
+
+-- Every unary relation induces a preorder and, for symmetric kinds,
+-- an equivalence. (No claim is made that these relations are unique.)
+
+InducedRelation₁ : Kind → ∀ {a s} {A : Set a} →
+ (A → Set s) → A → A → Set _
+InducedRelation₁ k S = λ x y → S x ∼[ k ] S y
+
+InducedPreorder₁ : Kind → ∀ {a s} {A : Set a} →
+ (A → Set s) → Preorder _ _ _
+InducedPreorder₁ k S = record
+ { _≈_ = P._≡_
+ ; _∼_ = InducedRelation₁ k S
+ ; isPreorder = record
+ { isEquivalence = P.isEquivalence
+ ; reflexive = reflexive ∘
+ Setoid.reflexive (setoid bijection _) ∘
+ P.cong S
+ ; trans = trans
+ }
+ } where open Preorder (preorder _ _)
+
+InducedEquivalence₁ : Symmetric-kind → ∀ {a s} {A : Set a} →
+ (A → Set s) → Setoid _ _
+InducedEquivalence₁ k S = record
+ { _≈_ = InducedRelation₁ ⌊ k ⌋ S
+ ; isEquivalence = record {refl = refl; sym = sym; trans = trans}
+ } where open Setoid (setoid _ _)
+
+-- Every binary relation induces a preorder and, for symmetric kinds,
+-- an equivalence. (No claim is made that these relations are unique.)
+
+InducedRelation₂ : Kind → ∀ {a b s} {A : Set a} {B : Set b} →
+ (A → B → Set s) → B → B → Set _
+InducedRelation₂ k _S_ = λ x y → ∀ {z} → (z S x) ∼[ k ] (z S y)
+
+InducedPreorder₂ : Kind → ∀ {a b s} {A : Set a} {B : Set b} →
+ (A → B → Set s) → Preorder _ _ _
+InducedPreorder₂ k _S_ = record
+ { _≈_ = P._≡_
+ ; _∼_ = InducedRelation₂ k _S_
+ ; isPreorder = record
+ { isEquivalence = P.isEquivalence
+ ; reflexive = λ x≡y {z} →
+ reflexive $
+ Setoid.reflexive (setoid bijection _) $
+ P.cong (_S_ z) x≡y
+
+ ; trans = λ i↝j j↝k → trans i↝j j↝k
+ }
+ } where open Preorder (preorder _ _)
+
+InducedEquivalence₂ : Symmetric-kind →
+ ∀ {a b s} {A : Set a} {B : Set b} →
+ (A → B → Set s) → Setoid _ _
+InducedEquivalence₂ k _S_ = record
+ { _≈_ = InducedRelation₂ ⌊ k ⌋ _S_
+ ; isEquivalence = record
+ { refl = refl
+ ; sym = λ i↝j → sym i↝j
+ ; trans = λ i↝j j↝k → trans i↝j j↝k
+ }
+ } where open Setoid (setoid _ _)
diff --git a/src/Function/Inverse/TypeIsomorphisms.agda b/src/Function/Related/TypeIsomorphisms.agda
index 6a26889..dfd57fb 100644
--- a/src/Function/Inverse/TypeIsomorphisms.agda
+++ b/src/Function/Related/TypeIsomorphisms.agda
@@ -1,10 +1,11 @@
------------------------------------------------------------------------
--- Various basic type isomorphisms
+-- The Agda standard library
+--
+-- Basic lemmas showing that various types are related (isomorphic or
+-- equivalent or…)
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-module Function.Inverse.TypeIsomorphisms where
+module Function.Related.TypeIsomorphisms where
open import Algebra
import Algebra.FunctionProperties as FP
@@ -15,10 +16,9 @@ open import Data.Unit
open import Level
open import Function
open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence
- using (_⇔_; equivalent; module Equivalent)
-open import Function.Inverse as Inv
- using (Kind; Isomorphism; _⇿_; module Inverse)
+open import Function.Equivalence as Eq using (_⇔_; module Equivalence)
+open import Function.Inverse using (_↔_; module Inverse)
+open import Function.Related as Related
open import Relation.Binary
open import Relation.Binary.Product.Pointwise
open import Relation.Binary.PropositionalEquality as P using (_≡_; _≗_)
@@ -26,48 +26,48 @@ open import Relation.Binary.Sum
open import Relation.Nullary
------------------------------------------------------------------------
+-- Σ is "associative"
+
+Σ-assoc : ∀ {a b c}
+ {A : Set a} {B : A → Set b} {C : (a : A) → B a → Set c} →
+ Σ (Σ A B) (uncurry C) ↔ Σ A (λ a → Σ (B a) (C a))
+Σ-assoc = record
+ { to = P.→-to-⟶ λ p →
+ proj₁ (proj₁ p) , (proj₂ (proj₁ p) , proj₂ p)
+ ; from = P.→-to-⟶ _
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.refl
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+
+------------------------------------------------------------------------
-- ⊥, ⊤, _×_ and _⊎_ form a commutative semiring
-×-CommutativeMonoid : Kind → (ℓ : Level) → CommutativeMonoid _ _
+×-CommutativeMonoid : Symmetric-kind → (ℓ : Level) →
+ CommutativeMonoid _ _
×-CommutativeMonoid k ℓ = record
{ Carrier = Set ℓ
- ; _≈_ = Isomorphism k
+ ; _≈_ = Related ⌊ 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
+ { isEquivalence = Setoid.isEquivalence $ Related.setoid k ℓ
+ ; assoc = λ _ _ _ → ↔⇒ Σ-assoc
+ ; ∙-cong = _×-cong_
}
- ; identityˡ = λ A → Inv.⇿⇒ $ left-identity A
- ; comm = λ A B → Inv.⇿⇒ $ comm A B
+ ; identityˡ = λ A → ↔⇒ $ left-identity A
+ ; comm = λ A B → ↔⇒ $ 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 = _×-⇿_
+ open FP _↔_
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))
+ { to = P.→-to-⟶ proj₂
+ ; from = P.→-to-⟶ λ y → _ , y
; inverse-of = record
{ left-inverse-of = λ _ → P.refl
; right-inverse-of = λ _ → P.refl
@@ -76,72 +76,53 @@ open import Relation.Nullary
comm : Commutative _×_
comm _ _ = record
- { to = P.→-to-⟶ $ Prod.swap {a = ℓ} {b = ℓ}
- ; from = P.→-to-⟶ $ Prod.swap {a = ℓ} {b = ℓ}
+ { to = P.→-to-⟶ Prod.swap
+ ; from = P.→-to-⟶ Prod.swap
; inverse-of = record
{ left-inverse-of = λ _ → P.refl
; right-inverse-of = λ _ → P.refl
}
}
-⊎-CommutativeMonoid : Kind → (ℓ : Level) → CommutativeMonoid _ _
+⊎-CommutativeMonoid : Symmetric-kind → (ℓ : Level) →
+ CommutativeMonoid _ _
⊎-CommutativeMonoid k ℓ = record
{ Carrier = Set ℓ
- ; _≈_ = Isomorphism k
+ ; _≈_ = Related ⌊ 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
+ { isEquivalence = Setoid.isEquivalence $ Related.setoid k ℓ
+ ; assoc = λ A B C → ↔⇒ $ assoc A B C
+ ; ∙-cong = _⊎-cong_
}
- ; identityˡ = λ A → Inv.⇿⇒ $ left-identity A
- ; comm = λ A B → Inv.⇿⇒ $ comm A B
+ ; identityˡ = λ A → ↔⇒ $ left-identity A
+ ; comm = λ A B → ↔⇒ $ 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 = _⊎-⇿_
+ open FP _↔_
left-identity : LeftIdentity (Lift ⊥) (_⊎_ {a = ℓ} {b = ℓ})
left-identity A = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
+ { to = P.→-to-⟶ [ (λ ()) ∘′ lower , id ]
+ ; from = P.→-to-⟶ inj₂
; inverse-of = record
{ right-inverse-of = λ _ → P.refl
- ; left-inverse-of =
- [ ⊥-elim {Whatever = _ ≡ _} ∘ lower , (λ _ → P.refl) ]
+ ; left-inverse-of = [ ⊥-elim ∘ 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
+ { to = P.→-to-⟶ [ [ inj₁ , inj₂ ∘ inj₁ ] , inj₂ ∘ inj₂ ]
+ ; from = P.→-to-⟶ [ inj₁ ∘ inj₁ , [ inj₁ ∘ inj₂ , inj₂ ] ]
; 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
@@ -159,10 +140,11 @@ open import Relation.Nullary
inv : ∀ {A B} → swap ∘ swap {A} {B} ≗ id
inv = [ (λ _ → P.refl) , (λ _ → P.refl) ]
-×⊎-CommutativeSemiring : Kind → (ℓ : Level) → CommutativeSemiring _ _
+×⊎-CommutativeSemiring : Symmetric-kind → (ℓ : Level) →
+ CommutativeSemiring _ _
×⊎-CommutativeSemiring k ℓ = record
{ Carrier = Set ℓ
- ; _≈_ = Isomorphism k
+ ; _≈_ = Related ⌊ k ⌋
; _+_ = _⊎_
; _*_ = _×_
; 0# = Lift ⊥
@@ -172,17 +154,17 @@ open import Relation.Nullary
⊎-CommutativeMonoid k ℓ
; *-isCommutativeMonoid = isCommutativeMonoid $
×-CommutativeMonoid k ℓ
- ; distribʳ = λ A B C → Inv.⇿⇒ $ right-distrib A B C
- ; zeroˡ = λ A → Inv.⇿⇒ $ left-zero A
+ ; distribʳ = λ A B C → ↔⇒ $ right-distrib A B C
+ ; zeroˡ = λ A → ↔⇒ $ left-zero A
}
}
where
open CommutativeMonoid
- open FP _⇿_
+ open FP _↔_
left-zero : LeftZero (Lift ⊥) (_×_ {a = ℓ} {b = ℓ})
left-zero A = record
- { to = P.→-to-⟶ $ proj₁ {a = ℓ} {b = ℓ}
+ { to = P.→-to-⟶ proj₁
; from = P.→-to-⟶ (⊥-elim ∘′ lower)
; inverse-of = record
{ left-inverse-of = λ p → ⊥-elim (lower $ proj₁ p)
@@ -192,7 +174,7 @@ open import Relation.Nullary
right-distrib : _×_ DistributesOverʳ _⊎_
right-distrib A B C = record
- { to = P.→-to-⟶ to
+ { to = P.→-to-⟶ $ uncurry [ curry inj₁ , curry inj₂ ]
; from = P.→-to-⟶ from
; inverse-of = record
{ right-inverse-of = [ (λ _ → P.refl) , (λ _ → P.refl) ]
@@ -201,18 +183,15 @@ open import Relation.Nullary
}
}
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
+ΠΠ↔ΠΠ : ∀ {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
@@ -221,23 +200,23 @@ open import Relation.Nullary
}
}
-∃∃⇿∃∃ : ∀ {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))
+∃∃↔∃∃ : ∀ {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-⟶ λ p → (proj₁ (proj₂ p) , proj₁ p , proj₂ (proj₂ p))
+ ; from = P.→-to-⟶ λ 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
+Π↔Π : ∀ {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
@@ -247,75 +226,74 @@ open import Relation.Nullary
}
------------------------------------------------------------------------
--- _→_ preserves isomorphisms
+-- _→_ preserves the symmetric relations
_→-cong-⇔_ :
∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
A ⇔ B → C ⇔ D → (A → C) ⇔ (B → D)
A⇔B →-cong-⇔ C⇔D = record
{ to = P.→-to-⟶ λ f x →
- Equivalent.to C⇔D ⟨$⟩ f (Equivalent.from A⇔B ⟨$⟩ x)
+ Equivalence.to C⇔D ⟨$⟩ f (Equivalence.from A⇔B ⟨$⟩ x)
; from = P.→-to-⟶ λ f x →
- Equivalent.from C⇔D ⟨$⟩ f (Equivalent.to A⇔B ⟨$⟩ x)
+ Equivalence.from C⇔D ⟨$⟩ f (Equivalence.to A⇔B ⟨$⟩ x)
}
→-cong :
∀ {a b c d} →
P.Extensionality a c → P.Extensionality b d →
∀ {k} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
- Isomorphism k A B → Isomorphism k C D → Isomorphism k (A → C) (B → D)
-→-cong extAC extBD {Inv.equivalent} A⇔B C⇔D = A⇔B →-cong-⇔ C⇔D
-→-cong extAC extBD {Inv.inverse} A⇿B C⇿D = record
- { to = Equivalent.to A→C⇔B→D
- ; from = Equivalent.from A→C⇔B→D
+ A ∼[ ⌊ k ⌋ ] B → C ∼[ ⌊ k ⌋ ] D → (A → C) ∼[ ⌊ k ⌋ ] (B → D)
+→-cong extAC extBD {equivalence} A⇔B C⇔D = A⇔B →-cong-⇔ C⇔D
+→-cong extAC extBD {bijection} A↔B C↔D = record
+ { to = Equivalence.to A→C⇔B→D
+ ; from = Equivalence.from A→C⇔B→D
; inverse-of = record
{ left-inverse-of = λ f → extAC λ x → begin
- Inverse.from C⇿D ⟨$⟩ (Inverse.to C⇿D ⟨$⟩
- f (Inverse.from A⇿B ⟨$⟩ (Inverse.to A⇿B ⟨$⟩ x))) ≡⟨ Inverse.left-inverse-of C⇿D _ ⟩
- f (Inverse.from A⇿B ⟨$⟩ (Inverse.to A⇿B ⟨$⟩ x)) ≡⟨ P.cong f $ Inverse.left-inverse-of A⇿B x ⟩
+ Inverse.from C↔D ⟨$⟩ (Inverse.to C↔D ⟨$⟩
+ f (Inverse.from A↔B ⟨$⟩ (Inverse.to A↔B ⟨$⟩ x))) ≡⟨ Inverse.left-inverse-of C↔D _ ⟩
+ f (Inverse.from A↔B ⟨$⟩ (Inverse.to A↔B ⟨$⟩ x)) ≡⟨ P.cong f $ Inverse.left-inverse-of A↔B x ⟩
f x ∎
; right-inverse-of = λ f → extBD λ x → begin
- Inverse.to C⇿D ⟨$⟩ (Inverse.from C⇿D ⟨$⟩
- f (Inverse.to A⇿B ⟨$⟩ (Inverse.from A⇿B ⟨$⟩ x))) ≡⟨ Inverse.right-inverse-of C⇿D _ ⟩
- f (Inverse.to A⇿B ⟨$⟩ (Inverse.from A⇿B ⟨$⟩ x)) ≡⟨ P.cong f $ Inverse.right-inverse-of A⇿B x ⟩
+ Inverse.to C↔D ⟨$⟩ (Inverse.from C↔D ⟨$⟩
+ f (Inverse.to A↔B ⟨$⟩ (Inverse.from A↔B ⟨$⟩ x))) ≡⟨ Inverse.right-inverse-of C↔D _ ⟩
+ f (Inverse.to A↔B ⟨$⟩ (Inverse.from A↔B ⟨$⟩ x)) ≡⟨ P.cong f $ Inverse.right-inverse-of A↔B x ⟩
f x ∎
}
}
where
open P.≡-Reasoning
- A→C⇔B→D = Inv.⇿⇒ A⇿B →-cong-⇔ Inv.⇿⇒ C⇿D
+ A→C⇔B→D = ↔⇒ A↔B →-cong-⇔ ↔⇒ C↔D
------------------------------------------------------------------------
--- ¬_ preserves isomorphisms
+-- ¬_ preserves the symmetric relations
¬-cong-⇔ : ∀ {a b} {A : Set a} {B : Set b} →
A ⇔ B → (¬ A) ⇔ (¬ B)
¬-cong-⇔ A⇔B = A⇔B →-cong-⇔ (⊥ ∎)
- where open Inv.EquationalReasoning
+ where open Related.EquationalReasoning
¬-cong : ∀ {a b} →
P.Extensionality a zero → P.Extensionality b zero →
∀ {k} {A : Set a} {B : Set b} →
- Isomorphism k A B → Isomorphism k (¬ A) (¬ B)
+ A ∼[ ⌊ k ⌋ ] B → (¬ A) ∼[ ⌊ k ⌋ ] (¬ B)
¬-cong extA extB A≈B = →-cong extA extB A≈B (⊥ ∎)
- where open Inv.EquationalReasoning
+ where open Related.EquationalReasoning
------------------------------------------------------------------------
-- _⇔_ preserves _⇔_
-- The type of the following proof is a bit more general.
-Isomorphism-cong :
+Related-cong :
∀ {k a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
- Isomorphism k A B → Isomorphism k C D →
- Isomorphism k A C ⇔ Isomorphism k B D
-Isomorphism-cong {A = A} {B} {C} {D} A≈B C≈D =
- equivalent (λ A≈C → B ≈⟨ sym A≈B ⟩
- A ≈⟨ A≈C ⟩
- C ≈⟨ C≈D ⟩
- D ∎)
- (λ B≈D → A ≈⟨ A≈B ⟩
- B ≈⟨ B≈D ⟩
- D ≈⟨ sym C≈D ⟩
- C ∎)
- where open Inv.EquationalReasoning
+ A ∼[ ⌊ k ⌋ ] B → C ∼[ ⌊ k ⌋ ] D → (A ∼[ ⌊ k ⌋ ] C) ⇔ (B ∼[ ⌊ k ⌋ ] D)
+Related-cong {A = A} {B} {C} {D} A≈B C≈D =
+ Eq.equivalence (λ A≈C → B ∼⟨ sym A≈B ⟩
+ A ∼⟨ A≈C ⟩
+ C ∼⟨ C≈D ⟩
+ D ∎)
+ (λ B≈D → A ∼⟨ A≈B ⟩
+ B ∼⟨ B≈D ⟩
+ D ∼⟨ sym C≈D ⟩
+ C ∎)
+ where open Related.EquationalReasoning
diff --git a/src/Function/Surjection.agda b/src/Function/Surjection.agda
index fcf586d..6d27a29 100644
--- a/src/Function/Surjection.agda
+++ b/src/Function/Surjection.agda
@@ -1,15 +1,15 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- 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.Equivalence using (Equivalence)
open import Function.Injection hiding (id; _∘_)
open import Function.LeftInverse as Left hiding (id; _∘_)
open import Data.Product
@@ -50,12 +50,25 @@ record Surjection {f₁ f₂ t₁ t₂}
injection : Injection To From
injection = LeftInverse.injection right-inverse
- equivalent : Equivalent From To
- equivalent = record
+ equivalence : Equivalence From To
+ equivalence = record
{ to = to
; from = from
}
+-- Right inverses can be turned into surjections.
+
+fromRightInverse :
+ ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} →
+ RightInverse From To → Surjection From To
+fromRightInverse r = record
+ { to = from
+ ; surjective = record
+ { from = to
+ ; right-inverse-of = left-inverse-of
+ }
+ } where open LeftInverse r
+
-- The set of all surjections from one set to another.
infix 3 _↠_
diff --git a/src/IO.agda b/src/IO.agda
index 5d9ad1a..f89bbad 100644
--- a/src/IO.agda
+++ b/src/IO.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- IO
------------------------------------------------------------------------
@@ -10,8 +12,9 @@ open import Coinduction
open import Data.Unit
open import Data.String
open import Data.Colist
-import Foreign.Haskell as Haskell
+open import Function
import IO.Primitive as Prim
+open import Level
------------------------------------------------------------------------
-- The IO monad
@@ -25,18 +28,18 @@ import IO.Primitive as Prim
infixl 1 _>>=_ _>>_
-data IO : Set → Set₁ where
- lift : ∀ {A} (m : Prim.IO A) → IO A
- return : ∀ {A} (x : A) → IO A
- _>>=_ : ∀ {A B} (m : ∞ (IO A)) (f : (x : A) → ∞ (IO B)) → IO B
- _>>_ : ∀ {A B} (m₁ : ∞ (IO A)) (m₂ : ∞ (IO B)) → IO B
+data IO {a} (A : Set a) : Set (suc a) where
+ lift : (m : Prim.IO A) → IO A
+ return : (x : A) → IO A
+ _>>=_ : {B : Set a} (m : ∞ (IO B)) (f : (x : B) → ∞ (IO A)) → IO A
+ _>>_ : {B : Set a} (m₁ : ∞ (IO B)) (m₂ : ∞ (IO A)) → IO A
-- The use of abstract ensures that the run function will not be
-- unfolded infinitely by the type checker.
abstract
- run : ∀ {A} → IO A → Prim.IO A
+ run : ∀ {a} {A : Set a} → IO A → Prim.IO A
run (lift m) = m
run (return x) = Prim.return x
run (m >>= f) = Prim._>>=_ (run (♭ m )) λ x → run (♭ (f x))
@@ -45,21 +48,27 @@ abstract
------------------------------------------------------------------------
-- Utilities
--- Because IO A lives in Set₁ I hesitate to define sequence, which
--- would require defining a Set₁ variant of Colist.
+sequence : ∀ {a} {A : Set a} → Colist (IO A) → IO (Colist A)
+sequence [] = return []
+sequence (c ∷ cs) = ♯ c >>= λ x →
+ ♯ (♯ sequence (♭ cs) >>= λ xs →
+ ♯ return (x ∷ ♯ xs))
-mapM : {A B : Set} → (A → IO B) → Colist A → IO (Colist B)
-mapM f [] = return []
-mapM f (x ∷ xs) = ♯ f x >>= λ y →
- ♯ (♯ mapM f (♭ xs) >>= λ ys →
- ♯ return (y ∷ ♯ ys))
+-- The reason for not defining sequence′ in terms of sequence is
+-- efficiency (the unused results could cause unnecessary memory use).
--- The reason for not defining mapM′ in terms of mapM is efficiency
--- (the unused results could cause unnecessary memory use).
+sequence′ : ∀ {a} {A : Set a} → Colist (IO A) → IO (Lift ⊤)
+sequence′ [] = return _
+sequence′ (c ∷ cs) = ♯ c >>
+ ♯ (♯ sequence′ (♭ cs) >>
+ ♯ return _)
-mapM′ : {A B : Set} → (A → IO B) → Colist A → IO ⊤
-mapM′ f [] = return _
-mapM′ f (x ∷ xs) = ♯ f x >> ♯ mapM′ f (♭ xs)
+mapM : ∀ {a b} {A : Set a} {B : Set b} →
+ (A → IO B) → Colist A → IO (Colist B)
+mapM f = sequence ∘ map f
+
+mapM′ : {A B : Set} → (A → IO B) → Colist A → IO (Lift ⊤)
+mapM′ f = sequence′ ∘ map f
------------------------------------------------------------------------
-- Simple lazy IO
@@ -71,18 +80,14 @@ mapM′ f (x ∷ xs) = ♯ f x >> ♯ mapM′ f (♭ xs)
-- version 3) the functions use ISO-8859-1.
getContents : IO Costring
-getContents =
- ♯ lift Prim.getContents >>= λ s →
- ♯ return (Haskell.toColist s)
+getContents = lift Prim.getContents
readFile : String → IO Costring
-readFile f =
- ♯ lift (Prim.readFile f) >>= λ s →
- ♯ return (Haskell.toColist s)
+readFile f = lift (Prim.readFile f)
writeFile∞ : String → Costring → IO ⊤
writeFile∞ f s =
- ♯ lift (Prim.writeFile f (Haskell.fromColist s)) >>
+ ♯ lift (Prim.writeFile f s) >>
♯ return _
writeFile : String → String → IO ⊤
@@ -90,7 +95,7 @@ writeFile f s = writeFile∞ f (toCostring s)
appendFile∞ : String → Costring → IO ⊤
appendFile∞ f s =
- ♯ lift (Prim.appendFile f (Haskell.fromColist s)) >>
+ ♯ lift (Prim.appendFile f s) >>
♯ return _
appendFile : String → String → IO ⊤
@@ -98,7 +103,7 @@ appendFile f s = appendFile∞ f (toCostring s)
putStr∞ : Costring → IO ⊤
putStr∞ s =
- ♯ lift (Prim.putStr (Haskell.fromColist s)) >>
+ ♯ lift (Prim.putStr s) >>
♯ return _
putStr : String → IO ⊤
@@ -106,7 +111,7 @@ putStr s = putStr∞ (toCostring s)
putStrLn∞ : Costring → IO ⊤
putStrLn∞ s =
- ♯ lift (Prim.putStrLn (Haskell.fromColist s)) >>
+ ♯ lift (Prim.putStrLn s) >>
♯ return _
putStrLn : String → IO ⊤
diff --git a/src/IO/Primitive.agda b/src/IO/Primitive.agda
index 7764519..b046e0e 100644
--- a/src/IO/Primitive.agda
+++ b/src/IO/Primitive.agda
@@ -1,10 +1,12 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Primitive IO: simple bindings to Haskell types and functions
------------------------------------------------------------------------
module IO.Primitive where
-open import Data.String hiding (Costring)
+open import Data.String
open import Data.Char
open import Foreign.Haskell
@@ -12,18 +14,20 @@ open import Foreign.Haskell
-- The IO monad
postulate
- IO : Set → Set
+ IO : ∀ {ℓ} → Set ℓ → Set ℓ
-{-# COMPILED_TYPE IO IO #-}
+{-# IMPORT IO.FFI #-}
+{-# COMPILED_TYPE IO IO.FFI.AgdaIO #-}
+{-# BUILTIN IO IO #-}
infixl 1 _>>=_
postulate
- return : ∀ {A} → A → IO A
- _>>=_ : ∀ {A B} → IO A → (A → IO B) → IO B
+ return : ∀ {a} {A : Set a} → A → IO A
+ _>>=_ : ∀ {a b} {A : Set a} {B : Set b} → IO A → (A → IO B) → IO B
-{-# COMPILED return (\_ -> return :: a -> IO a) #-}
-{-# COMPILED _>>=_ (\_ _ -> (>>=) :: IO a -> (a -> IO b) -> IO b) #-}
+{-# COMPILED return (\_ _ -> return) #-}
+{-# COMPILED _>>=_ (\_ _ _ _ -> (>>=)) #-}
------------------------------------------------------------------------
-- Simple lazy IO
@@ -34,9 +38,6 @@ postulate
-- locale. For older versions of the library (going back to at least
-- version 3) the functions use ISO-8859-1.
-private
- Costring = Colist Char
-
postulate
getContents : IO Costring
readFile : String → IO Costring
diff --git a/src/Induction.agda b/src/Induction.agda
index 3902c65..f145c34 100644
--- a/src/Induction.agda
+++ b/src/Induction.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- An abstraction of various forms of recursion/induction
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- The idea underlying Induction.* comes from Epigram 1, see Section 4
-- of "The view from the left" by McBride and McKinna.
diff --git a/src/Induction/Lexicographic.agda b/src/Induction/Lexicographic.agda
index 6cf5a7c..d4065e9 100644
--- a/src/Induction/Lexicographic.agda
+++ b/src/Induction/Lexicographic.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Lexicographic induction
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Induction.Lexicographic where
open import Induction
diff --git a/src/Induction/Nat.agda b/src/Induction/Nat.agda
index 78b6b81..eabd6aa 100644
--- a/src/Induction/Nat.agda
+++ b/src/Induction/Nat.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Various forms of induction for natural numbers
------------------------------------------------------------------------
diff --git a/src/Induction/WellFounded.agda b/src/Induction/WellFounded.agda
index 9ed9e52..c455aba 100644
--- a/src/Induction/WellFounded.agda
+++ b/src/Induction/WellFounded.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Well-founded induction
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Induction.WellFounded where
diff --git a/src/Level.agda b/src/Level.agda
index 2aefb6a..4c2bd5e 100644
--- a/src/Level.agda
+++ b/src/Level.agda
@@ -1,31 +1,33 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Universe levels
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Level where
-- Levels.
-data Level : Set where
- zero : Level
- suc : (i : Level) → Level
-
-{-# BUILTIN LEVEL Level #-}
-{-# BUILTIN LEVELZERO zero #-}
-{-# BUILTIN LEVELSUC suc #-}
+infixl 6 _⊔_
--- Maximum.
+postulate
+ Level : Set
+ zero : Level
+ suc : Level → Level
+ _⊔_ : Level → Level → Level
-infixl 6 _⊔_
+-- MAlonzo compiles Level to (). This should be safe, because it is
+-- not possible to pattern match on levels.
-_⊔_ : Level → Level → Level
-zero ⊔ j = j
-suc i ⊔ zero = suc i
-suc i ⊔ suc j = suc (i ⊔ j)
+{-# COMPILED_TYPE Level () #-}
+{-# COMPILED zero () #-}
+{-# COMPILED suc (\_ -> ()) #-}
+{-# COMPILED _⊔_ (\_ _ -> ()) #-}
-{-# BUILTIN LEVELMAX _⊔_ #-}
+{-# BUILTIN LEVEL Level #-}
+{-# BUILTIN LEVELZERO zero #-}
+{-# BUILTIN LEVELSUC suc #-}
+{-# BUILTIN LEVELMAX _⊔_ #-}
-- Lifting.
diff --git a/src/Record.agda b/src/Record.agda
new file mode 100644
index 0000000..f6a2000
--- /dev/null
+++ b/src/Record.agda
@@ -0,0 +1,230 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Record types with manifest fields and "with", based on Randy
+-- Pollack's "Dependently Typed Records in Type Theory"
+------------------------------------------------------------------------
+
+-- For an example of how this module can be used, see README.Record.
+
+-- The module is parametrised by the type of labels, which should come
+-- with decidable equality.
+
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality
+
+module Record (Label : Set) (_≟_ : Decidable (_≡_ {A = Label})) where
+
+open import Data.Bool
+open import Data.Empty
+open import Data.List
+open import Data.Product hiding (proj₁; proj₂)
+open import Data.Unit
+open import Function hiding (type-signature)
+open import Level
+open import Relation.Nullary
+open import Relation.Nullary.Decidable
+
+------------------------------------------------------------------------
+-- A Σ-type with a manifest field
+
+-- A variant of Σ where the value of the second field is "manifest"
+-- (given by the first).
+
+infix 4 _,
+
+record Manifest-Σ {a b} (A : Set a) {B : A → Set b}
+ (f : (x : A) → B x) : Set a where
+ constructor _,
+ field proj₁ : A
+
+ proj₂ : B proj₁
+ proj₂ = f proj₁
+
+------------------------------------------------------------------------
+-- Signatures and records
+
+mutual
+
+ infixl 5 _,_∶_ _,_≔_
+
+ data Signature s : Set (suc s) where
+ ∅ : Signature s
+ _,_∶_ : (Sig : Signature s)
+ (ℓ : Label)
+ (A : Record Sig → Set s) →
+ Signature s
+ _,_≔_ : (Sig : Signature s)
+ (ℓ : Label)
+ {A : Record Sig → Set s}
+ (a : (r : Record Sig) → A r) →
+ Signature s
+
+ -- Record is a record type to ensure that the signature can be
+ -- inferred from a value of type Record Sig.
+
+ record Record {s} (Sig : Signature s) : Set s where
+ constructor rec
+ field fun : Record-fun Sig
+
+ Record-fun : ∀ {s} → Signature s → Set s
+ Record-fun ∅ = Lift ⊤
+ Record-fun (Sig , ℓ ∶ A) = Σ (Record Sig) A
+ Record-fun (Sig , ℓ ≔ a) = Manifest-Σ (Record Sig) a
+
+------------------------------------------------------------------------
+-- Variants of Signature and Record
+
+-- It may be easier to define values of type Record′ (Sig⇒Sig′ Sig)
+-- than of type Record Sig.
+
+mutual
+
+ data Signature′ s : Set (suc s) where
+ ∅ : Signature′ s
+ _,_∶_ : (Sig : Signature′ s)
+ (ℓ : Label)
+ (A : Record′ Sig → Set s) →
+ Signature′ s
+ _,_≔_ : (Sig : Signature′ s)
+ (ℓ : Label)
+ {A : Record′ Sig → Set s}
+ (a : (r : Record′ Sig) → A r) →
+ Signature′ s
+
+ Record′ : ∀ {s} → Signature′ s → Set s
+ Record′ ∅ = Lift ⊤
+ Record′ (Sig , ℓ ∶ A) = Σ (Record′ Sig) A
+ Record′ (Sig , ℓ ≔ a) = Manifest-Σ (Record′ Sig) a
+
+-- We can convert between the two kinds of signatures/records.
+
+mutual
+
+ Sig′⇒Sig : ∀ {s} → Signature′ s → Signature s
+ Sig′⇒Sig ∅ = ∅
+ Sig′⇒Sig (Sig , ℓ ∶ A) = Sig′⇒Sig Sig , ℓ ∶ (A ∘ Rec⇒Rec′ _)
+ Sig′⇒Sig (Sig , ℓ ≔ a) = Sig′⇒Sig Sig , ℓ ≔ (a ∘ Rec⇒Rec′ _)
+
+ Rec⇒Rec′ : ∀ {s} (Sig : Signature′ s) →
+ Record (Sig′⇒Sig Sig) → Record′ Sig
+ Rec⇒Rec′ ∅ (rec r) = r
+ Rec⇒Rec′ (Sig , ℓ ∶ A) (rec r) = (Rec⇒Rec′ _ (Σ.proj₁ r) , Σ.proj₂ r)
+ Rec⇒Rec′ (Sig , ℓ ≔ a) (rec r) = (Rec⇒Rec′ _ (Manifest-Σ.proj₁ r) ,)
+
+mutual
+
+ Sig⇒Sig′ : ∀ {s} → Signature s → Signature′ s
+ Sig⇒Sig′ ∅ = ∅
+ Sig⇒Sig′ (Sig , ℓ ∶ A) = Sig⇒Sig′ Sig , ℓ ∶ (A ∘ Rec′⇒Rec _)
+ Sig⇒Sig′ (Sig , ℓ ≔ a) = Sig⇒Sig′ Sig , ℓ ≔ (a ∘ Rec′⇒Rec _)
+
+ Rec′⇒Rec : ∀ {s} (Sig : Signature s) →
+ Record′ (Sig⇒Sig′ Sig) → Record Sig
+ Rec′⇒Rec ∅ r = rec r
+ Rec′⇒Rec (Sig , ℓ ∶ A) r = rec (Rec′⇒Rec _ (Σ.proj₁ r) , Σ.proj₂ r)
+ Rec′⇒Rec (Sig , ℓ ≔ a) r = rec (Rec′⇒Rec _ (Manifest-Σ.proj₁ r) ,)
+
+------------------------------------------------------------------------
+-- Labels
+
+-- A signature's labels, starting with the last one.
+
+labels : ∀ {s} → Signature s → List Label
+labels ∅ = []
+labels (Sig , ℓ ∶ A) = ℓ ∷ labels Sig
+labels (Sig , ℓ ≔ a) = ℓ ∷ labels Sig
+
+-- Inhabited if the label is part of the signature.
+
+infix 4 _∈_
+
+_∈_ : ∀ {s} → Label → Signature s → Set
+ℓ ∈ Sig =
+ foldr (λ ℓ′ → if ⌊ ℓ ≟ ℓ′ ⌋ then (λ _ → ⊤) else id) ⊥ (labels Sig)
+
+------------------------------------------------------------------------
+-- Projections
+
+-- Signature restriction and projection. (Restriction means removal of
+-- a given field and all subsequent fields.)
+
+Restrict : ∀ {s} (Sig : Signature s) (ℓ : Label) → ℓ ∈ Sig →
+ Signature s
+Restrict ∅ ℓ ()
+Restrict (Sig , ℓ′ ∶ A) ℓ ℓ∈ with ℓ ≟ ℓ′
+... | yes _ = Sig
+... | no _ = Restrict Sig ℓ ℓ∈
+Restrict (Sig , ℓ′ ≔ a) ℓ ℓ∈ with ℓ ≟ ℓ′
+... | yes _ = Sig
+... | no _ = Restrict Sig ℓ ℓ∈
+
+Restricted : ∀ {s} (Sig : Signature s) (ℓ : Label) → ℓ ∈ Sig → Set s
+Restricted Sig ℓ ℓ∈ = Record (Restrict Sig ℓ ℓ∈)
+
+Proj : ∀ {s} (Sig : Signature s) (ℓ : Label) {ℓ∈ : ℓ ∈ Sig} →
+ Restricted Sig ℓ ℓ∈ → Set s
+Proj ∅ ℓ {}
+Proj (Sig , ℓ′ ∶ A) ℓ with ℓ ≟ ℓ′
+... | yes _ = A
+... | no _ = Proj Sig ℓ
+Proj (_,_≔_ Sig ℓ′ {A = A} a) ℓ with ℓ ≟ ℓ′
+... | yes _ = A
+... | no _ = Proj Sig ℓ
+
+-- Record restriction and projection.
+
+infixl 5 _∣_
+
+_∣_ : ∀ {s} {Sig : Signature s} → Record Sig →
+ (ℓ : Label) {ℓ∈ : ℓ ∈ Sig} → Restricted Sig ℓ ℓ∈
+_∣_ {Sig = ∅} r ℓ {}
+_∣_ {Sig = Sig , ℓ′ ∶ A} (rec r) ℓ with ℓ ≟ ℓ′
+... | yes _ = Σ.proj₁ r
+... | no _ = Σ.proj₁ r ∣ ℓ
+_∣_ {Sig = Sig , ℓ′ ≔ a} (rec r) ℓ with ℓ ≟ ℓ′
+... | yes _ = Manifest-Σ.proj₁ r
+... | no _ = Manifest-Σ.proj₁ r ∣ ℓ
+
+infixl 5 _·_
+
+_·_ : ∀ {s} {Sig : Signature s} (r : Record Sig)
+ (ℓ : Label) {ℓ∈ : ℓ ∈ Sig} →
+ Proj Sig ℓ {ℓ∈} (r ∣ ℓ)
+_·_ {Sig = ∅} r ℓ {}
+_·_ {Sig = Sig , ℓ′ ∶ A} (rec r) ℓ with ℓ ≟ ℓ′
+... | yes _ = Σ.proj₂ r
+... | no _ = Σ.proj₁ r · ℓ
+_·_ {Sig = Sig , ℓ′ ≔ a} (rec r) ℓ with ℓ ≟ ℓ′
+... | yes _ = Manifest-Σ.proj₂ r
+... | no _ = Manifest-Σ.proj₁ r · ℓ
+
+------------------------------------------------------------------------
+-- With
+
+-- Sig With ℓ ≔ a is the signature Sig, but with the ℓ field set to a.
+
+mutual
+
+ infixl 5 _With_≔_
+
+ _With_≔_ : ∀ {s} (Sig : Signature s) (ℓ : Label) {ℓ∈ : ℓ ∈ Sig} →
+ ((r : Restricted Sig ℓ ℓ∈) → Proj Sig ℓ r) → Signature s
+ _With_≔_ ∅ ℓ {} a
+ Sig , ℓ′ ∶ A With ℓ ≔ a with ℓ ≟ ℓ′
+ ... | yes _ = Sig , ℓ′ ≔ a
+ ... | no _ = Sig With ℓ ≔ a , ℓ′ ∶ A ∘ drop-With
+ Sig , ℓ′ ≔ a′ With ℓ ≔ a with ℓ ≟ ℓ′
+ ... | yes _ = Sig , ℓ′ ≔ a
+ ... | no _ = Sig With ℓ ≔ a , ℓ′ ≔ a′ ∘ drop-With
+
+ drop-With : ∀ {s} {Sig : Signature s} {ℓ : Label} {ℓ∈ : ℓ ∈ Sig}
+ {a : (r : Restricted Sig ℓ ℓ∈) → Proj Sig ℓ r} →
+ Record (Sig With ℓ ≔ a) → Record Sig
+ drop-With {Sig = ∅} {ℓ∈ = ()} r
+ drop-With {Sig = Sig , ℓ′ ∶ A} {ℓ} (rec r) with ℓ ≟ ℓ′
+ ... | yes _ = rec (Manifest-Σ.proj₁ r , Manifest-Σ.proj₂ r)
+ ... | no _ = rec (drop-With (Σ.proj₁ r) , Σ.proj₂ r)
+ drop-With {Sig = Sig , ℓ′ ≔ a} {ℓ} (rec r) with ℓ ≟ ℓ′
+ ... | yes _ = rec (Manifest-Σ.proj₁ r ,)
+ ... | no _ = rec (drop-With (Manifest-Σ.proj₁ r) ,)
diff --git a/src/Reflection.agda b/src/Reflection.agda
index 4f18ede..62d7f54 100644
--- a/src/Reflection.agda
+++ b/src/Reflection.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Support for reflection
------------------------------------------------------------------------
@@ -6,12 +8,15 @@ module Reflection where
open import Data.Bool as Bool using (Bool); open Bool.Bool
open import Data.List using (List); open Data.List.List
-open import Data.Nat as Nat using (ℕ)
+open import Data.Nat using (ℕ) renaming (_≟_ to _≟-ℕ_)
+open import Data.Product
open import Function
open import Relation.Binary
open import Relation.Binary.PropositionalEquality
open import Relation.Binary.PropositionalEquality.TrustMe
open import Relation.Nullary
+open import Relation.Nullary.Decidable as Dec
+open import Relation.Nullary.Product
------------------------------------------------------------------------
-- Names
@@ -44,37 +49,66 @@ s₁ ≟-Name s₂ with s₁ == s₂
------------------------------------------------------------------------
-- Terms
--- Is the argument implicit? (Here true stands for implicit and false
--- for explicit.)
+-- Is the argument visible (explicit), hidden (implicit), or an
+-- instance argument?
-Implicit? = Bool
+data Visibility : Set where
+ visible hidden instance : Visibility
+
+{-# BUILTIN HIDING Visibility #-}
+{-# BUILTIN VISIBLE visible #-}
+{-# BUILTIN HIDDEN hidden #-}
+{-# BUILTIN INSTANCE instance #-}
+
+-- Arguments can be relevant or irrelevant.
+
+data Relevance : Set where
+ relevant irrelevant : Relevance
+
+{-# BUILTIN RELEVANCE Relevance #-}
+{-# BUILTIN RELEVANT relevant #-}
+{-# BUILTIN IRRELEVANT irrelevant #-}
-- Arguments.
data Arg A : Set where
- arg : (im? : Implicit?) (x : A) → Arg A
+ arg : (v : Visibility) (r : Relevance) (x : A) → Arg A
{-# BUILTIN ARG Arg #-}
{-# BUILTIN ARGARG arg #-}
-- Terms.
-data Term : Set where
- -- Variable applied to arguments.
- var : (x : ℕ) (args : List (Arg Term)) → Term
- -- Constructor applied to arguments.
- con : (c : Name) (args : List (Arg Term)) → Term
- -- Identifier applied to arguments.
- def : (f : Name) (args : List (Arg Term)) → Term
- -- Explicit or implicit λ abstraction.
- lam : (im? : Implicit?) (t : Term) → Term
- -- Pi-type.
- pi : (t₁ : Arg Term) (t₂ : Term) → Term
- -- An arbitrary sort (Set, for instance).
- sort : Term
- -- Anything.
- unknown : Term
-
+mutual
+ data Term : Set where
+ -- Variable applied to arguments.
+ var : (x : ℕ) (args : List (Arg Term)) → Term
+ -- Constructor applied to arguments.
+ con : (c : Name) (args : List (Arg Term)) → Term
+ -- Identifier applied to arguments.
+ def : (f : Name) (args : List (Arg Term)) → Term
+ -- Different kinds of λ-abstraction.
+ lam : (v : Visibility) (t : Term) → Term
+ -- Pi-type.
+ pi : (t₁ : Arg Type) (t₂ : Type) → Term
+ -- A sort.
+ sort : Sort → Term
+ -- Anything else.
+ unknown : Term
+
+ data Type : Set where
+ el : (s : Sort) (t : Term) → Type
+
+ data Sort : Set where
+ -- A Set of a given (possibly neutral) level.
+ set : (t : Term) → Sort
+ -- A Set of a given concrete level.
+ lit : (n : ℕ) → Sort
+ -- Anything else.
+ unknown : Sort
+
+{-# BUILTIN AGDASORT Sort #-}
+{-# BUILTIN AGDATYPE Type #-}
{-# BUILTIN AGDATERM Term #-}
{-# BUILTIN AGDATERMVAR var #-}
{-# BUILTIN AGDATERMCON con #-}
@@ -83,6 +117,64 @@ data Term : Set where
{-# BUILTIN AGDATERMPI pi #-}
{-# BUILTIN AGDATERMSORT sort #-}
{-# BUILTIN AGDATERMUNSUPPORTED unknown #-}
+{-# BUILTIN AGDATYPEEL el #-}
+{-# BUILTIN AGDASORTSET set #-}
+{-# BUILTIN AGDASORTLIT lit #-}
+{-# BUILTIN AGDASORTUNSUPPORTED unknown #-}
+
+------------------------------------------------------------------------
+-- Definitions
+
+postulate
+ -- Function definition.
+ Function : Set
+ -- Data type definition.
+ Data-type : Set
+ -- Record type definition.
+ Record : Set
+
+{-# BUILTIN AGDAFUNDEF Function #-}
+{-# BUILTIN AGDADATADEF Data-type #-}
+{-# BUILTIN AGDARECORDDEF Record #-}
+
+-- Definitions.
+
+data Definition : Set where
+ function : Function → Definition
+ data-type : Data-type → Definition
+ record′ : Record → Definition
+ constructor′ : Definition
+ axiom : Definition
+ primitive′ : Definition
+
+{-# BUILTIN AGDADEFINITION Definition #-}
+{-# BUILTIN AGDADEFINITIONFUNDEF function #-}
+{-# BUILTIN AGDADEFINITIONDATADEF data-type #-}
+{-# BUILTIN AGDADEFINITIONRECORDDEF record′ #-}
+{-# BUILTIN AGDADEFINITIONDATACONSTRUCTOR constructor′ #-}
+{-# BUILTIN AGDADEFINITIONPOSTULATE axiom #-}
+{-# BUILTIN AGDADEFINITIONPRIMITIVE primitive′ #-}
+
+private
+ primitive
+ primQNameType : Name → Type
+ primQNameDefinition : Name → Definition
+ primDataConstructors : Data-type → List Name
+
+-- The type of the thing with the given name.
+
+type : Name → Type
+type = primQNameType
+
+-- The definition of the thing with the given name.
+
+definition : Name → Definition
+definition = primQNameDefinition
+
+-- The constructors of the given data type.
+
+constructors : Data-type → List Name
+constructors = primDataConstructors
------------------------------------------------------------------------
-- Term equality is decidable
@@ -91,14 +183,23 @@ data Term : Set where
private
- arg₁ : ∀ {A im? im?′} {x x′ : A} →
- arg im? x ≡ arg im?′ x′ → im? ≡ im?′
+ cong₂′ : ∀ {A B C : Set} (f : A → B → C) {x y u v} →
+ x ≡ y × u ≡ v → f x u ≡ f y v
+ cong₂′ f = uncurry (cong₂ f)
+
+ cong₃′ : ∀ {A B C D : Set} (f : A → B → C → D) {x y u v r s} →
+ x ≡ y × u ≡ v × r ≡ s → f x u r ≡ f y v s
+ cong₃′ f (refl , refl , refl) = refl
+
+ arg₁ : ∀ {A v v′ r r′} {x x′ : A} → arg v r x ≡ arg v′ r′ x′ → v ≡ v′
arg₁ refl = refl
- arg₂ : ∀ {A im? im?′} {x x′ : A} →
- arg im? x ≡ arg im?′ x′ → x ≡ x′
+ arg₂ : ∀ {A v v′ r r′} {x x′ : A} → arg v r x ≡ arg v′ r′ x′ → r ≡ r′
arg₂ refl = refl
+ arg₃ : ∀ {A v v′ r r′} {x x′ : A} → arg v r x ≡ arg v′ r′ x′ → x ≡ x′
+ arg₃ refl = refl
+
cons₁ : ∀ {A : Set} {x y} {xs ys : List A} → x ∷ xs ≡ y ∷ ys → x ≡ y
cons₁ refl = refl
@@ -123,10 +224,10 @@ private
def₂ : ∀ {f f′ args args′} → def f args ≡ def f′ args′ → args ≡ args′
def₂ refl = refl
- lam₁ : ∀ {im? im?′ t t′} → lam im? t ≡ lam im?′ t′ → im? ≡ im?′
+ lam₁ : ∀ {v v′ t t′} → lam v t ≡ lam v′ t′ → v ≡ v′
lam₁ refl = refl
- lam₂ : ∀ {im? im?′ t t′} → lam im? t ≡ lam im?′ t′ → t ≡ t′
+ lam₂ : ∀ {v v′ t t′} → lam v t ≡ lam v′ t′ → t ≡ t′
lam₂ refl = refl
pi₁ : ∀ {t₁ t₁′ t₂ t₂′} → pi t₁ t₂ ≡ pi t₁′ t₂′ → t₁ ≡ t₁′
@@ -135,88 +236,123 @@ private
pi₂ : ∀ {t₁ t₁′ t₂ t₂′} → pi t₁ t₂ ≡ pi t₁′ t₂′ → t₂ ≡ t₂′
pi₂ refl = refl
-mutual
+ sort₁ : ∀ {x y} → sort x ≡ sort y → x ≡ y
+ sort₁ refl = refl
- infix 4 _≟-Arg_ _≟-Args_ _≟_
+ set₁ : ∀ {x y} → set x ≡ set y → x ≡ y
+ set₁ refl = refl
- _≟-Arg_ : Decidable (_≡_ {A = Arg Term})
- arg e a ≟-Arg arg e′ a′ with Bool._≟_ e e′ | a ≟ a′
- arg e a ≟-Arg arg .e .a | yes refl | yes refl = yes refl
- arg e a ≟-Arg arg e′ a′ | no e≢e′ | _ = no (e≢e′ ∘ arg₁)
- arg e a ≟-Arg arg e′ a′ | _ | no a≢a′ = no (a≢a′ ∘ arg₂)
+ lit₁ : ∀ {x y} → lit x ≡ lit y → x ≡ y
+ lit₁ refl = refl
+
+ el₁ : ∀ {s s′ t t′} → el s t ≡ el s′ t′ → s ≡ s′
+ el₁ refl = refl
+
+ el₂ : ∀ {s s′ t t′} → el s t ≡ el s′ t′ → t ≡ t′
+ el₂ refl = refl
+
+_≟-Visibility_ : Decidable (_≡_ {A = Visibility})
+visible ≟-Visibility visible = yes refl
+hidden ≟-Visibility hidden = yes refl
+instance ≟-Visibility instance = yes refl
+visible ≟-Visibility hidden = no λ()
+visible ≟-Visibility instance = no λ()
+hidden ≟-Visibility visible = no λ()
+hidden ≟-Visibility instance = no λ()
+instance ≟-Visibility visible = no λ()
+instance ≟-Visibility hidden = no λ()
+
+_≟-Relevance_ : Decidable (_≡_ {A = Relevance})
+relevant ≟-Relevance relevant = yes refl
+irrelevant ≟-Relevance irrelevant = yes refl
+relevant ≟-Relevance irrelevant = no λ()
+irrelevant ≟-Relevance relevant = no λ()
+
+mutual
+ infix 4 _≟_ _≟-Args_ _≟-ArgType_
+
+ _≟-ArgTerm_ : Decidable (_≡_ {A = Arg Term})
+ arg e r a ≟-ArgTerm arg e′ r′ a′ =
+ Dec.map′ (cong₃′ arg)
+ < arg₁ , < arg₂ , arg₃ > >
+ (e ≟-Visibility e′ ×-dec r ≟-Relevance r′ ×-dec a ≟ a′)
+
+ _≟-ArgType_ : Decidable (_≡_ {A = Arg Type})
+ arg e r a ≟-ArgType arg e′ r′ a′ =
+ Dec.map′ (cong₃′ arg)
+ < arg₁ , < arg₂ , arg₃ > >
+ (e ≟-Visibility e′ ×-dec
+ r ≟-Relevance r′ ×-dec
+ a ≟-Type a′)
_≟-Args_ : Decidable (_≡_ {A = List (Arg Term)})
- [] ≟-Args [] = yes refl
- (a ∷ args) ≟-Args (a′ ∷ args′) with a ≟-Arg a′ | args ≟-Args args′
- (a ∷ args) ≟-Args (.a ∷ .args) | yes refl | yes refl = yes refl
- (a ∷ args) ≟-Args (a′ ∷ args′) | no a≢a′ | _ = no (a≢a′ ∘ cons₁)
- (a ∷ args) ≟-Args (a′ ∷ args′) | _ | no args≢args′ = no (args≢args′ ∘ cons₂)
- [] ≟-Args (_ ∷ _) = no λ()
- (_ ∷ _) ≟-Args [] = no λ()
+ [] ≟-Args [] = yes refl
+ (x ∷ xs) ≟-Args (y ∷ ys) = Dec.map′ (cong₂′ _∷_) < cons₁ , cons₂ > (x ≟-ArgTerm y ×-dec xs ≟-Args ys)
+ [] ≟-Args (_ ∷ _) = no λ()
+ (_ ∷ _) ≟-Args [] = no λ()
_≟_ : Decidable (_≡_ {A = Term})
- var x args ≟ var x′ args′ with Nat._≟_ x x′ | args ≟-Args args′
- var x args ≟ var .x .args | yes refl | yes refl = yes refl
- var x args ≟ var x′ args′ | no x≢x′ | _ = no (x≢x′ ∘ var₁)
- var x args ≟ var x′ args′ | _ | no args≢args′ = no (args≢args′ ∘ var₂)
- con c args ≟ con c′ args′ with c ≟-Name c′ | args ≟-Args args′
- con c args ≟ con .c .args | yes refl | yes refl = yes refl
- con c args ≟ con c′ args′ | no c≢c′ | _ = no (c≢c′ ∘ con₁)
- con c args ≟ con c′ args′ | _ | no args≢args′ = no (args≢args′ ∘ con₂)
- def f args ≟ def f′ args′ with f ≟-Name f′ | args ≟-Args args′
- def f args ≟ def .f .args | yes refl | yes refl = yes refl
- def f args ≟ def f′ args′ | no f≢f′ | _ = no (f≢f′ ∘ def₁)
- def f args ≟ def f′ args′ | _ | no args≢args′ = no (args≢args′ ∘ def₂)
- lam im? t ≟ lam im?′ t′ with Bool._≟_ im? im?′ | t ≟ t′
- lam im? t ≟ lam .im? .t | yes refl | yes refl = yes refl
- lam im? t ≟ lam im?′ t′ | no im?≢im?′ | _ = no (im?≢im?′ ∘ lam₁)
- lam im? t ≟ lam im?′ t′ | _ | no t≢t′ = no (t≢t′ ∘ lam₂)
- pi t₁ t₂ ≟ pi t₁′ t₂′ with t₁ ≟-Arg t₁′ | t₂ ≟ t₂′
- pi t₁ t₂ ≟ pi .t₁ .t₂ | yes refl | yes refl = yes refl
- pi t₁ t₂ ≟ pi t₁′ t₂′ | no t₁≢t₁′ | _ = no (t₁≢t₁′ ∘ pi₁)
- pi t₁ t₂ ≟ pi t₁′ t₂′ | _ | no t₂≢t₂′ = no (t₂≢t₂′ ∘ pi₂)
- sort ≟ sort = yes refl
+ var x args ≟ var x′ args′ = Dec.map′ (cong₂′ var) < var₁ , var₂ > (x ≟-ℕ x′ ×-dec args ≟-Args args′)
+ con c args ≟ con c′ args′ = Dec.map′ (cong₂′ con) < con₁ , con₂ > (c ≟-Name c′ ×-dec args ≟-Args args′)
+ def f args ≟ def f′ args′ = Dec.map′ (cong₂′ def) < def₁ , def₂ > (f ≟-Name f′ ×-dec args ≟-Args args′)
+ lam v t ≟ lam v′ t′ = Dec.map′ (cong₂′ lam) < lam₁ , lam₂ > (v ≟-Visibility v′ ×-dec t ≟ t′)
+ pi t₁ t₂ ≟ pi t₁′ t₂′ = Dec.map′ (cong₂′ pi) < pi₁ , pi₂ > (t₁ ≟-ArgType t₁′ ×-dec t₂ ≟-Type t₂′)
+ sort s ≟ sort s′ = Dec.map′ (cong sort) sort₁ (s ≟-Sort s′)
unknown ≟ unknown = yes refl
- var x args ≟ con c args′ = no λ()
- var x args ≟ def f args′ = no λ()
- var x args ≟ lam im? t = no λ()
- var x args ≟ pi t₁ t₂ = no λ()
- var x args ≟ sort = no λ()
- var x args ≟ unknown = no λ()
- con c args ≟ var x args′ = no λ()
- con c args ≟ def f args′ = no λ()
- con c args ≟ lam im? t = no λ()
- con c args ≟ pi t₁ t₂ = no λ()
- con c args ≟ sort = no λ()
- con c args ≟ unknown = no λ()
- def f args ≟ var x args′ = no λ()
- def f args ≟ con c args′ = no λ()
- def f args ≟ lam im? t = no λ()
- def f args ≟ pi t₁ t₂ = no λ()
- def f args ≟ sort = no λ()
- def f args ≟ unknown = no λ()
- lam im? t ≟ var x args = no λ()
- lam im? t ≟ con c args = no λ()
- lam im? t ≟ def f args = no λ()
- lam im? t ≟ pi t₁ t₂ = no λ()
- lam im? t ≟ sort = no λ()
- lam im? t ≟ unknown = no λ()
- pi t₁ t₂ ≟ var x args = no λ()
- pi t₁ t₂ ≟ con c args = no λ()
- pi t₁ t₂ ≟ def f args = no λ()
- pi t₁ t₂ ≟ lam im? t = no λ()
- pi t₁ t₂ ≟ sort = no λ()
- pi t₁ t₂ ≟ unknown = no λ()
- sort ≟ var x args = no λ()
- sort ≟ con c args = no λ()
- sort ≟ def f args = no λ()
- sort ≟ lam im? t = no λ()
- sort ≟ pi t₁ t₂ = no λ()
- sort ≟ unknown = no λ()
- unknown ≟ var x args = no λ()
- unknown ≟ con c args = no λ()
- unknown ≟ def f args = no λ()
- unknown ≟ lam im? t = no λ()
- unknown ≟ pi t₁ t₂ = no λ()
- unknown ≟ sort = no λ()
+ var x args ≟ con c args′ = no λ()
+ var x args ≟ def f args′ = no λ()
+ var x args ≟ lam v t = no λ()
+ var x args ≟ pi t₁ t₂ = no λ()
+ var x args ≟ sort _ = no λ()
+ var x args ≟ unknown = no λ()
+ con c args ≟ var x args′ = no λ()
+ con c args ≟ def f args′ = no λ()
+ con c args ≟ lam v t = no λ()
+ con c args ≟ pi t₁ t₂ = no λ()
+ con c args ≟ sort _ = no λ()
+ con c args ≟ unknown = no λ()
+ def f args ≟ var x args′ = no λ()
+ def f args ≟ con c args′ = no λ()
+ def f args ≟ lam v t = no λ()
+ def f args ≟ pi t₁ t₂ = no λ()
+ def f args ≟ sort _ = no λ()
+ def f args ≟ unknown = no λ()
+ lam v t ≟ var x args = no λ()
+ lam v t ≟ con c args = no λ()
+ lam v t ≟ def f args = no λ()
+ lam v t ≟ pi t₁ t₂ = no λ()
+ lam v t ≟ sort _ = no λ()
+ lam v t ≟ unknown = no λ()
+ pi t₁ t₂ ≟ var x args = no λ()
+ pi t₁ t₂ ≟ con c args = no λ()
+ pi t₁ t₂ ≟ def f args = no λ()
+ pi t₁ t₂ ≟ lam v t = no λ()
+ pi t₁ t₂ ≟ sort _ = no λ()
+ pi t₁ t₂ ≟ unknown = no λ()
+ sort _ ≟ var x args = no λ()
+ sort _ ≟ con c args = no λ()
+ sort _ ≟ def f args = no λ()
+ sort _ ≟ lam v t = no λ()
+ sort _ ≟ pi t₁ t₂ = no λ()
+ sort _ ≟ unknown = no λ()
+ unknown ≟ var x args = no λ()
+ unknown ≟ con c args = no λ()
+ unknown ≟ def f args = no λ()
+ unknown ≟ lam v t = no λ()
+ unknown ≟ pi t₁ t₂ = no λ()
+ unknown ≟ sort _ = no λ()
+
+ _≟-Type_ : Decidable (_≡_ {A = Type})
+ el s t ≟-Type el s′ t′ = Dec.map′ (cong₂′ el) < el₁ , el₂ > (s ≟-Sort s′ ×-dec t ≟ t′)
+
+ _≟-Sort_ : Decidable (_≡_ {A = Sort})
+ set t ≟-Sort set t′ = Dec.map′ (cong set) set₁ (t ≟ t′)
+ lit n ≟-Sort lit n′ = Dec.map′ (cong lit) lit₁ (n ≟-ℕ n′)
+ unknown ≟-Sort unknown = yes refl
+ set _ ≟-Sort lit _ = no λ()
+ set _ ≟-Sort unknown = no λ()
+ lit _ ≟-Sort set _ = no λ()
+ lit _ ≟-Sort unknown = no λ()
+ unknown ≟-Sort set _ = no λ()
+ unknown ≟-Sort lit _ = no λ()
diff --git a/src/Relation/Binary.agda b/src/Relation/Binary.agda
index c5aaed4..c49ea00 100644
--- a/src/Relation/Binary.agda
+++ b/src/Relation/Binary.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties of homogeneous binary relations
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Binary where
open import Data.Product
diff --git a/src/Relation/Binary/Consequences.agda b/src/Relation/Binary/Consequences.agda
index d07c2a2..0f36c7b 100644
--- a/src/Relation/Binary/Consequences.agda
+++ b/src/Relation/Binary/Consequences.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some properties imply others
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Binary.Consequences where
open import Relation.Binary.Core hiding (refl)
diff --git a/src/Relation/Binary/Consequences/Core.agda b/src/Relation/Binary/Consequences/Core.agda
index 80759f1..515c0dd 100644
--- a/src/Relation/Binary/Consequences/Core.agda
+++ b/src/Relation/Binary/Consequences/Core.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some properties imply others
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- This file contains some core definitions which are reexported by
-- Relation.Binary.Consequences.
diff --git a/src/Relation/Binary/Core.agda b/src/Relation/Binary/Core.agda
index 472ca58..47a207a 100644
--- a/src/Relation/Binary/Core.agda
+++ b/src/Relation/Binary/Core.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties of homogeneous binary relations
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- This file contains some core definitions which are reexported by
-- Relation.Binary or Relation.Binary.PropositionalEquality.
diff --git a/src/Relation/Binary/EqReasoning.agda b/src/Relation/Binary/EqReasoning.agda
index 44321d1..d56716d 100644
--- a/src/Relation/Binary/EqReasoning.agda
+++ b/src/Relation/Binary/EqReasoning.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Convenient syntax for equational reasoning
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- Example use:
-- n*0≡0 : ∀ n → n * 0 ≡ 0
diff --git a/src/Relation/Binary/Flip.agda b/src/Relation/Binary/Flip.agda
index 74b1300..9d52c36 100644
--- a/src/Relation/Binary/Flip.agda
+++ b/src/Relation/Binary/Flip.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Many properties which hold for _∼_ also hold for flip _∼_
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Relation.Binary.Flip where
diff --git a/src/Relation/Binary/HeterogeneousEquality.agda b/src/Relation/Binary/HeterogeneousEquality.agda
index b82e3ce..29c2ffb 100644
--- a/src/Relation/Binary/HeterogeneousEquality.agda
+++ b/src/Relation/Binary/HeterogeneousEquality.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Heterogeneous equality
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Binary.HeterogeneousEquality where
open import Data.Product
@@ -14,16 +14,16 @@ 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)
+open import Relation.Binary.PropositionalEquality as P using (_≡_; refl)
+
+import Relation.Binary.HeterogeneousEquality.Core as Core
------------------------------------------------------------------------
-- Heterogeneous equality
-infix 4 _≅_ _≇_
+infix 4 _≇_
-data _≅_ {a} {A : Set a} (x : A) : ∀ {b} {B : Set b} → B → Set where
- refl : x ≅ x
+open Core public using (_≅_; refl)
-- Nonequality.
@@ -36,8 +36,7 @@ x ≇ y = ¬ x ≅ y
≡-to-≅ : ∀ {a} {A : Set a} {x y : A} → x ≡ y → x ≅ y
≡-to-≅ refl = refl
-≅-to-≡ : ∀ {a} {A : Set a} {x y : A} → x ≅ y → x ≡ y
-≅-to-≡ refl = refl
+open Core public using (≅-to-≡)
------------------------------------------------------------------------
-- Some properties
@@ -67,7 +66,7 @@ subst-removable P refl z = refl
≡-subst-removable : ∀ {a p} {A : Set a}
(P : A → Set p) {x y} (eq : x ≡ y) z →
- PropEq.subst P eq z ≅ z
+ P.subst P eq z ≅ z
≡-subst-removable P refl z = refl
cong : ∀ {a b} {A : Set a} {B : A → Set b} {x y}
@@ -112,9 +111,9 @@ indexedSetoid B = record
}
}
-≡⇿≅ : ∀ {a b} {A : Set a} (B : A → Set b) {x : A} →
- Inverse (PropEq.setoid (B x)) (indexedSetoid B at x)
-≡⇿≅ B = record
+≡↔≅ : ∀ {a b} {A : Set a} (B : A → Set b) {x : A} →
+ Inverse (P.setoid (B x)) (indexedSetoid B at x)
+≡↔≅ B = record
{ to = record { _⟨$⟩_ = id; cong = ≡-to-≅ }
; from = record { _⟨$⟩_ = id; cong = ≅-to-≡ }
; inverse-of = record
@@ -145,7 +144,7 @@ isPreorder = record
isPreorder-≡ : ∀ {a} {A : Set a} →
IsPreorder {A = A} _≡_ (λ x y → x ≅ y)
isPreorder-≡ = record
- { isEquivalence = PropEq.isEquivalence
+ { isEquivalence = P.isEquivalence
; reflexive = reflexive
; trans = trans
}
@@ -159,17 +158,6 @@ preorder A = record
}
------------------------------------------------------------------------
--- The inspect idiom
-
--- See Relation.Binary.PropositionalEquality.Inspect.
-
-data Inspect {a} {A : Set a} (x : A) : Set a where
- _with-≅_ : (y : A) (eq : y ≅ x) → Inspect x
-
-inspect : ∀ {a} {A : Set a} (x : A) → Inspect x
-inspect x = x with-≅ refl
-
-------------------------------------------------------------------------
-- Convenient syntax for equational reasoning
module ≅-Reasoning where
@@ -201,3 +189,25 @@ module ≅-Reasoning where
_∎ : ∀ {a} {A : Set a} (x : A) → x IsRelatedTo x
_∎ _ = relTo refl
+
+------------------------------------------------------------------------
+-- Functional extensionality
+
+-- A form of functional extensionality for _≅_.
+
+Extensionality : (a b : Level) → Set _
+Extensionality a b =
+ {A : Set a} {B₁ B₂ : A → Set b}
+ {f₁ : (x : A) → B₁ x} {f₂ : (x : A) → B₂ x} →
+ (∀ x → B₁ x ≡ B₂ x) → (∀ x → f₁ x ≅ f₂ x) → f₁ ≅ f₂
+
+-- This form of extensionality follows from extensionality for _≡_.
+
+≡-ext-to-≅-ext : ∀ {ℓ₁ ℓ₂} →
+ P.Extensionality ℓ₁ (suc ℓ₂) → Extensionality ℓ₁ ℓ₂
+≡-ext-to-≅-ext ext B₁≡B₂ f₁≅f₂ with ext B₁≡B₂
+≡-ext-to-≅-ext {ℓ₁} {ℓ₂} ext B₁≡B₂ f₁≅f₂ | P.refl =
+ ≡-to-≅ $ ext′ (≅-to-≡ ∘ f₁≅f₂)
+ where
+ ext′ : P.Extensionality ℓ₁ ℓ₂
+ ext′ = P.extensionality-for-lower-levels ℓ₁ (suc ℓ₂) ext
diff --git a/src/Relation/Binary/HeterogeneousEquality/Core.agda b/src/Relation/Binary/HeterogeneousEquality/Core.agda
new file mode 100644
index 0000000..6a2f24b
--- /dev/null
+++ b/src/Relation/Binary/HeterogeneousEquality/Core.agda
@@ -0,0 +1,26 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Heterogeneous equality
+------------------------------------------------------------------------
+
+-- This file contains some core definitions which are reexported by
+-- Relation.Binary.HeterogeneousEquality.
+
+module Relation.Binary.HeterogeneousEquality.Core where
+
+open import Relation.Binary.Core using (_≡_; refl)
+
+------------------------------------------------------------------------
+-- Heterogeneous equality
+
+infix 4 _≅_
+
+data _≅_ {a} {A : Set a} (x : A) : ∀ {b} {B : Set b} → B → Set where
+ refl : x ≅ x
+
+------------------------------------------------------------------------
+-- Conversion
+
+≅-to-≡ : ∀ {a} {A : Set a} {x y : A} → x ≅ y → x ≡ y
+≅-to-≡ refl = refl
diff --git a/src/Relation/Binary/Indexed.agda b/src/Relation/Binary/Indexed.agda
index fdb750a..603fe97 100644
--- a/src/Relation/Binary/Indexed.agda
+++ b/src/Relation/Binary/Indexed.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Indexed binary relations
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Binary.Indexed where
import Relation.Binary as B
diff --git a/src/Relation/Binary/Indexed/Core.agda b/src/Relation/Binary/Indexed/Core.agda
index ecc3b03..c6f1332 100644
--- a/src/Relation/Binary/Indexed/Core.agda
+++ b/src/Relation/Binary/Indexed/Core.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Indexed binary relations
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- This file contains some core definitions which are reexported by
-- Relation.Binary.Indexed.
diff --git a/src/Relation/Binary/InducedPreorders.agda b/src/Relation/Binary/InducedPreorders.agda
index 057421e..53d343b 100644
--- a/src/Relation/Binary/InducedPreorders.agda
+++ b/src/Relation/Binary/InducedPreorders.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Induced preorders
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Relation.Binary.InducedPreorders
diff --git a/src/Relation/Binary/List/NonStrictLex.agda b/src/Relation/Binary/List/NonStrictLex.agda
index 4ba9842..ed023b5 100644
--- a/src/Relation/Binary/List/NonStrictLex.agda
+++ b/src/Relation/Binary/List/NonStrictLex.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Lexicographic ordering of lists
------------------------------------------------------------------------
diff --git a/src/Relation/Binary/List/Pointwise.agda b/src/Relation/Binary/List/Pointwise.agda
index 04a4fe0..243a63a 100644
--- a/src/Relation/Binary/List/Pointwise.agda
+++ b/src/Relation/Binary/List/Pointwise.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Pointwise lifting of relations to lists
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Binary.List.Pointwise where
open import Function
diff --git a/src/Relation/Binary/List/StrictLex.agda b/src/Relation/Binary/List/StrictLex.agda
index 7530895..870c427 100644
--- a/src/Relation/Binary/List/StrictLex.agda
+++ b/src/Relation/Binary/List/StrictLex.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Lexicographic ordering of lists
------------------------------------------------------------------------
diff --git a/src/Relation/Binary/NonStrictToStrict.agda b/src/Relation/Binary/NonStrictToStrict.agda
index 2ec6806..73e2eed 100644
--- a/src/Relation/Binary/NonStrictToStrict.agda
+++ b/src/Relation/Binary/NonStrictToStrict.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Conversion of ≤ to <, along with a number of properties
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- Possible TODO: Prove that a conversion ≤ → < → ≤ returns a
-- relation equivalent to the original one (and similarly for
-- < → ≤ → <).
diff --git a/src/Relation/Binary/On.agda b/src/Relation/Binary/On.agda
index 6e5f7d9..8c86599 100644
--- a/src/Relation/Binary/On.agda
+++ b/src/Relation/Binary/On.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Many properties which hold for _∼_ also hold for _∼_ on f
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Relation.Binary.On {a b} {A : Set a} {B : Set b}
diff --git a/src/Relation/Binary/OrderMorphism.agda b/src/Relation/Binary/OrderMorphism.agda
index 0c2f981..706742d 100644
--- a/src/Relation/Binary/OrderMorphism.agda
+++ b/src/Relation/Binary/OrderMorphism.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Order morphisms
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Binary.OrderMorphism where
open import Relation.Binary
diff --git a/src/Relation/Binary/PartialOrderReasoning.agda b/src/Relation/Binary/PartialOrderReasoning.agda
index 28bc30b..e337b04 100644
--- a/src/Relation/Binary/PartialOrderReasoning.agda
+++ b/src/Relation/Binary/PartialOrderReasoning.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Convenient syntax for "equational reasoning" using a partial order
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Relation.Binary.PartialOrderReasoning
diff --git a/src/Relation/Binary/PreorderReasoning.agda b/src/Relation/Binary/PreorderReasoning.agda
index 4601b41..b101199 100644
--- a/src/Relation/Binary/PreorderReasoning.agda
+++ b/src/Relation/Binary/PreorderReasoning.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Convenient syntax for "equational reasoning" using a preorder
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- I think that the idea behind this library is originally Ulf
-- Norell's. I have adapted it to my tastes and mixfix operators,
-- though.
diff --git a/src/Relation/Binary/Product/NonStrictLex.agda b/src/Relation/Binary/Product/NonStrictLex.agda
index 6467a24..36724bd 100644
--- a/src/Relation/Binary/Product/NonStrictLex.agda
+++ b/src/Relation/Binary/Product/NonStrictLex.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- 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.
diff --git a/src/Relation/Binary/Product/Pointwise.agda b/src/Relation/Binary/Product/Pointwise.agda
index 38d167d..09f5431 100644
--- a/src/Relation/Binary/Product/Pointwise.agda
+++ b/src/Relation/Binary/Product/Pointwise.agda
@@ -1,24 +1,27 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Pointwise products of binary relations
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Binary.Product.Pointwise where
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.Equality as F using (_⟶_; _⟨$⟩_)
open import Function.Equivalence as Eq
- using (Equivalent; _⇔_; module Equivalent)
- renaming (_∘_ to _⟨∘⟩_)
+ using (Equivalence; _⇔_; module Equivalence)
+open import Function.Injection as Inj
+ using (Injection; _↣_; module Injection)
open import Function.Inverse as Inv
- using (Inverse; _⇿_; module Inverse)
- renaming (_∘_ to _⟪∘⟫_)
-open import Function.LeftInverse
- using (_LeftInverseOf_; _RightInverseOf_)
+ using (Inverse; _↔_; module Inverse)
+open import Function.LeftInverse as LeftInv
+ using (LeftInverse; _↞_; _LeftInverseOf_; module LeftInverse)
+open import Function.Related
+open import Function.Surjection as Surj
+ using (Surjection; _↠_; module Surjection)
open import Level
open import Relation.Nullary.Product
open import Relation.Binary
@@ -248,11 +251,11 @@ s₁ ×-strictPartialOrder s₂ = record
} where open StrictPartialOrder
------------------------------------------------------------------------
--- Some properties related to equivalences and inverses
+-- Some properties related to "relatedness"
-×-Rel⇿≡ : ∀ {a b} {A : Set a} {B : Set b} →
+×-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
+×-Rel↔≡ = record
{ to = record { _⟨$⟩_ = id; cong = to-cong }
; from = record { _⟨$⟩_ = id; cong = from-cong }
; inverse-of = record
@@ -262,82 +265,156 @@ s₁ ×-strictPartialOrder s₂ = record
}
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
+ to-cong {i = (x , y)} {j = (.x , .y)} (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} }
+_×-⟶_ :
+ ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈}
+ {A : Setoid s₁ s₂} {B : Setoid s₃ s₄}
+ {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} →
+ A ⟶ B → C ⟶ D → (A ×-setoid C) ⟶ (B ×-setoid D)
+_×-⟶_ {A = A} {B} {C} {D} f g = record
+ { _⟨$⟩_ = fg
+ ; cong = fg-cong
}
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) _∼₂_)
+ fg = Prod.map (_⟨$⟩_ f) (_⟨$⟩_ g)
- from = Prod.map (_⟨$⟩_ (Equivalent.from A⇔B))
- (_⟨$⟩_ (Equivalent.from C⇔D))
+ fg-cong : _≈AC_ =[ fg ]⇒ _≈BD_
+ fg-cong (_∼₁_ , _∼₂_) = (F.cong f _∼₁_ , F.cong g _∼₂_)
- from-cong : _≈BD_ =[ from ]⇒ _≈AC_
- from-cong (_∼₁_ , _∼₂_) =
- (F.cong (Equivalent.from A⇔B) _∼₁_ ,
- F.cong (Equivalent.from C⇔D) _∼₂_)
+_×-equivalence_ :
+ ∀ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
+ {A : Setoid a₁ a₂} {B : Setoid b₁ b₂}
+ {C : Setoid c₁ c₂} {D : Setoid d₁ d₂} →
+ Equivalence A B → Equivalence C D →
+ Equivalence (A ×-setoid C) (B ×-setoid D)
+_×-equivalence_ {A = A} {B} {C} {D} A⇔B C⇔D = record
+ { to = to A⇔B ×-⟶ to C⇔D
+ ; from = from A⇔B ×-⟶ from C⇔D
+ } where open Equivalence
_×-⇔_ : ∀ {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.equivalence (×-Rel↔≡ {A = B} {B = D}) ⟨∘⟩
+ A⇔B ×-equivalence C⇔D ⟨∘⟩
+ Eq.sym (Inverse.equivalence (×-Rel↔≡ {A = A} {B = C}))
+ where open Eq using () renaming (_∘_ to _⟨∘⟩_)
+
+_×-injection_ :
+ ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈}
+ {A : Setoid s₁ s₂} {B : Setoid s₃ s₄}
+ {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} →
+ Injection A B → Injection C D →
+ Injection (A ×-setoid C) (B ×-setoid D)
+A↣B ×-injection C↣D = record
+ { to = to A↣B ×-⟶ to C↣D
+ ; injective = Prod.map (injective A↣B) (injective C↣D)
+ } where open Injection
+
+_×-↣_ : ∀ {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.injection (×-Rel↔≡ {A = B} {B = D}) ⟨∘⟩
+ A↣B ×-injection C↣D ⟨∘⟩
+ Inverse.injection (Inv.sym (×-Rel↔≡ {A = A} {B = C}))
+ where open Inj using () renaming (_∘_ to _⟨∘⟩_)
+
+_×-left-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₈} →
+ LeftInverse A B → LeftInverse C D →
+ LeftInverse (A ×-setoid C) (B ×-setoid D)
+A↞B ×-left-inverse C↞D = record
+ { to = Equivalence.to eq
+ ; from = Equivalence.from eq
+ ; left-inverse-of = left
+ }
+ where
+ open LeftInverse
+ eq = LeftInverse.equivalence A↞B ×-equivalence
+ LeftInverse.equivalence C↞D
+
+ left : Equivalence.from eq LeftInverseOf Equivalence.to eq
+ left (x , y) = ( left-inverse-of A↞B x
+ , left-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 =
+ Inverse.left-inverse (×-Rel↔≡ {A = B} {B = D}) ⟨∘⟩
+ A↞B ×-left-inverse C↞D ⟨∘⟩
+ Inverse.left-inverse (Inv.sym (×-Rel↔≡ {A = A} {B = C}))
+ where open LeftInv using () renaming (_∘_ to _⟨∘⟩_)
+
+_×-surjection_ :
+ ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈}
+ {A : Setoid s₁ s₂} {B : Setoid s₃ s₄}
+ {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} →
+ Surjection A B → Surjection C D →
+ Surjection (A ×-setoid C) (B ×-setoid D)
+A↠B ×-surjection C↠D = record
+ { to = LeftInverse.from inv
+ ; surjective = record
+ { from = LeftInverse.to inv
+ ; right-inverse-of = LeftInverse.left-inverse-of inv
+ }
+ }
+ where
+ open Surjection
+ inv = right-inverse A↠B ×-left-inverse right-inverse 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.surjection (×-Rel↔≡ {A = B} {B = D}) ⟨∘⟩
+ A↠B ×-surjection C↠D ⟨∘⟩
+ Inverse.surjection (Inv.sym (×-Rel↔≡ {A = A} {B = C}))
+ where open Surj using () renaming (_∘_ to _⟨∘⟩_)
_×-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₂} →
+ ∀ {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
+A↔B ×-inverse C↔D = record
+ { to = Surjection.to surj
+ ; from = Surjection.from surj
; inverse-of = record
- { left-inverse-of = left
- ; right-inverse-of = right
+ { left-inverse-of = LeftInverse.left-inverse-of inv
+ ; right-inverse-of = Surjection.right-inverse-of surj
}
}
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})
+ open Inverse
+ surj = Inverse.surjection A↔B ×-surjection
+ Inverse.surjection C↔D
+ inv = Inverse.left-inverse A↔B ×-left-inverse
+ Inverse.left-inverse 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↔≡ {A = B} {B = D} ⟨∘⟩
+ A↔B ×-inverse C↔D ⟨∘⟩
+ Inv.sym (×-Rel↔≡ {A = A} {B = C})
+ where open Inv using () renaming (_∘_ to _⟨∘⟩_)
+
+_×-cong_ : ∀ {k a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
+ A ∼[ k ] B → C ∼[ k ] D → (A × C) ∼[ k ] (B × D)
+_×-cong_ {implication} = λ f g → Prod.map f g
+_×-cong_ {reverse-implication} = λ f g → lam (Prod.map (app-← f) (app-← g))
+_×-cong_ {equivalence} = _×-⇔_
+_×-cong_ {injection} = _×-↣_
+_×-cong_ {reverse-injection} = λ f g → lam (app-↢ f ×-↣ app-↢ g)
+_×-cong_ {left-inverse} = _×-↞_
+_×-cong_ {surjection} = _×-↠_
+_×-cong_ {bijection} = _×-↔_
diff --git a/src/Relation/Binary/Product/StrictLex.agda b/src/Relation/Binary/Product/StrictLex.agda
index 42755b6..1b73943 100644
--- a/src/Relation/Binary/Product/StrictLex.agda
+++ b/src/Relation/Binary/Product/StrictLex.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- 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.
@@ -77,12 +77,9 @@ private
antisym {x} {y}
where
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) =
- ⊥-elim {Whatever = _ × _} $ irrefl₁ (sym₁ $ proj₁ y≈≤x) x₁<y₁
- antisym (inj₂ x≈≤y) (inj₁ y₁<x₁) =
- ⊥-elim {Whatever = _ × _} $ irrefl₁ (sym₁ $ proj₁ x≈≤y) y₁<x₁
+ antisym (inj₁ x₁<y₁) (inj₁ y₁<x₁) = ⊥-elim $ asym₁ x₁<y₁ y₁<x₁
+ antisym (inj₁ x₁<y₁) (inj₂ y≈≤x) = ⊥-elim $ irrefl₁ (sym₁ $ proj₁ y≈≤x) x₁<y₁
+ antisym (inj₂ x≈≤y) (inj₁ y₁<x₁) = ⊥-elim $ irrefl₁ (sym₁ $ proj₁ x≈≤y) y₁<x₁
antisym (inj₂ x≈≤y) (inj₂ y≈≤x) =
proj₁ x≈≤y , antisym₂ (proj₂ x≈≤y) (proj₂ y≈≤x)
diff --git a/src/Relation/Binary/PropositionalEquality.agda b/src/Relation/Binary/PropositionalEquality.agda
index 680a3b1..5805877 100644
--- a/src/Relation/Binary/PropositionalEquality.agda
+++ b/src/Relation/Binary/PropositionalEquality.agda
@@ -1,18 +1,20 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Propositional (intensional) equality
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Binary.PropositionalEquality where
open import Function
open import Function.Equality using (Π; _⟶_; ≡-setoid)
open import Data.Product
+open import Data.Unit.Core
open import Level
open import Relation.Binary
import Relation.Binary.Indexed as I
open import Relation.Binary.Consequences
+open import Relation.Binary.HeterogeneousEquality.Core as H using (_≅_)
-- Some of the definitions can be found in the following modules:
@@ -93,22 +95,51 @@ _≗_ {A = A} {B} = Setoid._≈_ (A →-setoid B)
→-to-⟶ = :→-to-Π
------------------------------------------------------------------------
--- The inspect idiom
+-- The old inspect idiom
+
+-- The old inspect idiom has been deprecated, and may be removed in
+-- the future. Use inspect on steroids instead.
+
+module Deprecated-inspect where
+
+ -- The inspect idiom can be used when you want to pattern match on
+ -- the result r of some expression e, and you also need to
+ -- "remember" that r ≡ e.
+
+ -- The inspect idiom has a problem: sometimes you can only pattern
+ -- match on the p part of p with-≡ eq if you also pattern match on
+ -- the eq part, and then you no longer have access to the equality.
+ -- Inspect on steroids solves this problem.
+
+ data Inspect {a} {A : Set a} (x : A) : Set a where
+ _with-≡_ : (y : A) (eq : x ≡ y) → Inspect x
+
+ inspect : ∀ {a} {A : Set a} (x : A) → Inspect x
+ inspect x = x with-≡ refl
+
+ -- Example usage:
+
+ -- f x y with inspect (g x)
+ -- f x y | c z with-≡ eq = ...
+
+------------------------------------------------------------------------
+-- Inspect on steroids
--- The inspect idiom can be used when you want to pattern match on the
--- result r of some expression e, and you also need to "remember" that
--- r ≡ e.
+-- Inspect on steroids can be used when you want to pattern match on
+-- the result r of some expression e, and you also need to "remember"
+-- that r ≡ e.
-data Inspect {a} {A : Set a} (x : A) : Set a where
- _with-≡_ : (y : A) (eq : x ≡ y) → Inspect x
+data Reveal_is_ {a} {A : Set a} (x : Hidden A) (y : A) : Set a where
+ [_] : (eq : reveal x ≡ y) → Reveal x is y
-inspect : ∀ {a} {A : Set a} (x : A) → Inspect x
-inspect x = x with-≡ refl
+inspect : ∀ {a b} {A : Set a} {B : A → Set b}
+ (f : (x : A) → B x) (x : A) → Reveal (hide f x) is (f x)
+inspect f x = [ refl ]
-- Example usage:
--- f x y with inspect (g x)
--- f x y | c z with-≡ eq = ...
+-- f x y with g x | inspect g x
+-- f x y | c z | [ eq ] = ...
------------------------------------------------------------------------
-- Convenient syntax for equational reasoning
@@ -127,8 +158,14 @@ module ≡-Reasoning where
hiding (_≡⟨_⟩_) renaming (_≈⟨_⟩_ to _≡⟨_⟩_)
open Dummy public
+ infixr 2 _≅⟨_⟩_
+
+ _≅⟨_⟩_ : ∀ {a} {A : Set a} (x : A) {y z : A} →
+ x ≅ y → y IsRelatedTo z → x IsRelatedTo z
+ _ ≅⟨ x≅y ⟩ y≡z = _ ≡⟨ H.≅-to-≡ x≅y ⟩ y≡z
+
------------------------------------------------------------------------
--- Definition of functional extensionality
+-- Functional extensionality
-- If _≡_ were extensional, then the following statement could be
-- proved.
@@ -137,3 +174,24 @@ 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
+
+-- If extensionality holds for a given universe level, then it also
+-- holds for lower ones.
+
+extensionality-for-lower-levels :
+ ∀ {a₁ b₁} a₂ b₂ →
+ Extensionality (a₁ ⊔ a₂) (b₁ ⊔ b₂) → Extensionality a₁ b₁
+extensionality-for-lower-levels a₂ b₂ ext f≡g =
+ cong (λ h → lower ∘ h ∘ lift) $
+ ext (cong (lift {ℓ = b₂}) ∘ f≡g ∘ lower {ℓ = a₂})
+
+-- Functional extensionality implies a form of extensionality for
+-- Π-types.
+
+∀-extensionality :
+ ∀ {a b} →
+ Extensionality a (suc b) →
+ {A : Set a} (B₁ B₂ : A → Set b) →
+ (∀ x → B₁ x ≡ B₂ x) → (∀ x → B₁ x) ≡ (∀ x → B₂ x)
+∀-extensionality ext B₁ B₂ B₁≡B₂ with ext B₁≡B₂
+∀-extensionality ext B .B B₁≡B₂ | refl = refl
diff --git a/src/Relation/Binary/PropositionalEquality/Core.agda b/src/Relation/Binary/PropositionalEquality/Core.agda
index 540ce64..17fdf9b 100644
--- a/src/Relation/Binary/PropositionalEquality/Core.agda
+++ b/src/Relation/Binary/PropositionalEquality/Core.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Propositional equality
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- This file contains some core definitions which are reexported by
-- Relation.Binary.PropositionalEquality.
diff --git a/src/Relation/Binary/PropositionalEquality/TrustMe.agda b/src/Relation/Binary/PropositionalEquality/TrustMe.agda
index f363eed..5c85b46 100644
--- a/src/Relation/Binary/PropositionalEquality/TrustMe.agda
+++ b/src/Relation/Binary/PropositionalEquality/TrustMe.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- An equality postulate which evaluates
------------------------------------------------------------------------
@@ -8,12 +10,12 @@ open import Relation.Binary.PropositionalEquality
private
primitive
- primTrustMe : {A : Set}{a b : A} → a ≡ b
+ primTrustMe : ∀ {a} {A : Set a} {x y : A} → x ≡ y
--- trustMe {a = x} {b = y} evaluates to refl if x and y are
+-- trustMe {x = x} {y = y} evaluates to refl if x and y are
-- definitionally equal.
--
-- For an example of the use of trustMe, see Data.String._≟_.
-trustMe : {A : Set}{a b : A} → a ≡ b
+trustMe : ∀ {a} {A : Set a} {x y : A} → x ≡ y
trustMe = primTrustMe
diff --git a/src/Relation/Binary/Props/DecTotalOrder.agda b/src/Relation/Binary/Props/DecTotalOrder.agda
index d9a59fd..2623d62 100644
--- a/src/Relation/Binary/Props/DecTotalOrder.agda
+++ b/src/Relation/Binary/Props/DecTotalOrder.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties satisfied by decidable total orders
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Relation.Binary.Props.DecTotalOrder
diff --git a/src/Relation/Binary/Props/Poset.agda b/src/Relation/Binary/Props/Poset.agda
index 2dfaa58..d40b234 100644
--- a/src/Relation/Binary/Props/Poset.agda
+++ b/src/Relation/Binary/Props/Poset.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties satisfied by posets
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Relation.Binary.Props.Poset
diff --git a/src/Relation/Binary/Props/Preorder.agda b/src/Relation/Binary/Props/Preorder.agda
index 82bfe38..41e8e68 100644
--- a/src/Relation/Binary/Props/Preorder.agda
+++ b/src/Relation/Binary/Props/Preorder.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties satisfied by preorders
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Relation.Binary.Props.Preorder
diff --git a/src/Relation/Binary/Props/StrictPartialOrder.agda b/src/Relation/Binary/Props/StrictPartialOrder.agda
index 9c65caa..e9faf98 100644
--- a/src/Relation/Binary/Props/StrictPartialOrder.agda
+++ b/src/Relation/Binary/Props/StrictPartialOrder.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties satisfied by strict partial orders
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Relation.Binary.Props.StrictPartialOrder
diff --git a/src/Relation/Binary/Props/StrictTotalOrder.agda b/src/Relation/Binary/Props/StrictTotalOrder.agda
index 84a2c04..2e42554 100644
--- a/src/Relation/Binary/Props/StrictTotalOrder.agda
+++ b/src/Relation/Binary/Props/StrictTotalOrder.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties satisfied by strict partial orders
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Relation.Binary.Props.StrictTotalOrder
diff --git a/src/Relation/Binary/Props/TotalOrder.agda b/src/Relation/Binary/Props/TotalOrder.agda
index 6906082..4ae8b96 100644
--- a/src/Relation/Binary/Props/TotalOrder.agda
+++ b/src/Relation/Binary/Props/TotalOrder.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties satisfied by total orders
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Relation.Binary.Props.TotalOrder
diff --git a/src/Relation/Binary/Reflection.agda b/src/Relation/Binary/Reflection.agda
index 5f54c6a..e56e75d 100644
--- a/src/Relation/Binary/Reflection.agda
+++ b/src/Relation/Binary/Reflection.agda
@@ -1,16 +1,16 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Helpers intended to ease the development of "tactics" which use
-- proof by reflection
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Data.Fin
open import Data.Nat
open import Data.Vec as Vec
open import Function
open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence using (module Equivalent)
+open import Function.Equivalence using (module Equivalence)
open import Level
open import Relation.Binary
import Relation.Binary.PropositionalEquality as P
@@ -60,7 +60,7 @@ prove ρ e₁ e₂ hyp = begin
-- Applies the function to all possible "variables".
-close : ∀ {A} n → N-ary n (Expr n) A → A
+close : ∀ {A : Set e} n → N-ary n (Expr n) A → A
close n f = f $ⁿ Vec.map var (allFin n)
-- A variant of prove which should in many cases be easier to use,
@@ -89,7 +89,7 @@ solve₁ : ∀ n (f : N-ary n (Expr n) (Expr n × Expr n)) →
⟦ proj₁ (close n f) ⇓⟧ ρ ≈ ⟦ proj₂ (close n f) ⇓⟧ ρ →
⟦ proj₁ (close n f) ⟧ ρ ≈ ⟦ proj₂ (close n f) ⟧ ρ)
solve₁ n f =
- Equivalent.from (uncurry-∀ⁿ n) ⟨$⟩ λ ρ →
+ Equivalence.from (uncurry-∀ⁿ n) ⟨$⟩ λ ρ →
P.subst id (P.sym (left-inverse (λ _ → _ ≈ _ → _ ≈ _) ρ))
(prove ρ (proj₁ (close n f)) (proj₂ (close n f)))
diff --git a/src/Relation/Binary/Sigma/Pointwise.agda b/src/Relation/Binary/Sigma/Pointwise.agda
index dadec84..9634629 100644
--- a/src/Relation/Binary/Sigma/Pointwise.agda
+++ b/src/Relation/Binary/Sigma/Pointwise.agda
@@ -1,25 +1,28 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- 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
open import Function.Equality as F using (_⟶_; _⟨$⟩_)
- renaming (_∘_ to _⊚_)
open import Function.Equivalence as Eq
- using (Equivalent; _⇔_; module Equivalent)
- renaming (_∘_ to _⟨∘⟩_)
+ using (Equivalence; _⇔_; module Equivalence)
+open import Function.Injection as Inj
+ using (Injection; _↣_; module Injection; Injective)
open import Function.Inverse as Inv
- using (Inverse; _⇿_; module Inverse; Isomorphism)
- renaming (_∘_ to _⟪∘⟫_)
-open import Function.LeftInverse
- using (_LeftInverseOf_; _RightInverseOf_)
-open import Function.Surjection using (_↠_; module Surjection)
+ using (Inverse; _↔_; module Inverse)
+open import Function.LeftInverse as LeftInv
+ using (LeftInverse; _↞_; module LeftInverse;
+ _LeftInverseOf_; _RightInverseOf_)
+open import Function.Related as Related
+ using (_∼[_]_; lam; app-←; app-↢)
+open import Function.Surjection as Surj
+ using (Surjection; _↠_; module Surjection)
import Relation.Binary as B
open import Relation.Binary.Indexed as I using (_at_)
import Relation.Binary.HeterogeneousEquality as H
@@ -88,10 +91,10 @@ setoid s₁ s₂ = record
-- The propositional equality setoid over sigma types can be
-- decomposed using Rel
-Rel⇿≡ : ∀ {a b} {A : Set a} {B : A → Set b} →
+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
+Rel↔≡ {a} {b} {A} {B} = record
{ to = record { _⟨$⟩_ = id; cong = to-cong }
; from = record { _⟨$⟩_ = id; cong = from-cong }
; inverse-of = record
@@ -109,19 +112,16 @@ Rel⇿≡ {a} {b} {A} {B} = record
from-cong {i = (x , y)} P.refl = (P.refl , H.refl)
------------------------------------------------------------------------
--- Equivalences and inverses are also preserved
+-- Some properties related to "relatedness"
-equivalent :
- ∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
- {A₁ : Set a₁} {A₂ : Set a₂}
- {B₁ : I.Setoid A₁ b₁ b₁′} {B₂ : I.Setoid A₂ b₂ b₂′}
- (A₁⇔A₂ : A₁ ⇔ A₂) →
- (∀ {x} → B₁ at x ⟶ B₂ at (Equivalent.to A₁⇔A₂ ⟨$⟩ x)) →
- (∀ {y} → B₂ at y ⟶ B₁ at (Equivalent.from A₁⇔A₂ ⟨$⟩ y)) →
- Equivalent (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
-equivalent {A₁ = A₁} {A₂} {B₁} {B₂} A₁⇔A₂ B-to B-from = record
- { to = record { _⟨$⟩_ = to; cong = to-cong }
- ; from = record { _⟨$⟩_ = from; cong = from-cong }
+⟶ : ∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : I.Setoid A₁ b₁ b₁′} (B₂ : I.Setoid A₂ b₂ b₂′)
+ (f : A₁ → A₂) → (∀ {x} → (B₁ at x) ⟶ (B₂ at f x)) →
+ setoid (P.setoid A₁) B₁ ⟶ setoid (P.setoid A₂) B₂
+⟶ {A₁ = A₁} {A₂} {B₁} B₂ f g = record
+ { _⟨$⟩_ = fg
+ ; cong = fg-cong
}
where
open B.Setoid (setoid (P.setoid A₁) B₁)
@@ -130,28 +130,25 @@ equivalent {A₁ = A₁} {A₂} {B₁} {B₂} A₁⇔A₂ B-to B-from = record
using () renaming (_≈_ to _≈₂_)
open B using (_=[_]⇒_)
- to = Prod.map (_⟨$⟩_ (Equivalent.to A₁⇔A₂)) (_⟨$⟩_ B-to)
+ fg = Prod.map f (_⟨$⟩_ g)
- to-cong : _≈₁_ =[ to ]⇒ _≈₂_
- to-cong (P.refl , ∼) = (P.refl , F.cong B-to ∼)
+ fg-cong : _≈₁_ =[ fg ]⇒ _≈₂_
+ fg-cong (P.refl , ∼) = (P.refl , F.cong g ∼)
- from = Prod.map (_⟨$⟩_ (Equivalent.from A₁⇔A₂)) (_⟨$⟩_ B-from)
-
- from-cong : _≈₂_ =[ from ]⇒ _≈₁_
- from-cong (P.refl , ∼) = (P.refl , F.cong B-from ∼)
-
-equivalent′ :
+equivalence :
∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
{A₁ : Set a₁} {A₂ : Set a₂}
- {B₁ : I.Setoid A₁ b₁ b₁′} (B₂ : I.Setoid A₂ b₂ b₂′)
- (A₁↠A₂ : A₁ ↠ A₂) →
- (∀ {x} → Equivalent (B₁ at x) (B₂ at (Surjection.to A₁↠A₂ ⟨$⟩ x))) →
- Equivalent (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
-equivalent′ {B₁ = B₁} B₂ A₁↠A₂ B₁⇔B₂ =
- equivalent (Surjection.equivalent A₁↠A₂) B-to B-from
- where
- B-to : ∀ {x} → B₁ at x ⟶ B₂ at (Surjection.to A₁↠A₂ ⟨$⟩ x)
- B-to = Equivalent.to B₁⇔B₂
+ {B₁ : I.Setoid A₁ b₁ b₁′} {B₂ : I.Setoid A₂ b₂ b₂′}
+ (A₁⇔A₂ : A₁ ⇔ A₂) →
+ (∀ {x} → _⟶_ (B₁ at x) (B₂ at (Equivalence.to A₁⇔A₂ ⟨$⟩ x))) →
+ (∀ {y} → _⟶_ (B₂ at y) (B₁ at (Equivalence.from A₁⇔A₂ ⟨$⟩ y))) →
+ Equivalence (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+equivalence {B₁ = B₁} {B₂} A₁⇔A₂ B-to B-from = record
+ { to = ⟶ B₂ (_⟨$⟩_ (to A₁⇔A₂)) B-to
+ ; from = ⟶ B₁ (_⟨$⟩_ (from A₁⇔A₂)) B-from
+ } where open Equivalence
+
+private
subst-cong : ∀ {i a p} {I : Set i} {A : I → Set a}
(P : ∀ {i} → A i → A i → Set p) {i i′} {x y : A i}
@@ -159,13 +156,51 @@ equivalent′ {B₁ = B₁} B₂ A₁↠A₂ B₁⇔B₂ =
P x y → P (P.subst A i≡i′ x) (P.subst A i≡i′ y)
subst-cong P P.refl p = p
+equivalence-↞ :
+ ∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ (B₁ : I.Setoid A₁ b₁ b₁′) {B₂ : I.Setoid A₂ b₂ b₂′}
+ (A₁↞A₂ : A₁ ↞ A₂) →
+ (∀ {x} → Equivalence (B₁ at (LeftInverse.from A₁↞A₂ ⟨$⟩ x))
+ (B₂ at x)) →
+ Equivalence (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+equivalence-↞ B₁ {B₂} A₁↞A₂ B₁⇔B₂ =
+ equivalence (LeftInverse.equivalence A₁↞A₂) B-to B-from
+ where
+ B-to : ∀ {x} → _⟶_ (B₁ at x) (B₂ at (LeftInverse.to A₁↞A₂ ⟨$⟩ x))
+ B-to = record
+ { _⟨$⟩_ = λ x → Equivalence.to B₁⇔B₂ ⟨$⟩
+ P.subst (I.Setoid.Carrier B₁)
+ (P.sym $ LeftInverse.left-inverse-of A₁↞A₂ _)
+ x
+ ; cong = F.cong (Equivalence.to B₁⇔B₂) ∘
+ subst-cong (λ {x} → I.Setoid._≈_ B₁ {x} {x})
+ (P.sym (LeftInverse.left-inverse-of A₁↞A₂ _))
+ }
+
+ B-from : ∀ {y} → _⟶_ (B₂ at y) (B₁ at (LeftInverse.from A₁↞A₂ ⟨$⟩ y))
+ B-from = Equivalence.from B₁⇔B₂
+
+equivalence-↠ :
+ ∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : I.Setoid A₁ b₁ b₁′} (B₂ : I.Setoid A₂ b₂ b₂′)
+ (A₁↠A₂ : A₁ ↠ A₂) →
+ (∀ {x} → Equivalence (B₁ at x) (B₂ at (Surjection.to A₁↠A₂ ⟨$⟩ x))) →
+ Equivalence (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+equivalence-↠ {B₁ = B₁} B₂ A₁↠A₂ B₁⇔B₂ =
+ equivalence (Surjection.equivalence A₁↠A₂) B-to B-from
+ where
+ B-to : ∀ {x} → B₁ at x ⟶ B₂ at (Surjection.to A₁↠A₂ ⟨$⟩ x)
+ B-to = Equivalence.to B₁⇔B₂
+
B-from : ∀ {y} → B₂ at y ⟶ B₁ at (Surjection.from A₁↠A₂ ⟨$⟩ y)
B-from = record
- { _⟨$⟩_ = λ x → Equivalent.from B₁⇔B₂ ⟨$⟩
+ { _⟨$⟩_ = λ x → Equivalence.from B₁⇔B₂ ⟨$⟩
P.subst (I.Setoid.Carrier B₂)
(P.sym $ Surjection.right-inverse-of A₁↠A₂ _)
x
- ; cong = F.cong (Equivalent.from B₁⇔B₂) ∘
+ ; cong = F.cong (Equivalence.from B₁⇔B₂) ∘
subst-cong (λ {x} → I.Setoid._≈_ B₂ {x} {x})
(P.sym (Surjection.right-inverse-of A₁↠A₂ _))
}
@@ -174,100 +209,241 @@ equivalent′ {B₁ = B₁} B₂ A₁↠A₂ B₁⇔B₂ =
{A₁ : Set a₁} {A₂ : Set a₂}
{B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
(A₁⇔A₂ : A₁ ⇔ A₂) →
- (∀ {x} → B₁ x → B₂ (Equivalent.to A₁⇔A₂ ⟨$⟩ x)) →
- (∀ {y} → B₂ y → B₁ (Equivalent.from A₁⇔A₂ ⟨$⟩ y)) →
+ (∀ {x} → B₁ x → B₂ (Equivalence.to A₁⇔A₂ ⟨$⟩ x)) →
+ (∀ {y} → B₂ y → B₁ (Equivalence.from A₁⇔A₂ ⟨$⟩ y)) →
Σ A₁ B₁ ⇔ Σ A₂ B₂
⇔ {B₁ = B₁} {B₂} A₁⇔A₂ B-to B-from =
- Inverse.equivalent (Rel⇿≡ {B = B₂}) ⟨∘⟩
- equivalent A₁⇔A₂
- (Inverse.to (H.≡⇿≅ B₂) ⊚ P.→-to-⟶ B-to ⊚ Inverse.from (H.≡⇿≅ B₁))
- (Inverse.to (H.≡⇿≅ B₁) ⊚ P.→-to-⟶ B-from ⊚ Inverse.from (H.≡⇿≅ B₂))
+ Inverse.equivalence (Rel↔≡ {B = B₂}) ⟨∘⟩
+ equivalence A₁⇔A₂
+ (Inverse.to (H.≡↔≅ B₂) ⊚ P.→-to-⟶ B-to ⊚ Inverse.from (H.≡↔≅ B₁))
+ (Inverse.to (H.≡↔≅ B₁) ⊚ P.→-to-⟶ B-from ⊚ Inverse.from (H.≡↔≅ B₂))
⟨∘⟩
- Eq.sym (Inverse.equivalent (Rel⇿≡ {B = B₁}))
-
-⇔′ : ∀ {a₁ a₂ b₁ b₂}
- {A₁ : Set a₁} {A₂ : Set a₂}
- {B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
- (A₁↠A₂ : A₁ ↠ A₂) →
- (∀ {x} → _⇔_ (B₁ x) (B₂ (Surjection.to A₁↠A₂ ⟨$⟩ x))) →
- _⇔_ (Σ A₁ B₁) (Σ A₂ B₂)
-⇔′ {B₁ = B₁} {B₂} A₁↠A₂ B₁⇔B₂ =
- Inverse.equivalent (Rel⇿≡ {B = B₂}) ⟨∘⟩
- equivalent′ (H.indexedSetoid B₂) A₁↠A₂
- (λ {x} → Inverse.equivalent (H.≡⇿≅ B₂) ⟨∘⟩
+ Eq.sym (Inverse.equivalence (Rel↔≡ {B = B₁}))
+ where
+ open Eq using () renaming (_∘_ to _⟨∘⟩_)
+ open F using () renaming (_∘_ to _⊚_)
+
+⇔-↠ : ∀ {a₁ a₂ b₁ b₂}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
+ (A₁↠A₂ : A₁ ↠ A₂) →
+ (∀ {x} → _⇔_ (B₁ x) (B₂ (Surjection.to A₁↠A₂ ⟨$⟩ x))) →
+ _⇔_ (Σ A₁ B₁) (Σ A₂ B₂)
+⇔-↠ {B₁ = B₁} {B₂} A₁↠A₂ B₁⇔B₂ =
+ Inverse.equivalence (Rel↔≡ {B = B₂}) ⟨∘⟩
+ equivalence-↠ (H.indexedSetoid B₂) A₁↠A₂
+ (λ {x} → Inverse.equivalence (H.≡↔≅ B₂) ⟨∘⟩
B₁⇔B₂ {x} ⟨∘⟩
- Inverse.equivalent (Inv.sym (H.≡⇿≅ B₁))) ⟨∘⟩
- Eq.sym (Inverse.equivalent (Rel⇿≡ {B = B₁}))
+ Inverse.equivalence (Inv.sym (H.≡↔≅ B₁))) ⟨∘⟩
+ Eq.sym (Inverse.equivalence (Rel↔≡ {B = B₁}))
+ where open Eq using () renaming (_∘_ to _⟨∘⟩_)
+
+injection :
+ ∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : I.Setoid A₁ b₁ b₁′} (B₂ : I.Setoid A₂ b₂ b₂′) →
+ (A₁↣A₂ : A₁ ↣ A₂) →
+ (∀ {x} → Injection (B₁ at x) (B₂ at (Injection.to A₁↣A₂ ⟨$⟩ x))) →
+ Injection (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+injection {B₁ = B₁} B₂ A₁↣A₂ B₁↣B₂ = record
+ { to = to
+ ; injective = inj
+ }
+ where
+ to = ⟶ B₂ (_⟨$⟩_ (Injection.to A₁↣A₂)) (Injection.to B₁↣B₂)
+
+ inj : Injective to
+ inj (x , y) =
+ Injection.injective A₁↣A₂ x ,
+ lemma (Injection.injective A₁↣A₂ x) y
+ where
+ lemma :
+ ∀ {x x′}
+ {y : I.Setoid.Carrier B₁ x} {y′ : I.Setoid.Carrier B₁ x′} →
+ P._≡_ x x′ →
+ (eq : I.Setoid._≈_ B₂ (Injection.to B₁↣B₂ ⟨$⟩ y)
+ (Injection.to B₁↣B₂ ⟨$⟩ y′)) →
+ I.Setoid._≈_ B₁ y y′
+ lemma P.refl = Injection.injective B₁↣B₂
+
+↣ : ∀ {a₁ a₂ b₁ b₂}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
+ (A₁↣A₂ : A₁ ↣ A₂) →
+ (∀ {x} → B₁ x ↣ B₂ (Injection.to A₁↣A₂ ⟨$⟩ x)) →
+ Σ A₁ B₁ ↣ Σ A₂ B₂
+↣ {B₁ = B₁} {B₂} A₁↣A₂ B₁↣B₂ =
+ Inverse.injection (Rel↔≡ {B = B₂}) ⟨∘⟩
+ injection (H.indexedSetoid B₂) A₁↣A₂
+ (λ {x} → Inverse.injection (H.≡↔≅ B₂) ⟨∘⟩
+ B₁↣B₂ {x} ⟨∘⟩
+ Inverse.injection (Inv.sym (H.≡↔≅ B₁))) ⟨∘⟩
+ Inverse.injection (Inv.sym (Rel↔≡ {B = B₁}))
+ where open Inj using () renaming (_∘_ to _⟨∘⟩_)
+
+left-inverse :
+ ∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ (B₁ : I.Setoid A₁ b₁ b₁′) {B₂ : I.Setoid A₂ b₂ b₂′} →
+ (A₁↞A₂ : A₁ ↞ A₂) →
+ (∀ {x} → LeftInverse (B₁ at (LeftInverse.from A₁↞A₂ ⟨$⟩ x))
+ (B₂ at x)) →
+ LeftInverse (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+left-inverse B₁ {B₂} A₁↞A₂ B₁↞B₂ = record
+ { to = Equivalence.to eq
+ ; from = Equivalence.from eq
+ ; left-inverse-of = left
+ }
+ where
+ eq = equivalence-↞ B₁ A₁↞A₂ (LeftInverse.equivalence B₁↞B₂)
+
+ left : Equivalence.from eq LeftInverseOf Equivalence.to eq
+ left (x , y) =
+ LeftInverse.left-inverse-of A₁↞A₂ x ,
+ I.Setoid.trans B₁
+ (LeftInverse.left-inverse-of B₁↞B₂ _)
+ (lemma (P.sym (LeftInverse.left-inverse-of A₁↞A₂ x)))
+ where
+ lemma :
+ ∀ {x x′ y} (eq : P._≡_ x x′) →
+ I.Setoid._≈_ B₁ (P.subst (I.Setoid.Carrier B₁) eq y) y
+ lemma P.refl = I.Setoid.refl B₁
+
+↞ : ∀ {a₁ a₂ b₁ b₂}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
+ (A₁↞A₂ : A₁ ↞ A₂) →
+ (∀ {x} → B₁ (LeftInverse.from A₁↞A₂ ⟨$⟩ x) ↞ B₂ x) →
+ Σ A₁ B₁ ↞ Σ A₂ B₂
+↞ {B₁ = B₁} {B₂} A₁↞A₂ B₁↞B₂ =
+ Inverse.left-inverse (Rel↔≡ {B = B₂}) ⟨∘⟩
+ left-inverse (H.indexedSetoid B₁) A₁↞A₂
+ (λ {x} → Inverse.left-inverse (H.≡↔≅ B₂) ⟨∘⟩
+ B₁↞B₂ {x} ⟨∘⟩
+ Inverse.left-inverse (Inv.sym (H.≡↔≅ B₁))) ⟨∘⟩
+ Inverse.left-inverse (Inv.sym (Rel↔≡ {B = B₁}))
+ where open LeftInv using () renaming (_∘_ to _⟨∘⟩_)
+
+surjection :
+ ∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : I.Setoid A₁ b₁ b₁′} (B₂ : I.Setoid A₂ b₂ b₂′) →
+ (A₁↠A₂ : A₁ ↠ A₂) →
+ (∀ {x} → Surjection (B₁ at x) (B₂ at (Surjection.to A₁↠A₂ ⟨$⟩ x))) →
+ Surjection (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+surjection {B₁} B₂ A₁↠A₂ B₁↠B₂ = record
+ { to = Equivalence.to eq
+ ; surjective = record
+ { from = Equivalence.from eq
+ ; right-inverse-of = right
+ }
+ }
+ where
+ eq = equivalence-↠ B₂ A₁↠A₂ (Surjection.equivalence B₁↠B₂)
+
+ right : Equivalence.from eq RightInverseOf Equivalence.to eq
+ right (x , y) =
+ Surjection.right-inverse-of A₁↠A₂ x ,
+ I.Setoid.trans B₂
+ (Surjection.right-inverse-of B₁↠B₂ _)
+ (lemma (P.sym $ Surjection.right-inverse-of A₁↠A₂ x))
+ where
+ lemma : ∀ {x x′ y} (eq : P._≡_ x x′) →
+ I.Setoid._≈_ B₂ (P.subst (I.Setoid.Carrier B₂) eq y) y
+ lemma P.refl = I.Setoid.refl B₂
+
+↠ : ∀ {a₁ a₂ b₁ b₂}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
+ (A₁↠A₂ : A₁ ↠ A₂) →
+ (∀ {x} → B₁ x ↠ B₂ (Surjection.to A₁↠A₂ ⟨$⟩ x)) →
+ Σ A₁ B₁ ↠ Σ A₂ B₂
+↠ {B₁ = B₁} {B₂} A₁↠A₂ B₁↠B₂ =
+ Inverse.surjection (Rel↔≡ {B = B₂}) ⟨∘⟩
+ surjection (H.indexedSetoid B₂) A₁↠A₂
+ (λ {x} → Inverse.surjection (H.≡↔≅ B₂) ⟨∘⟩
+ B₁↠B₂ {x} ⟨∘⟩
+ Inverse.surjection (Inv.sym (H.≡↔≅ B₁))) ⟨∘⟩
+ Inverse.surjection (Inv.sym (Rel↔≡ {B = B₁}))
+ where open Surj using () renaming (_∘_ to _⟨∘⟩_)
inverse :
∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
{A₁ : Set a₁} {A₂ : Set a₂}
{B₁ : I.Setoid A₁ b₁ b₁′} (B₂ : I.Setoid A₂ b₂ b₂′) →
- (A₁⇿A₂ : A₁ ⇿ A₂) →
- (∀ {x} → Inverse (B₁ at x) (B₂ at (Inverse.to A₁⇿A₂ ⟨$⟩ x))) →
+ (A₁↔A₂ : A₁ ↔ A₂) →
+ (∀ {x} → Inverse (B₁ at x) (B₂ at (Inverse.to A₁↔A₂ ⟨$⟩ x))) →
Inverse (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
-inverse {A₁ = A₁} {A₂} {B₁} B₂ A₁⇿A₂ B₁⇿B₂ = record
- { to = Equivalent.to eq
- ; from = Equivalent.from eq
+inverse {B₁ = B₁} B₂ A₁↔A₂ B₁↔B₂ = record
+ { to = Surjection.to surj
+ ; from = Surjection.from surj
; inverse-of = record
{ left-inverse-of = left
- ; right-inverse-of = right
+ ; right-inverse-of = Surjection.right-inverse-of surj
}
}
where
- eq = equivalent′ B₂ (Inverse.surjection A₁⇿A₂)
- (Inverse.equivalent B₁⇿B₂)
+ surj = surjection B₂ (Inverse.surjection A₁↔A₂)
+ (Inverse.surjection B₁↔B₂)
- left : Equivalent.from eq LeftInverseOf Equivalent.to eq
+ left : Surjection.from surj LeftInverseOf Surjection.to surj
left (x , y) =
- Inverse.left-inverse-of A₁⇿A₂ x ,
+ Inverse.left-inverse-of A₁↔A₂ x ,
I.Setoid.trans B₁
- (lemma (P.sym (Inverse.left-inverse-of A₁⇿A₂ x))
- (P.sym (Inverse.right-inverse-of A₁⇿A₂
- (Inverse.to A₁⇿A₂ ⟨$⟩ x))))
- (Inverse.left-inverse-of B₁⇿B₂ y)
+ (lemma (P.sym (Inverse.left-inverse-of A₁↔A₂ x))
+ (P.sym (Inverse.right-inverse-of A₁↔A₂
+ (Inverse.to A₁↔A₂ ⟨$⟩ x))))
+ (Inverse.left-inverse-of B₁↔B₂ y)
where
lemma :
∀ {x x′ y} → P._≡_ x x′ →
- (eq : P._≡_ (Inverse.to A₁⇿A₂ ⟨$⟩ x) (Inverse.to A₁⇿A₂ ⟨$⟩ x′)) →
+ (eq : P._≡_ (Inverse.to A₁↔A₂ ⟨$⟩ x) (Inverse.to A₁↔A₂ ⟨$⟩ x′)) →
I.Setoid._≈_ B₁
- (Inverse.from B₁⇿B₂ ⟨$⟩ P.subst (I.Setoid.Carrier B₂) eq y)
- (Inverse.from B₁⇿B₂ ⟨$⟩ y)
+ (Inverse.from B₁↔B₂ ⟨$⟩ P.subst (I.Setoid.Carrier B₂) eq y)
+ (Inverse.from B₁↔B₂ ⟨$⟩ y)
lemma P.refl P.refl = I.Setoid.refl B₁
- right : Equivalent.from eq RightInverseOf Equivalent.to eq
- right (x , y) =
- Inverse.right-inverse-of A₁⇿A₂ x ,
- I.Setoid.trans B₂
- (Inverse.right-inverse-of B₁⇿B₂
- (P.subst (I.Setoid.Carrier B₂) p y))
- (lemma p)
- where
- p = P.sym $ Inverse.right-inverse-of A₁⇿A₂ x
-
- lemma : ∀ {x x′ y} (eq : P._≡_ x x′) →
- I.Setoid._≈_ B₂
- (P.subst (I.Setoid.Carrier B₂) eq y)
- y
- lemma P.refl = I.Setoid.refl B₂
-
-⇿ : ∀ {a₁ a₂ b₁ b₂}
+↔ : ∀ {a₁ a₂ b₁ b₂}
{A₁ : Set a₁} {A₂ : Set a₂}
{B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
- (A₁⇿A₂ : A₁ ⇿ A₂) →
- (∀ {x} → B₁ x ⇿ B₂ (Inverse.to A₁⇿A₂ ⟨$⟩ x)) →
- Σ A₁ B₁ ⇿ Σ A₂ B₂
-⇿ {B₁ = B₁} {B₂} A₁⇿A₂ B₁⇿B₂ =
- Rel⇿≡ {B = B₂} ⟪∘⟫
- inverse (H.indexedSetoid B₂) A₁⇿A₂
- (λ {x} → H.≡⇿≅ B₂ ⟪∘⟫ B₁⇿B₂ {x} ⟪∘⟫ Inv.sym (H.≡⇿≅ B₁)) ⟪∘⟫
- Inv.sym (Rel⇿≡ {B = B₁})
+ (A₁↔A₂ : A₁ ↔ A₂) →
+ (∀ {x} → B₁ x ↔ B₂ (Inverse.to A₁↔A₂ ⟨$⟩ x)) →
+ Σ A₁ B₁ ↔ Σ A₂ B₂
+↔ {B₁ = B₁} {B₂} A₁↔A₂ B₁↔B₂ =
+ Rel↔≡ {B = B₂} ⟨∘⟩
+ inverse (H.indexedSetoid B₂) A₁↔A₂
+ (λ {x} → H.≡↔≅ B₂ ⟨∘⟩ B₁↔B₂ {x} ⟨∘⟩ Inv.sym (H.≡↔≅ B₁)) ⟨∘⟩
+ Inv.sym (Rel↔≡ {B = B₁})
+ where open Inv using () renaming (_∘_ to _⟨∘⟩_)
+
+private
+
+ swap-coercions :
+ ∀ {k a₁ a₂ b₁ b₂}
+ {A₁ : Set a₁} {A₂ : Set a₂} {B₁ : A₁ → Set b₁} (B₂ : A₂ → Set b₂)
+ (A₁↔A₂ : _↔_ A₁ A₂) →
+ (∀ {x} → B₁ x ∼[ k ] B₂ (Inverse.to A₁↔A₂ ⟨$⟩ x)) →
+ ∀ {x} → B₁ (Inverse.from A₁↔A₂ ⟨$⟩ x) ∼[ k ] B₂ x
+ swap-coercions {k} {B₁ = B₁} B₂ A₁↔A₂ eq {x} =
+ B₁ (Inverse.from A₁↔A₂ ⟨$⟩ x) ∼⟨ eq ⟩
+ B₂ (Inverse.to A₁↔A₂ ⟨$⟩ (Inverse.from A₁↔A₂ ⟨$⟩ x)) ↔⟨ B.Setoid.reflexive (Related.setoid Related.bijection _)
+ (P.cong B₂ $ Inverse.right-inverse-of A₁↔A₂ x) ⟩
+ B₂ x ∎
+ where open Related.EquationalReasoning
cong : ∀ {k a₁ a₂ b₁ b₂}
{A₁ : Set a₁} {A₂ : Set a₂}
{B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
- (A₁⇿A₂ : _⇿_ A₁ A₂) →
- (∀ {x} → Isomorphism k (B₁ x) (B₂ (Inverse.to A₁⇿A₂ ⟨$⟩ x))) →
- Isomorphism k (Σ A₁ B₁) (Σ A₂ B₂)
-cong {k = Inv.equivalent} = ⇔′ ∘ Inverse.surjection
-cong {k = Inv.inverse} = ⇿
+ (A₁↔A₂ : _↔_ A₁ A₂) →
+ (∀ {x} → B₁ x ∼[ k ] B₂ (Inverse.to A₁↔A₂ ⟨$⟩ x)) →
+ Σ A₁ B₁ ∼[ k ] Σ A₂ B₂
+cong {Related.implication} = λ A₁↔A₂ → Prod.map (_⟨$⟩_ (Inverse.to A₁↔A₂))
+cong {Related.reverse-implication} {B₂ = B₂} = λ A₁↔A₂ B₁←B₂ → lam (Prod.map (_⟨$⟩_ (Inverse.from A₁↔A₂))
+ (app-← (swap-coercions B₂ A₁↔A₂ B₁←B₂)))
+cong {Related.equivalence} = ⇔-↠ ∘ Inverse.surjection
+cong {Related.injection} = ↣ ∘ Inverse.injection
+cong {Related.reverse-injection} {B₂ = B₂} = λ A₁↔A₂ B₁↢B₂ → lam (↣ (Inverse.injection (Inv.sym A₁↔A₂))
+ (app-↢ (swap-coercions B₂ A₁↔A₂ B₁↢B₂)))
+cong {Related.left-inverse} = λ A₁↔A₂ → ↞ (Inverse.left-inverse A₁↔A₂) ∘ swap-coercions _ A₁↔A₂
+cong {Related.surjection} = ↠ ∘ Inverse.surjection
+cong {Related.bijection} = ↔
diff --git a/src/Relation/Binary/Simple.agda b/src/Relation/Binary/Simple.agda
index a1e6da4..dc4d8f4 100644
--- a/src/Relation/Binary/Simple.agda
+++ b/src/Relation/Binary/Simple.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Some simple binary relations
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Binary.Simple where
open import Relation.Binary
@@ -18,10 +18,10 @@ Const I = λ _ _ → I
-- The universally true relation.
-Always : ∀ {a} {A : Set a} → Rel A zero
-Always = Const ⊤
+Always : ∀ {a ℓ} {A : Set a} → Rel A ℓ
+Always = Const (Lift ⊤)
-Always-setoid : ∀ {a} (A : Set a) → Setoid _ _
+Always-setoid : ∀ {a ℓ} (A : Set a) → Setoid a ℓ
Always-setoid A = record
{ Carrier = A
; _≈_ = Always
@@ -30,5 +30,5 @@ Always-setoid A = record
-- The universally false relation.
-Never : ∀ {a} {A : Set a} → Rel A zero
-Never = Const ⊥
+Never : ∀ {a ℓ} {A : Set a} → Rel A ℓ
+Never = Const (Lift ⊥)
diff --git a/src/Relation/Binary/StrictPartialOrderReasoning.agda b/src/Relation/Binary/StrictPartialOrderReasoning.agda
index 1d76b83..facb252 100644
--- a/src/Relation/Binary/StrictPartialOrderReasoning.agda
+++ b/src/Relation/Binary/StrictPartialOrderReasoning.agda
@@ -1,10 +1,10 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Convenient syntax for "equational reasoning" using a strict partial
-- order
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
open import Relation.Binary
module Relation.Binary.StrictPartialOrderReasoning
diff --git a/src/Relation/Binary/StrictToNonStrict.agda b/src/Relation/Binary/StrictToNonStrict.agda
index be7bb1f..85434fd 100644
--- a/src/Relation/Binary/StrictToNonStrict.agda
+++ b/src/Relation/Binary/StrictToNonStrict.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Conversion of < to ≤, along with a number of properties
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- Possible TODO: Prove that a conversion ≤ → < → ≤ returns a
-- relation equivalent to the original one (and similarly for
-- < → ≤ → <).
diff --git a/src/Relation/Binary/Sum.agda b/src/Relation/Binary/Sum.agda
index 6f40f9c..1a1f2e7 100644
--- a/src/Relation/Binary/Sum.agda
+++ b/src/Relation/Binary/Sum.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Sums of binary relations
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Binary.Sum where
open import Data.Sum as Sum
@@ -11,13 +11,18 @@ 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.Equality as F using (_⟶_; _⟨$⟩_)
open import Function.Equivalence as Eq
- using (Equivalent; _⇔_; module Equivalent)
- renaming (_∘_ to _⟨∘⟩_)
+ using (Equivalence; _⇔_; module Equivalence)
+open import Function.Injection as Inj
+ using (Injection; _↣_; module Injection)
open import Function.Inverse as Inv
- using (Inverse; _⇿_; module Inverse)
- renaming (_∘_ to _⟪∘⟫_)
+ using (Inverse; _↔_; module Inverse)
+open import Function.LeftInverse as LeftInv
+ using (LeftInverse; _↞_; module LeftInverse)
+open import Function.Related
+open import Function.Surjection as Surj
+ using (Surjection; _↠_; module Surjection)
open import Level
open import Relation.Nullary
open import Relation.Binary
@@ -306,6 +311,20 @@ private
}
where open IsDecTotalOrder
+ _⊎-<-isStrictTotalOrder_ :
+ ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {<₁ : Rel A₁ ℓ₁′}
+ {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {<₂ : Rel A₂ ℓ₂′} →
+ IsStrictTotalOrder ≈₁ <₁ → IsStrictTotalOrder ≈₂ <₂ →
+ IsStrictTotalOrder (≈₁ ⊎-Rel ≈₂) (<₁ ⊎-< <₂)
+ sto₁ ⊎-<-isStrictTotalOrder sto₂ = record
+ { isEquivalence = isEquivalence sto₁ ⊎-isEquivalence
+ isEquivalence sto₂
+ ; trans = trans sto₁ ⊎-transitive trans sto₂
+ ; compare = compare sto₁ ⊎-<-trichotomous compare sto₂
+ ; <-resp-≈ = <-resp-≈ sto₁ ⊎-≈-respects₂ <-resp-≈ sto₂
+ }
+ where open IsStrictTotalOrder
+
open Dummy public
------------------------------------------------------------------------
@@ -373,12 +392,23 @@ to₁ ⊎-<-decTotalOrder to₂ = record
isDecTotalOrder to₂
} where open DecTotalOrder
+_⊎-<-strictTotalOrder_ :
+ ∀ {p₁ p₂ p₃ p₄ p₅ p₆} →
+ StrictTotalOrder p₁ p₂ p₃ → StrictTotalOrder p₄ p₅ p₆ →
+ StrictTotalOrder _ _ _
+sto₁ ⊎-<-strictTotalOrder sto₂ = record
+ { _<_ = _<_ sto₁ ⊎-< _<_ sto₂
+ ; isStrictTotalOrder = isStrictTotalOrder sto₁
+ ⊎-<-isStrictTotalOrder
+ isStrictTotalOrder sto₂
+ } where open StrictTotalOrder
+
------------------------------------------------------------------------
--- Some properties related to equivalences and inverses
+-- Some properties related to "relatedness"
-⊎-Rel⇿≡ : ∀ {a b} (A : Set a) (B : Set b) →
+⊎-Rel↔≡ : ∀ {a b} (A : Set a) (B : Set b) →
Inverse (P.setoid A ⊎-setoid P.setoid B) (P.setoid (A ⊎ B))
-⊎-Rel⇿≡ _ _ = record
+⊎-Rel↔≡ _ _ = record
{ to = record { _⟨$⟩_ = id; cong = to-cong }
; from = record { _⟨$⟩_ = id; cong = from-cong }
; inverse-of = record
@@ -395,65 +425,160 @@ to₁ ⊎-<-decTotalOrder to₂ = record
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 }
+ A ⟶ B → C ⟶ D → (A ⊎-setoid C) ⟶ (B ⊎-setoid D)
+_⊎-⟶_ {A = A} {B} {C} {D} f g = record
+ { _⟨$⟩_ = fg
+ ; cong = fg-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
+ fg = Sum.map (_⟨$⟩_ f) (_⟨$⟩_ g)
- from = Sum.map (_⟨$⟩_ (Equivalent.from A⇔B))
- (_⟨$⟩_ (Equivalent.from C⇔D))
+ fg-cong : _≈AC_ =[ fg ]⇒ _≈BD_
+ fg-cong (₁∼₂ ())
+ fg-cong (₁∼₁ x∼₁y) = ₁∼₁ $ F.cong f x∼₁y
+ fg-cong (₂∼₂ x∼₂y) = ₂∼₂ $ F.cong g x∼₂y
- 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
+_⊎-equivalence_ :
+ ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈}
+ {A : Setoid s₁ s₂} {B : Setoid s₃ s₄}
+ {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} →
+ Equivalence A B → Equivalence C D →
+ Equivalence (A ⊎-setoid C) (B ⊎-setoid D)
+A⇔B ⊎-equivalence C⇔D = record
+ { to = to A⇔B ⊎-⟶ to C⇔D
+ ; from = from A⇔B ⊎-⟶ from C⇔D
+ } where open Equivalence
_⊎-⇔_ : ∀ {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.equivalence (⊎-Rel↔≡ B D) ⟨∘⟩
+ A⇔B ⊎-equivalence C⇔D ⟨∘⟩
+ Eq.sym (Inverse.equivalence (⊎-Rel↔≡ A C))
+ where open Eq using () renaming (_∘_ to _⟨∘⟩_)
+
+_⊎-injection_ :
+ ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈}
+ {A : Setoid s₁ s₂} {B : Setoid s₃ s₄}
+ {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} →
+ Injection A B → Injection C D →
+ Injection (A ⊎-setoid C) (B ⊎-setoid D)
+_⊎-injection_ {A = A} {B} {C} {D} A↣B C↣D = record
+ { to = to A↣B ⊎-⟶ to C↣D
+ ; injective = inj _ _
+ }
+ where
+ open Injection
+ open Setoid (A ⊎-setoid C) using () renaming (_≈_ to _≈AC_)
+ open Setoid (B ⊎-setoid D) using () renaming (_≈_ to _≈BD_)
+
+ inj : ∀ x y →
+ (to A↣B ⊎-⟶ to C↣D) ⟨$⟩ x ≈BD (to A↣B ⊎-⟶ to C↣D) ⟨$⟩ y →
+ x ≈AC y
+ inj (inj₁ x) (inj₁ y) (₁∼₁ x∼₁y) = ₁∼₁ (injective A↣B x∼₁y)
+ inj (inj₂ x) (inj₂ y) (₂∼₂ x∼₂y) = ₂∼₂ (injective C↣D x∼₂y)
+ inj (inj₁ x) (inj₂ y) (₁∼₂ ())
+ inj (inj₂ x) (inj₁ 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.injection (⊎-Rel↔≡ B D) ⟨∘⟩
+ A↣B ⊎-injection C↣D ⟨∘⟩
+ Inverse.injection (Inv.sym (⊎-Rel↔≡ A C))
+ where open Inj using () renaming (_∘_ to _⟨∘⟩_)
+
+_⊎-left-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₈} →
+ LeftInverse A B → LeftInverse C D →
+ LeftInverse (A ⊎-setoid C) (B ⊎-setoid D)
+A↞B ⊎-left-inverse C↞D = record
+ { to = Equivalence.to eq
+ ; from = Equivalence.from eq
+ ; left-inverse-of = [ ₁∼₁ ∘ left-inverse-of A↞B
+ , ₂∼₂ ∘ left-inverse-of C↞D
+ ]
+ }
+ where
+ open LeftInverse
+ eq = LeftInverse.equivalence A↞B ⊎-equivalence
+ LeftInverse.equivalence 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.left-inverse (⊎-Rel↔≡ B D) ⟨∘⟩
+ A↞B ⊎-left-inverse C↞D ⟨∘⟩
+ Inverse.left-inverse (Inv.sym (⊎-Rel↔≡ A C))
+ where open LeftInv using () renaming (_∘_ to _⟨∘⟩_)
+
+_⊎-surjection_ :
+ ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈}
+ {A : Setoid s₁ s₂} {B : Setoid s₃ s₄}
+ {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} →
+ Surjection A B → Surjection C D →
+ Surjection (A ⊎-setoid C) (B ⊎-setoid D)
+A↠B ⊎-surjection C↠D = record
+ { to = LeftInverse.from inv
+ ; surjective = record
+ { from = LeftInverse.to inv
+ ; right-inverse-of = LeftInverse.left-inverse-of inv
+ }
+ }
+ where
+ open Surjection
+ inv = right-inverse A↠B ⊎-left-inverse right-inverse 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.surjection (⊎-Rel↔≡ B D) ⟨∘⟩
+ A↠B ⊎-surjection C↠D ⟨∘⟩
+ Inverse.surjection (Inv.sym (⊎-Rel↔≡ A C))
+ where open Surj using () renaming (_∘_ to _⟨∘⟩_)
_⊎-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
+A↔B ⊎-inverse C↔D = record
+ { to = Surjection.to surj
+ ; from = Surjection.from surj
; 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
- ]
+ { left-inverse-of = LeftInverse.left-inverse-of inv
+ ; right-inverse-of = Surjection.right-inverse-of surj
}
}
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)
+ surj = Inverse.surjection A↔B ⊎-surjection
+ Inverse.surjection C↔D
+ inv = Inverse.left-inverse A↔B ⊎-left-inverse
+ Inverse.left-inverse 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)
+ where open Inv using () renaming (_∘_ to _⟨∘⟩_)
+
+_⊎-cong_ : ∀ {k a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
+ A ∼[ k ] B → C ∼[ k ] D → (A ⊎ C) ∼[ k ] (B ⊎ D)
+_⊎-cong_ {implication} = Sum.map
+_⊎-cong_ {reverse-implication} = λ f g → lam (Sum.map (app-← f) (app-← g))
+_⊎-cong_ {equivalence} = _⊎-⇔_
+_⊎-cong_ {injection} = _⊎-↣_
+_⊎-cong_ {reverse-injection} = λ f g → lam (app-↢ f ⊎-↣ app-↢ g)
+_⊎-cong_ {left-inverse} = _⊎-↞_
+_⊎-cong_ {surjection} = _⊎-↠_
+_⊎-cong_ {bijection} = _⊎-↔_
diff --git a/src/Relation/Binary/Vec/Pointwise.agda b/src/Relation/Binary/Vec/Pointwise.agda
index 9613aa4..fdf9edf 100644
--- a/src/Relation/Binary/Vec/Pointwise.agda
+++ b/src/Relation/Binary/Vec/Pointwise.agda
@@ -1,21 +1,22 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Pointwise lifting of relations to vectors
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Binary.Vec.Pointwise where
open import Category.Applicative.Indexed
open import Data.Fin
+open import Data.Nat
open import Data.Plus as Plus hiding (equivalent; map)
open import Data.Vec as Vec hiding ([_]; map)
import Data.Vec.Properties as VecProp
open import Function
open import Function.Equality using (_⟨$⟩_)
open import Function.Equivalence as Equiv
- using (_⇔_; module Equivalent)
-open import Level
+ using (_⇔_; module Equivalence)
+import Level
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as P using (_≡_)
open import Relation.Nullary
@@ -44,7 +45,7 @@ private
equivalent : ∀ {_∼_ : Rel A a} {n} {xs ys : Vec A n} →
Pointwise _∼_ xs ys ⇔ Pointwise′ _∼_ xs ys
- equivalent {_∼_} = Equiv.equivalent (to _ _) from
+ equivalent {_∼_} = Equiv.equivalence (to _ _) from
where
to : ∀ {n} (xs ys : Vec A n) →
Pointwise _∼_ xs ys → Pointwise′ _∼_ xs ys
@@ -104,8 +105,8 @@ private
Pointwise-≡ : ∀ {n} {xs ys : Vec A n} →
Pointwise _≡_ xs ys ⇔ xs ≡ ys
Pointwise-≡ =
- Equiv.equivalent
- (to ∘ _⟨$⟩_ (Equivalent.to equivalent))
+ Equiv.equivalence
+ (to ∘ _⟨$⟩_ (Equivalence.to equivalent))
(λ xs≡ys → P.subst (Pointwise _≡_ _) xs≡ys (refl P.refl))
where
to : ∀ {n} {xs ys : Vec A n} → Pointwise′ _≡_ xs ys → xs ≡ ys
@@ -126,9 +127,9 @@ private
Reflexive _∼_ →
Pointwise (Plus _∼_) xs ys → Plus (Pointwise _∼_) xs ys
∙⁺⇒⁺∙ {_∼_} x∼x =
- Plus.map (_⟨$⟩_ (Equivalent.from equivalent)) ∘
+ Plus.map (_⟨$⟩_ (Equivalence.from equivalent)) ∘
helper ∘
- _⟨$⟩_ (Equivalent.to equivalent)
+ _⟨$⟩_ (Equivalence.to equivalent)
where
helper : ∀ {n} {xs ys : Vec A n} →
Pointwise′ (Plus _∼_) xs ys → Plus (Pointwise′ _∼_) xs ys
@@ -139,7 +140,7 @@ private
y ∷ ys ∎
where
xs∼xs : Pointwise′ _∼_ xs xs
- xs∼xs = Equivalent.to equivalent ⟨$⟩ refl x∼x
+ xs∼xs = Equivalence.to equivalent ⟨$⟩ refl x∼x
open Dummy public
@@ -153,7 +154,7 @@ private
data D : Set where
i j x y z : D
- data _R_ : Rel D zero where
+ data _R_ : Rel D Level.zero where
iRj : i R j
xRy : x R y
yRz : y R z
@@ -164,7 +165,10 @@ private
y ∼⁺⟨ [ yRz ] ⟩∎
z ∎
+ ix : Vec D 2
ix = i ∷ x ∷ []
+
+ jz : Vec D 2
jz = j ∷ z ∷ []
ix∙⁺jz : Pointwise′ (Plus _R_) ix jz
@@ -180,9 +184,9 @@ private
Pointwise (Plus _R_) xs ys →
Plus (Pointwise _R_) xs ys)
counterexample ∙⁺⇒⁺∙ =
- ¬ix⁺∙jz (Equivalent.to Plus.equivalent ⟨$⟩
- Plus.map (_⟨$⟩_ (Equivalent.to equivalent))
- (∙⁺⇒⁺∙ (Equivalent.from equivalent ⟨$⟩ ix∙⁺jz)))
+ ¬ix⁺∙jz (Equivalence.to Plus.equivalent ⟨$⟩
+ Plus.map (_⟨$⟩_ (Equivalence.to equivalent))
+ (∙⁺⇒⁺∙ (Equivalence.from equivalent ⟨$⟩ ix∙⁺jz)))
-- Map.
diff --git a/src/Relation/Nullary.agda b/src/Relation/Nullary.agda
index c3efa69..088c99c 100644
--- a/src/Relation/Nullary.agda
+++ b/src/Relation/Nullary.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Operations on nullary relations (like negation and decidability)
------------------------------------------------------------------------
diff --git a/src/Relation/Nullary/Core.agda b/src/Relation/Nullary/Core.agda
index 47e5859..d8b0757 100644
--- a/src/Relation/Nullary/Core.agda
+++ b/src/Relation/Nullary/Core.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Nullary relations (some core definitions)
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
-- The definitions in this file are reexported by Relation.Nullary.
module Relation.Nullary.Core where
diff --git a/src/Relation/Nullary/Decidable.agda b/src/Relation/Nullary/Decidable.agda
index d8493cc..3c5a2d5 100644
--- a/src/Relation/Nullary/Decidable.agda
+++ b/src/Relation/Nullary/Decidable.agda
@@ -1,20 +1,22 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Operations on and properties of decidable relations
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Nullary.Decidable where
+open import Data.Bool
open import Data.Empty
+open import Data.Product hiding (map)
+open import Data.Unit
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
+ using (_⇔_; equivalence; module Equivalence)
+open import Level using (Lift)
open import Relation.Binary.PropositionalEquality
+open import Relation.Nullary
⌊_⌋ : ∀ {p} {P : Set p} → Dec P → Bool
⌊ yes _ ⌋ = true
@@ -39,19 +41,31 @@ 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 P⇔Q (yes p) = yes (Equivalent.to P⇔Q ⟨$⟩ p)
-map P⇔Q (no ¬p) = no (¬p ∘ _⟨$⟩_ (Equivalent.from P⇔Q))
+map P⇔Q (yes p) = yes (Equivalence.to P⇔Q ⟨$⟩ p)
+map P⇔Q (no ¬p) = no (¬p ∘ _⟨$⟩_ (Equivalence.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)
+map′ P→Q Q→P = map (equivalence P→Q Q→P)
+
+-- If a decision procedure returns "yes", then we can extract the
+-- proof using from-yes.
+
+From-yes : ∀ {p} {P : Set p} → Dec P → Set p
+From-yes {P = P} (yes _) = P
+From-yes (no _) = Lift ⊤
+
+from-yes : ∀ {p} {P : Set p} (p : Dec P) → From-yes p
+from-yes (yes p) = p
+from-yes (no _) = _
+
+-- If a decision procedure returns "no", then we can extract the proof
+-- using from-no.
-fromYes : ∀ {p} {P : Set p} → P → Dec P → P
-fromYes _ (yes p) = p
-fromYes p (no ¬p) = ⊥-elim (¬p p)
+From-no : ∀ {p} {P : Set p} → Dec P → Set p
+From-no {P = P} (no _) = ¬ P
+From-no (yes _) = Lift ⊤
-fromYes-map-commute :
- ∀ {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)
+from-no : ∀ {p} {P : Set p} (p : Dec P) → From-no p
+from-no (no ¬p) = ¬p
+from-no (yes _) = _
diff --git a/src/Relation/Nullary/Negation.agda b/src/Relation/Nullary/Negation.agda
index a9ec05c..2838317 100644
--- a/src/Relation/Nullary/Negation.agda
+++ b/src/Relation/Nullary/Negation.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Properties related to negation
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Nullary.Negation where
open import Relation.Nullary
@@ -12,7 +12,7 @@ open import Data.Bool
open import Data.Empty
open import Function
open import Data.Product as Prod
-open import Data.Fin
+open import Data.Fin using (Fin)
open import Data.Fin.Dec
open import Data.Sum as Sum
open import Category.Monad
@@ -33,6 +33,12 @@ private
(P → ¬ Q) → Q → ¬ P
note = flip
+-- If we can decide P, then we can decide its negation.
+
+¬? : ∀ {p} {P : Set p} → Dec P → Dec (¬ P)
+¬? (yes p) = no (λ ¬p → ¬p p)
+¬? (no ¬p) = yes ¬p
+
------------------------------------------------------------------------
-- Quantifier juggling
diff --git a/src/Relation/Nullary/Product.agda b/src/Relation/Nullary/Product.agda
index 2cedd3e..4045c65 100644
--- a/src/Relation/Nullary/Product.agda
+++ b/src/Relation/Nullary/Product.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Products of nullary relations
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Nullary.Product where
open import Data.Product
@@ -12,6 +12,8 @@ open import Relation.Nullary
-- Some properties which are preserved by _×_.
+infixr 2 _×-dec_
+
_×-dec_ : ∀ {p q} {P : Set p} {Q : Set q} →
Dec P → Dec Q → Dec (P × Q)
yes p ×-dec yes q = yes (p , q)
diff --git a/src/Relation/Nullary/Sum.agda b/src/Relation/Nullary/Sum.agda
index 2a9f8e4..c3663ab 100644
--- a/src/Relation/Nullary/Sum.agda
+++ b/src/Relation/Nullary/Sum.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Sums of nullary relations
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Nullary.Sum where
open import Data.Sum
diff --git a/src/Relation/Nullary/Universe.agda b/src/Relation/Nullary/Universe.agda
index fb85597..c85f858 100644
--- a/src/Relation/Nullary/Universe.agda
+++ b/src/Relation/Nullary/Universe.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- A universe of proposition functors, along with some properties
------------------------------------------------------------------------
@@ -29,36 +31,36 @@ infix 1 ⟨_⟩_≈_
-- The universe.
-data PropF : Set₁ where
- Id : PropF
- K : (P : Set) → PropF
- _∨_ : (F₁ F₂ : PropF) → PropF
- _∧_ : (F₁ F₂ : PropF) → PropF
- _⇒_ : (P₁ : Set) (F₂ : PropF) → PropF
- ¬¬_ : (F : PropF) → PropF
+data PropF p : Set (suc p) where
+ Id : PropF p
+ K : (P : Set p) → PropF p
+ _∨_ : (F₁ F₂ : PropF p) → PropF p
+ _∧_ : (F₁ F₂ : PropF p) → PropF p
+ _⇒_ : (P₁ : Set p) (F₂ : PropF p) → PropF p
+ ¬¬_ : (F : PropF p) → PropF p
-- Equalities for universe inhabitants.
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)
+ setoid : ∀ {p} → PropF p → Set p → Setoid p p
+ setoid Id P = PropEq.setoid P
+ setoid (K P) _ = PropEq.setoid P
+ setoid (F₁ ∨ F₂) P = (setoid F₁ P) ⊎-setoid (setoid F₂ P)
+ setoid (F₁ ∧ F₂) P = (setoid F₁ P) ×-setoid (setoid F₂ P)
+ setoid (P₁ ⇒ F₂) P = FunS.≡-setoid P₁
+ (Setoid.indexedSetoid (setoid F₂ P))
+ setoid (¬¬ F) P = Always-setoid (¬ ¬ ⟦ F ⟧ P)
- ⟦_⟧ : PropF → (Set → Set)
- ⟦ F ⟧ P = Setoid.Carrier (setoid F {P})
+ ⟦_⟧ : ∀ {p} → PropF p → (Set p → Set p)
+ ⟦ F ⟧ P = Setoid.Carrier (setoid F P)
-⟨_⟩_≈_ : (F : PropF) {P : Set} → Rel (⟦ F ⟧ P) zero
-⟨_⟩_≈_ F = Setoid._≈_ (setoid F)
+⟨_⟩_≈_ : ∀ {p} (F : PropF p) {P : Set p} → Rel (⟦ F ⟧ P) p
+⟨_⟩_≈_ F = Setoid._≈_ (setoid F _)
-- ⟦ F ⟧ is functorial.
-map : ∀ F {P Q} → (P → Q) → ⟦ F ⟧ P → ⟦ F ⟧ Q
+map : ∀ {p} (F : PropF p) {P Q} → (P → Q) → ⟦ F ⟧ P → ⟦ F ⟧ Q
map Id f p = f p
map (K P) f p = p
map (F₁ ∨ F₂) f FP = Sum.map (map F₁ f) (map F₂ f) FP
@@ -66,7 +68,7 @@ map (F₁ ∧ F₂) f FP = Prod.map (map F₁ f) (map F₂ f) FP
map (P₁ ⇒ F₂) f FP = map F₂ f ∘ FP
map (¬¬ F) f FP = ¬¬-map (map F f) FP
-map-id : ∀ F {P} → ⟨ ⟦ F ⟧ P ⇒ F ⟩ map F id ≈ id
+map-id : ∀ {p} (F : PropF p) {P} → ⟨ ⟦ F ⟧ P ⇒ F ⟩ map F id ≈ id
map-id Id x = refl
map-id (K P) x = refl
map-id (F₁ ∨ F₂) (inj₁ x) = ₁∼₁ (map-id F₁ x)
@@ -75,7 +77,7 @@ map-id (F₁ ∧ F₂) (x , y) = (map-id F₁ x , map-id F₂ y)
map-id (P₁ ⇒ F₂) f = λ x → map-id F₂ (f x)
map-id (¬¬ F) ¬¬x = _
-map-∘ : ∀ F {P Q R} (f : Q → R) (g : P → Q) →
+map-∘ : ∀ {p} (F : PropF p) {P Q R} (f : Q → R) (g : P → Q) →
⟨ ⟦ F ⟧ P ⇒ F ⟩ map F f ∘ map F g ≈ map F (f ∘ g)
map-∘ Id f g x = refl
map-∘ (K P) f g x = refl
@@ -88,11 +90,11 @@ map-∘ (¬¬ F) f g x = _
-- A variant of sequence can be implemented for ⟦ F ⟧.
-sequence : ∀ {AF} → RawApplicative AF →
- (AF ⊥ → ⊥) →
- ({A B : Set} → (A → AF B) → AF (A → B)) →
+sequence : ∀ {p AF} → RawApplicative AF →
+ (AF (Lift ⊥) → ⊥) →
+ ({A B : Set p} → (A → AF B) → AF (A → B)) →
∀ F {P} → ⟦ F ⟧ (AF P) → AF (⟦ F ⟧ P)
-sequence {AF} A extract-⊥ sequence-⇒ = helper
+sequence {AF = AF} A extract-⊥ sequence-⇒ = helper
where
open RawApplicative A
@@ -104,16 +106,19 @@ sequence {AF} A extract-⊥ sequence-⇒ = helper
helper (F₁ ∧ F₂) (x , y) = _,_ <$> helper F₁ x ⊛ helper F₂ y
helper (P₁ ⇒ F₂) f = sequence-⇒ (helper F₂ ∘ f)
helper (¬¬ F) x =
- pure (λ ¬FP → x (λ fp → extract-⊥ (¬FP <$> helper F fp)))
+ pure (λ ¬FP → x (λ fp → extract-⊥ (lift ∘ ¬FP <$> helper F fp)))
-- Some lemmas about double negation.
-open RawMonad ¬¬-Monad
+private
+ open module M {p} = RawMonad (¬¬-Monad {p = p})
-¬¬-pull : ∀ F {P} → ⟦ F ⟧ (¬ ¬ P) → ¬ ¬ ⟦ F ⟧ P
+¬¬-pull : ∀ {p} (F : PropF p) {P} →
+ ⟦ F ⟧ (¬ ¬ P) → ¬ ¬ ⟦ F ⟧ P
¬¬-pull = sequence rawIApplicative
- (λ f → f id)
+ (λ f → f lower)
(λ f g → g (λ x → ⊥-elim (f x (λ y → g (λ _ → y)))))
-¬¬-remove : ∀ F {P} → ¬ ¬ ⟦ F ⟧ (¬ ¬ P) → ¬ ¬ ⟦ F ⟧ P
+¬¬-remove : ∀ {p} (F : PropF p) {P} →
+ ¬ ¬ ⟦ F ⟧ (¬ ¬ P) → ¬ ¬ ⟦ F ⟧ P
¬¬-remove F = negated-stable ∘ ¬¬-pull (¬¬ F)
diff --git a/src/Relation/Unary.agda b/src/Relation/Unary.agda
index 562415a..dc00b32 100644
--- a/src/Relation/Unary.agda
+++ b/src/Relation/Unary.agda
@@ -1,9 +1,9 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Unary relations
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Relation.Unary where
open import Data.Empty
diff --git a/src/Size.agda b/src/Size.agda
index 4cb481b..25614ac 100644
--- a/src/Size.agda
+++ b/src/Size.agda
@@ -1,4 +1,6 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Sizes for Agda's sized types
------------------------------------------------------------------------
diff --git a/src/Universe.agda b/src/Universe.agda
index 4923a30..8949e96 100644
--- a/src/Universe.agda
+++ b/src/Universe.agda
@@ -1,11 +1,13 @@
------------------------------------------------------------------------
+-- The Agda standard library
+--
-- Universes
------------------------------------------------------------------------
-{-# OPTIONS --universe-polymorphism #-}
-
module Universe where
+open import Data.Product
+open import Function
open import Level
-- Universes.
@@ -17,3 +19,24 @@ record Universe u e : Set (suc (u ⊔ e)) where
-- Decoding function.
El : U → Set e
+
+-- Indexed universes.
+
+record Indexed-universe i u e : Set (suc (i ⊔ u ⊔ e)) where
+ field
+ -- Index set.
+ I : Set i
+
+ -- Codes.
+ U : I → Set u
+
+ -- Decoding function.
+ El : ∀ {i} → U i → Set e
+
+ -- An indexed universe can be turned into an unindexed one.
+
+ unindexed-universe : Universe (i ⊔ u) e
+ unindexed-universe = record
+ { U = ∃ λ i → U i
+ ; El = El ∘ proj₂
+ }