summaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorIain Lane <iain@orangesquash.org.uk>2014-08-05 09:11:47 +0100
committerIain Lane <iain@orangesquash.org.uk>2014-08-05 09:11:47 +0100
commit84d43131c1027130bd54e74a699fd1132cda66de (patch)
tree5f0cb3f183c8109d9c8d659219d3b02cb3cf649a
parent6d5228941fe84e810ba3e49cd3f2bbee902bdeef (diff)
Imported Upstream version 0.8
-rw-r--r--.gitignore7
-rw-r--r--AllNonAsciiChars.hs11
-rw-r--r--Everything.agda564
-rw-r--r--GNUmakefile5
-rw-r--r--LICENCE5
-rw-r--r--README.agda25
-rw-r--r--README.md6
-rw-r--r--README/AVL.agda8
-rw-r--r--README/Container/FreeMonad.agda64
-rw-r--r--ffi/agda-lib-ffi.cabal2
-rw-r--r--lib.cabal6
-rwxr-xr-xpublish-listings.sh22
-rw-r--r--release-notes22
-rw-r--r--src/Algebra/Monoid-solver.agda136
-rw-r--r--src/Algebra/Morphism.agda2
-rw-r--r--src/Algebra/Properties/AbelianGroup.agda44
-rw-r--r--src/Algebra/Properties/BooleanAlgebra.agda570
-rw-r--r--src/Algebra/Properties/BooleanAlgebra/Expression.agda (renamed from src/Algebra/Props/BooleanAlgebra/Expression.agda)2
-rw-r--r--src/Algebra/Properties/DistributiveLattice.agda85
-rw-r--r--src/Algebra/Properties/Group.agda70
-rw-r--r--src/Algebra/Properties/Lattice.agda106
-rw-r--r--src/Algebra/Properties/Ring.agda51
-rw-r--r--src/Algebra/Props/AbelianGroup.agda37
-rw-r--r--src/Algebra/Props/BooleanAlgebra.agda562
-rw-r--r--src/Algebra/Props/DistributiveLattice.agda77
-rw-r--r--src/Algebra/Props/Group.agda64
-rw-r--r--src/Algebra/Props/Lattice.agda100
-rw-r--r--src/Algebra/Props/Ring.agda45
-rw-r--r--src/Algebra/RingSolver/AlmostCommutativeRing.agda4
-rw-r--r--src/Category/Applicative/Predicate.agda48
-rw-r--r--src/Category/Functor/Predicate.agda25
-rw-r--r--src/Category/Monad/Indexed.agda12
-rw-r--r--src/Category/Monad/Predicate.agda70
-rw-r--r--src/Coinduction.agda3
-rw-r--r--src/Data/AVL.agda6
-rw-r--r--src/Data/AVL/IndexedMap.agda7
-rw-r--r--src/Data/AVL/Sets.agda7
-rw-r--r--src/Data/Bin.agda2
-rw-r--r--src/Data/Bool/Properties.agda12
-rw-r--r--src/Data/Char.agda31
-rw-r--r--src/Data/Colist.agda250
-rw-r--r--src/Data/Colist/Infinite-merge.agda231
-rw-r--r--src/Data/Container/Any.agda4
-rw-r--r--src/Data/Container/FreeMonad.agda55
-rw-r--r--src/Data/Container/Indexed.agda351
-rw-r--r--src/Data/Container/Indexed/Combinator.agda297
-rw-r--r--src/Data/Container/Indexed/Core.agda43
-rw-r--r--src/Data/Container/Indexed/FreeMonad.agda58
-rw-r--r--src/Data/Fin/Dec.agda2
-rw-r--r--src/Data/Fin/Properties.agda240
-rw-r--r--src/Data/Fin/Props.agda239
-rw-r--r--src/Data/Fin/Subset.agda4
-rw-r--r--src/Data/Fin/Subset/Properties.agda156
-rw-r--r--src/Data/Fin/Subset/Props.agda152
-rw-r--r--src/Data/Graph/Acyclic.agda4
-rw-r--r--src/Data/Integer.agda2
-rw-r--r--src/Data/Integer/Properties.agda34
-rw-r--r--src/Data/List/All/Properties.agda71
-rw-r--r--src/Data/List/Any/Membership.agda4
-rw-r--r--src/Data/List/Any/Properties.agda3
-rw-r--r--src/Data/List/NonEmpty.agda210
-rw-r--r--src/Data/List/NonEmpty/Properties.agda25
-rw-r--r--src/Data/List/Properties.agda9
-rw-r--r--src/Data/M/Indexed.agda42
-rw-r--r--src/Data/Maybe.agda58
-rw-r--r--src/Data/Nat.agda19
-rw-r--r--src/Data/Nat/Coprimality.agda2
-rw-r--r--src/Data/Nat/DivMod.agda2
-rw-r--r--src/Data/Nat/Divisibility.agda2
-rw-r--r--src/Data/Nat/Properties.agda122
-rw-r--r--src/Data/Nat/Properties/Simple.agda110
-rw-r--r--src/Data/Rational.agda136
-rw-r--r--src/Data/Star/BoundedVec.agda5
-rw-r--r--src/Data/Stream.agda22
-rw-r--r--src/Data/String.agda37
-rw-r--r--src/Data/Sum.agda4
-rw-r--r--src/Data/Vec.agda2
-rw-r--r--src/Data/Vec/Properties.agda53
-rw-r--r--src/Data/W/Indexed.agda42
-rw-r--r--src/Function.agda6
-rw-r--r--src/Function/Equality.agda13
-rw-r--r--src/Function/Inverse.agda8
-rw-r--r--src/Function/Related.agda6
-rw-r--r--src/Function/Related/TypeIsomorphisms.agda89
-rw-r--r--src/Induction.agda21
-rw-r--r--src/Induction/Lexicographic.agda22
-rw-r--r--src/Induction/Nat.agda200
-rw-r--r--src/Induction/WellFounded.agda34
-rw-r--r--src/Level.agda23
-rw-r--r--src/Record.agda4
-rw-r--r--src/Reflection.agda48
-rw-r--r--src/Relation/Binary.agda52
-rw-r--r--src/Relation/Binary/HeterogeneousEquality.agda42
-rw-r--r--src/Relation/Binary/HeterogeneousEquality/Core.agda4
-rw-r--r--src/Relation/Binary/List/Pointwise.agda47
-rw-r--r--src/Relation/Binary/Product/Pointwise.agda24
-rw-r--r--src/Relation/Binary/Properties/DecTotalOrder.agda26
-rw-r--r--src/Relation/Binary/Properties/Poset.agda29
-rw-r--r--src/Relation/Binary/Properties/Preorder.agda28
-rw-r--r--src/Relation/Binary/Properties/StrictPartialOrder.agda32
-rw-r--r--src/Relation/Binary/Properties/StrictTotalOrder.agda34
-rw-r--r--src/Relation/Binary/Properties/TotalOrder.agda22
-rw-r--r--src/Relation/Binary/PropositionalEquality.agda4
-rw-r--r--src/Relation/Binary/Props/DecTotalOrder.agda19
-rw-r--r--src/Relation/Binary/Props/Poset.agda22
-rw-r--r--src/Relation/Binary/Props/Preorder.agda21
-rw-r--r--src/Relation/Binary/Props/StrictPartialOrder.agda25
-rw-r--r--src/Relation/Binary/Props/StrictTotalOrder.agda26
-rw-r--r--src/Relation/Binary/Props/TotalOrder.agda15
-rw-r--r--src/Relation/Binary/SetoidReasoning.agda46
-rw-r--r--src/Relation/Binary/StrictPartialOrderReasoning.agda2
-rw-r--r--src/Relation/Binary/Sum.agda28
-rw-r--r--src/Relation/Nullary/Decidable.agda19
-rw-r--r--src/Relation/Nullary/Sum.agda2
-rw-r--r--src/Relation/Unary.agda74
-rw-r--r--src/Relation/Unary/PredicateTransformer.agda117
-rw-r--r--src/Size.agda14
117 files changed, 4823 insertions, 2452 deletions
diff --git a/.gitignore b/.gitignore
new file mode 100644
index 0000000..1435974
--- /dev/null
+++ b/.gitignore
@@ -0,0 +1,7 @@
+dist
+html
+Everything.agda
+*.agdai
+*.agda.el
+*.lagda.el
+.*.swp
diff --git a/AllNonAsciiChars.hs b/AllNonAsciiChars.hs
index fb90dbd..b87e8a0 100644
--- a/AllNonAsciiChars.hs
+++ b/AllNonAsciiChars.hs
@@ -7,6 +7,7 @@ import qualified Data.List as L
import Data.Char
import Data.Function
import Control.Applicative
+import Numeric ( showHex )
import System.FilePath.Find
import System.IO
@@ -27,4 +28,12 @@ main = do
L.sortBy (compare `on` snd) $
map (\cs -> (head cs, length cs)) $
L.group $ L.sort $ nonAsciiChars
- mapM_ (\(c, count) -> putStrLn (c : ": " ++ show count)) table
+
+ let codePoint :: Char -> String
+ codePoint c = showHex (ord c) ""
+
+ uPlus :: Char -> String
+ uPlus c = "(U+" ++ codePoint c ++ ")"
+
+ mapM_ (\(c, count) -> putStrLn (c : " " ++ uPlus c ++ ": " ++ show count))
+ table
diff --git a/Everything.agda b/Everything.agda
deleted file mode 100644
index 1703c01..0000000
--- a/Everything.agda
+++ /dev/null
@@ -1,564 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- All library modules, along with short descriptions
-------------------------------------------------------------------------
-
--- Note that core modules are not included.
-
-module Everything where
-
--- Definitions of algebraic structures like monoids and rings
--- (packed in records together with sets, operations, etc.)
-import Algebra
-
--- Properties of functions, such as associativity and commutativity
-import Algebra.FunctionProperties
-
--- Morphisms between algebraic structures
-import Algebra.Morphism
-
--- Some defined operations (multiplication by natural number and
--- exponentiation)
-import Algebra.Operations
-
--- Some derivable properties
-import Algebra.Props.AbelianGroup
-
--- Some derivable properties
-import Algebra.Props.BooleanAlgebra
-
--- Boolean algebra expressions
-import Algebra.Props.BooleanAlgebra.Expression
-
--- Some derivable properties
-import Algebra.Props.DistributiveLattice
-
--- Some derivable properties
-import Algebra.Props.Group
-
--- Some derivable properties
-import Algebra.Props.Lattice
-
--- Some derivable properties
-import Algebra.Props.Ring
-
--- Solver for commutative ring or semiring equalities
-import Algebra.RingSolver
-
--- Commutative semirings with some additional structure ("almost"
--- commutative rings), used by the ring solver
-import Algebra.RingSolver.AlmostCommutativeRing
-
--- Some boring lemmas used by the ring solver
-import Algebra.RingSolver.Lemmas
-
--- Instantiates the ring solver, using the natural numbers as the
--- coefficient "ring"
-import Algebra.RingSolver.Natural-coefficients
-
--- Instantiates the ring solver with two copies of the same ring with
--- decidable equality
-import Algebra.RingSolver.Simple
-
--- Some algebraic structures (not packed up with sets, operations,
--- etc.)
-import Algebra.Structures
-
--- Applicative functors
-import Category.Applicative
-
--- Indexed applicative functors
-import Category.Applicative.Indexed
-
--- Functors
-import Category.Functor
-
--- Monads
-import Category.Monad
-
--- A delimited continuation monad
-import Category.Monad.Continuation
-
--- The identity monad
-import Category.Monad.Identity
-
--- Indexed monads
-import Category.Monad.Indexed
-
--- The partiality monad
-import Category.Monad.Partiality
-
--- An All predicate for the partiality monad
-import Category.Monad.Partiality.All
-
--- The state monad
-import Category.Monad.State
-
--- Basic types related to coinduction
-import Coinduction
-
--- AVL trees
-import Data.AVL
-
--- Finite maps with indexed keys and values, based on AVL trees
-import Data.AVL.IndexedMap
-
--- Finite sets, based on AVL trees
-import Data.AVL.Sets
-
--- A binary representation of natural numbers
-import Data.Bin
-
--- Booleans
-import Data.Bool
-
--- A bunch of properties
-import Data.Bool.Properties
-
--- Showing booleans
-import Data.Bool.Show
-
--- Bounded vectors
-import Data.BoundedVec
-
--- Bounded vectors (inefficient, concrete implementation)
-import Data.BoundedVec.Inefficient
-
--- Characters
-import Data.Char
-
--- "Finite" sets indexed on coinductive "natural" numbers
-import Data.Cofin
-
--- Coinductive lists
-import Data.Colist
-
--- Coinductive "natural" numbers
-import Data.Conat
-
--- Containers, based on the work of Abbott and others
-import Data.Container
-
--- Properties related to ◇
-import Data.Container.Any
-
--- Container combinators
-import Data.Container.Combinator
-
--- Coinductive vectors
-import Data.Covec
-
--- Lists with fast append
-import Data.DifferenceList
-
--- Natural numbers with fast addition (for use together with
--- DifferenceVec)
-import Data.DifferenceNat
-
--- Vectors with fast append
-import Data.DifferenceVec
-
--- Digits and digit expansions
-import Data.Digit
-
--- Empty type
-import Data.Empty
-
--- Finite sets
-import Data.Fin
-
--- Decision procedures for finite sets and subsets of finite sets
-import Data.Fin.Dec
-
--- Properties related to Fin, and operations making use of these
--- properties (or other properties not available in Data.Fin)
-import Data.Fin.Props
-
--- Subsets of finite sets
-import Data.Fin.Subset
-
--- Some properties about subsets
-import Data.Fin.Subset.Props
-
--- Substitutions
-import Data.Fin.Substitution
-
--- An example of how Data.Fin.Substitution can be used: a definition
--- of substitution for the untyped λ-calculus, along with some lemmas
-import Data.Fin.Substitution.Example
-
--- Substitution lemmas
-import Data.Fin.Substitution.Lemmas
-
--- Application of substitutions to lists, along with various lemmas
-import Data.Fin.Substitution.List
-
--- Directed acyclic multigraphs
-import Data.Graph.Acyclic
-
--- Integers
-import Data.Integer
-
--- Properties related to addition of integers
-import Data.Integer.Addition.Properties
-
--- Divisibility and coprimality
-import Data.Integer.Divisibility
-
--- Properties related to multiplication of integers
-import Data.Integer.Multiplication.Properties
-
--- Some properties about integers
-import Data.Integer.Properties
-
--- Lists
-import Data.List
-
--- Lists where all elements satisfy a given property
-import Data.List.All
-
--- Properties relating All to various list functions
-import Data.List.All.Properties
-
--- Lists where at least one element satisfies a given property
-import Data.List.Any
-
--- Properties related to bag and set equality
-import Data.List.Any.BagAndSetEquality
-
--- Properties related to list membership
-import Data.List.Any.Membership
-
--- Properties related to Any
-import Data.List.Any.Properties
-
--- A data structure which keeps track of an upper bound on the number
--- of elements /not/ in a given list
-import Data.List.Countdown
-
--- Non-empty lists
-import Data.List.NonEmpty
-
--- Properties of non-empty lists
-import Data.List.NonEmpty.Properties
-
--- List-related properties
-import Data.List.Properties
-
--- Reverse view
-import Data.List.Reverse
-
--- M-types (the dual of W-types)
-import Data.M
-
--- The Maybe type
-import Data.Maybe
-
--- Natural numbers
-import Data.Nat
-
--- Coprimality
-import Data.Nat.Coprimality
-
--- Integer division
-import Data.Nat.DivMod
-
--- Divisibility
-import Data.Nat.Divisibility
-
--- Greatest common divisor
-import Data.Nat.GCD
-
--- Boring lemmas used in Data.Nat.GCD and Data.Nat.Coprimality
-import Data.Nat.GCD.Lemmas
-
--- Definition of and lemmas related to "true infinitely often"
-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
-
--- Showing natural numbers
-import Data.Nat.Show
-
--- Transitive closures
-import Data.Plus
-
--- Products
-import Data.Product
-
--- N-ary products
-import Data.Product.N-ary
-
--- Rational numbers
-import Data.Rational
-
--- Reflexive closures
-import Data.ReflexiveClosure
-
--- Signs
-import Data.Sign
-
--- Some properties about signs
-import Data.Sign.Properties
-
--- The reflexive transitive closures of McBride, Norell and Jansson
-import Data.Star
-
--- Bounded vectors (inefficient implementation)
-import Data.Star.BoundedVec
-
--- Decorated star-lists
-import Data.Star.Decoration
-
--- Environments (heterogeneous collections)
-import Data.Star.Environment
-
--- Finite sets defined in terms of Data.Star
-import Data.Star.Fin
-
--- Lists defined in terms of Data.Star
-import Data.Star.List
-
--- Natural numbers defined in terms of Data.Star
-import Data.Star.Nat
-
--- Pointers into star-lists
-import Data.Star.Pointer
-
--- Some properties related to Data.Star
-import Data.Star.Properties
-
--- Vectors defined in terms of Data.Star
-import Data.Star.Vec
-
--- Streams
-import Data.Stream
-
--- Strings
-import Data.String
-
--- Sums (disjoint unions)
-import Data.Sum
-
--- Some unit types
-import Data.Unit
-
--- Vectors
-import Data.Vec
-
--- Semi-heterogeneous vector equality
-import Data.Vec.Equality
-
--- Code for converting Vec A n → B to and from n-ary functions
-import Data.Vec.N-ary
-
--- Some Vec-related properties
-import Data.Vec.Properties
-
--- W-types
-import Data.W
-
--- Type(s) used (only) when calling out to Haskell via the FFI
-import Foreign.Haskell
-
--- Simple combinators working solely on and with functions
-import Function
-
--- Bijections
-import Function.Bijection
-
--- Function setoids and related constructions
-import Function.Equality
-
--- Equivalence (coinhabitance)
-import Function.Equivalence
-
--- Injections
-import Function.Injection
-
--- Inverses
-import Function.Inverse
-
--- 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
-
--- IO
-import IO
-
--- Primitive IO: simple bindings to Haskell types and functions
-import IO.Primitive
-
--- An abstraction of various forms of recursion/induction
-import Induction
-
--- Lexicographic induction
-import Induction.Lexicographic
-
--- Various forms of induction for natural numbers
-import Induction.Nat
-
--- Well-founded induction
-import Induction.WellFounded
-
--- The irrelevance axiom
-import Irrelevance
-
--- 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
-
--- Properties of homogeneous binary relations
-import Relation.Binary
-
--- Some properties imply others
-import Relation.Binary.Consequences
-
--- Convenient syntax for equational reasoning
-import Relation.Binary.EqReasoning
-
--- Many properties which hold for _∼_ also hold for flip _∼_
-import Relation.Binary.Flip
-
--- Heterogeneous equality
-import Relation.Binary.HeterogeneousEquality
-
--- Indexed binary relations
-import Relation.Binary.Indexed
-
--- Induced preorders
-import Relation.Binary.InducedPreorders
-
--- Lexicographic ordering of lists
-import Relation.Binary.List.NonStrictLex
-
--- Pointwise lifting of relations to lists
-import Relation.Binary.List.Pointwise
-
--- Lexicographic ordering of lists
-import Relation.Binary.List.StrictLex
-
--- Conversion of ≤ to <, along with a number of properties
-import Relation.Binary.NonStrictToStrict
-
--- Many properties which hold for _∼_ also hold for _∼_ on f
-import Relation.Binary.On
-
--- Order morphisms
-import Relation.Binary.OrderMorphism
-
--- Convenient syntax for "equational reasoning" using a partial order
-import Relation.Binary.PartialOrderReasoning
-
--- Convenient syntax for "equational reasoning" using a preorder
-import Relation.Binary.PreorderReasoning
-
--- Lexicographic products of binary relations
-import Relation.Binary.Product.NonStrictLex
-
--- Pointwise products of binary relations
-import Relation.Binary.Product.Pointwise
-
--- Lexicographic products of binary relations
-import Relation.Binary.Product.StrictLex
-
--- Propositional (intensional) equality
-import Relation.Binary.PropositionalEquality
-
--- An equality postulate which evaluates
-import Relation.Binary.PropositionalEquality.TrustMe
-
--- Properties satisfied by decidable total orders
-import Relation.Binary.Props.DecTotalOrder
-
--- Properties satisfied by posets
-import Relation.Binary.Props.Poset
-
--- Properties satisfied by preorders
-import Relation.Binary.Props.Preorder
-
--- Properties satisfied by strict partial orders
-import Relation.Binary.Props.StrictPartialOrder
-
--- Properties satisfied by strict partial orders
-import Relation.Binary.Props.StrictTotalOrder
-
--- Properties satisfied by total orders
-import Relation.Binary.Props.TotalOrder
-
--- Helpers intended to ease the development of "tactics" which use
--- proof by reflection
-import Relation.Binary.Reflection
-
--- Pointwise lifting of binary relations to sigma types
-import Relation.Binary.Sigma.Pointwise
-
--- Some simple binary relations
-import Relation.Binary.Simple
-
--- Convenient syntax for "equational reasoning" using a strict partial
--- order
-import Relation.Binary.StrictPartialOrderReasoning
-
--- Conversion of < to ≤, along with a number of properties
-import Relation.Binary.StrictToNonStrict
-
--- Sums of binary relations
-import Relation.Binary.Sum
-
--- Pointwise lifting of relations to vectors
-import Relation.Binary.Vec.Pointwise
-
--- Operations on nullary relations (like negation and decidability)
-import Relation.Nullary
-
--- Operations on and properties of decidable relations
-import Relation.Nullary.Decidable
-
--- Implications of nullary relations
-import Relation.Nullary.Implication
-
--- Properties related to negation
-import Relation.Nullary.Negation
-
--- Products of nullary relations
-import Relation.Nullary.Product
-
--- Sums of nullary relations
-import Relation.Nullary.Sum
-
--- A universe of proposition functors, along with some properties
-import Relation.Nullary.Universe
-
--- Unary relations
-import Relation.Unary
-
--- Sizes for Agda's sized types
-import Size
-
--- Universes
-import Universe
diff --git a/GNUmakefile b/GNUmakefile
index 342f661..f0a6606 100644
--- a/GNUmakefile
+++ b/GNUmakefile
@@ -13,3 +13,8 @@ Everything.agda:
.PHONY: agda-lib-ffi
agda-lib-ffi:
cd ffi && cabal install
+
+.PHONY: listings
+listings: Everything.agda
+ $(AGDA) -i. -isrc --html README.agda -v0
+
diff --git a/LICENCE b/LICENCE
index f57f794..2620467 100644
--- a/LICENCE
+++ b/LICENCE
@@ -1,10 +1,11 @@
-Copyright (c) 2007-2013 Nils Anders Danielsson, Ulf Norell, Shin-Cheng
+Copyright (c) 2007-2014 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, Andreas Abel, Alcatel-Lucent, Eric Mertens, Joachim
Breitner, Liyang Hu, Noam Zeilberger, Érdi Gergő, Stevan Andjelkovic,
-Helmut Grohne
+Helmut Grohne, Guilhem Moulin, Noriyuki OHKAWA, Evgeny Kotelnikov,
+James Chapman.
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 4374f70..40e3ecc 100644
--- a/README.agda
+++ b/README.agda
@@ -1,18 +1,19 @@
module README where
------------------------------------------------------------------------
--- The Agda standard library, version 0.7
+-- The Agda standard library, version 0.8
--
-- Author: Nils Anders Danielsson, with contributions from Andreas
-- Abel, Stevan Andjelkovic, Jean-Philippe Bernardy, Peter Berry,
--- Joachim Breitner, Samuel Bronson, Daniel Brown, Liang-Ting Chen,
--- Dominique Devriese, Dan Doel, Érdi Gergő, Helmut Grohne, Simon
--- Foster, Liyang Hu, Patrik Jansson, Alan Jeffrey, Eric Mertens,
--- Darin Morrison, Shin-Cheng Mu, Ulf Norell, Nicolas Pouillard,
--- Andrés Sicard-Ramírez and Noam Zeilberger
+-- Joachim Breitner, Samuel Bronson, Daniel Brown, James Chapman,
+-- Liang-Ting Chen, Dominique Devriese, Dan Doel, Érdi Gergő, Helmut
+-- Grohne, Simon Foster, Liyang Hu, Patrik Jansson, Alan Jeffrey,
+-- Evgeny Kotelnikov, Eric Mertens, Darin Morrison, Guilhem Moulin,
+-- Shin-Cheng Mu, Ulf Norell, Noriyuki OHKAWA, Nicolas Pouillard,
+-- Andrés Sicard-Ramírez and Noam Zeilberger.
------------------------------------------------------------------------
--- This version of the library has been tested using Agda 2.3.2.
+-- This version of the library has been tested using Agda 2.4.0.
-- Note that no guarantees are currently made about forwards or
-- backwards compatibility, the library is still at an experimental
@@ -34,12 +35,10 @@ module README where
-- 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
--- darcs send --output <patch file> and attach the patch file to an
--- email), and include a statement saying that you agree to release
--- your library patches under the library's licence. It is appreciated
--- if every patch contains a single, complete change, and if the
--- coding style used in the library is adhered to.
+-- them). Please send contributions in the form of git pull requests, patch
+-- bundles or ask for commmit rights to the repository. It is appreciated if
+-- every patch contains a single, complete change, and if the coding style used
+-- in the library is adhered to.
------------------------------------------------------------------------
-- Module hierarchy
diff --git a/README.md b/README.md
new file mode 100644
index 0000000..9100ee8
--- /dev/null
+++ b/README.md
@@ -0,0 +1,6 @@
+agda-stdlib
+===========
+
+The Agda standard library. You can browse the source in glorious clickable html here:
+
+http://agda.github.io/agda-stdlib/html/README.html
diff --git a/README/AVL.agda b/README/AVL.agda
index fcea8f7..181a8dc 100644
--- a/README/AVL.agda
+++ b/README/AVL.agda
@@ -103,23 +103,19 @@ q₅′ = refl
-- 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₂ ≡
+v₇ : Maybe.map (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₄ ≡
+v₉ : Maybe.map (Prod.map toList id) (initLast t₄) ≡
just (((1 , v₁) ∷ []) ,′ (2 , v₂))
v₉ = refl
diff --git a/README/Container/FreeMonad.agda b/README/Container/FreeMonad.agda
new file mode 100644
index 0000000..d2cbd1a
--- /dev/null
+++ b/README/Container/FreeMonad.agda
@@ -0,0 +1,64 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Example showing how the free monad construction on containers can be
+-- used
+------------------------------------------------------------------------
+
+module README.Container.FreeMonad where
+
+open import Category.Monad
+open import Data.Empty
+open import Data.Unit
+open import Data.Bool
+open import Data.Nat
+open import Data.Sum using (inj₁; inj₂)
+open import Data.Product renaming (_×_ to _⟨×⟩_)
+open import Data.Container
+open import Data.Container.Combinator
+open import Data.Container.FreeMonad
+open import Data.W
+open import Relation.Binary.PropositionalEquality
+
+------------------------------------------------------------------------
+
+-- The signature of state and its (generic) operations.
+
+State : Set → Container _
+State S = ⊤ ⟶ S ⊎ S ⟶ ⊤
+ where
+ _⟶_ : Set → Set → Container _
+ I ⟶ O = I ▷ λ _ → O
+
+get : ∀ {S} → State S ⋆ S
+get = do (inj₁ _ , return)
+ where
+ open RawMonad rawMonad
+
+put : ∀ {S} → S → State S ⋆ ⊤
+put s = do (inj₂ s , return)
+ where
+ open RawMonad rawMonad
+
+-- Using the above we can, for example, write a stateful program that
+-- delivers a boolean.
+prog : State ℕ ⋆ Bool
+prog =
+ get >>= λ n →
+ put (suc n) >>
+ return true
+ where
+ open RawMonad rawMonad
+
+runState : ∀ {S X} → State S ⋆ X → (S → X ⟨×⟩ S)
+runState (sup (inj₁ x) _) = λ s → x , s
+runState (sup (inj₂ (inj₁ _)) k) = λ s → runState (k s) s
+runState (sup (inj₂ (inj₂ s)) k) = λ _ → runState (k _) s
+
+test : runState prog 0 ≡ true , 1
+test = refl
+
+-- It should be noted that @State S ⋆ X@ is not the state monad. If we
+-- could quotient @State S ⋆ X@ by the seven axioms of state (see
+-- Plotkin and Power's "Notions of Computation Determine Monads", 2002)
+-- then we would get the state monad.
diff --git a/ffi/agda-lib-ffi.cabal b/ffi/agda-lib-ffi.cabal
index a7a32f9..dac3abf 100644
--- a/ffi/agda-lib-ffi.cabal
+++ b/ffi/agda-lib-ffi.cabal
@@ -5,6 +5,6 @@ build-type: Simple
description: Auxiliary Haskell code used by Agda's standard library.
library
- build-depends: base >= 3.0.3.1 && < 4.7
+ build-depends: base >= 3.0.3.1 && < 4.8
exposed-modules: Data.FFI
IO.FFI
diff --git a/lib.cabal b/lib.cabal
index 334023b..fbb447a 100644
--- a/lib.cabal
+++ b/lib.cabal
@@ -1,5 +1,5 @@
name: lib
-version: 0.1.0.1
+version: 0.1.0.2
cabal-version: >= 1.8
build-type: Simple
description: Helper programs.
@@ -7,12 +7,12 @@ description: Helper programs.
executable GenerateEverything
hs-source-dirs: .
main-is: GenerateEverything.hs
- build-depends: base >= 4.2 && < 4.7,
+ build-depends: base >= 4.2 && < 4.8,
filemanip == 0.3.*,
filepath >= 1.1 && < 1.4
executable AllNonAsciiChars
hs-source-dirs: .
main-is: AllNonAsciiChars.hs
- build-depends: base >= 4.2 && < 4.7,
+ build-depends: base >= 4.2 && < 4.8,
filemanip == 0.3.*
diff --git a/publish-listings.sh b/publish-listings.sh
new file mode 100755
index 0000000..dfe104b
--- /dev/null
+++ b/publish-listings.sh
@@ -0,0 +1,22 @@
+#!/bin/bash
+
+cd /tmp
+git clone git@github.com:agda/agda-stdlib.git
+cd agda-stdlib
+git checkout gh-pages
+git merge master -m "[auto] merge master into gh-pages"
+make listings
+
+if [ "`git status --porcelain`" != "" ]; then
+ echo "Updates:"
+ git status --porcelain
+ changed=`git status --porcelain | cut -c4-`
+ git add --all -- $changed
+ git commit -m "[auto] updated html listings"
+ git push
+else
+ echo "No changes!"
+fi
+
+cd ..
+rm -rf agda-stdlib
diff --git a/release-notes b/release-notes
index c755e35..345e7d7 100644
--- a/release-notes
+++ b/release-notes
@@ -1,4 +1,26 @@
------------------------------------------------------------------------
+Version 0.8
+------------------------------------------------------------------------
+
+Version 0.8 of the standard library has now been released, see
+http://wiki.portal.chalmers.se/agda/pmwiki.php?n=Libraries.StandardLibrary.
+
+The library has been tested using Agda version 2.4.0.
+
+Note that no guarantees are made about backwards or forwards
+compatibility, the library is still at an experimental stage.
+
+If you want to compile the library using the MAlonzo compiler, then
+you should first 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.
+
+------------------------------------------------------------------------
Version 0.7
------------------------------------------------------------------------
diff --git a/src/Algebra/Monoid-solver.agda b/src/Algebra/Monoid-solver.agda
new file mode 100644
index 0000000..97af872
--- /dev/null
+++ b/src/Algebra/Monoid-solver.agda
@@ -0,0 +1,136 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Solver for monoid equalities
+------------------------------------------------------------------------
+
+open import Algebra
+
+module Algebra.Monoid-solver {m₁ m₂} (M : Monoid m₁ m₂) where
+
+open import Data.Fin
+import Data.Fin.Properties as Fin
+open import Data.List
+open import Data.Maybe as Maybe
+ using (Maybe; decToMaybe; From-just; from-just)
+open import Data.Nat using (ℕ)
+open import Data.Product
+open import Data.Vec using (Vec; lookup)
+open import Function using (_∘_; _$_)
+import Relation.Binary.EqReasoning
+import Relation.Binary.List.Pointwise as Pointwise
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+import Relation.Binary.Reflection
+open import Relation.Nullary
+import Relation.Nullary.Decidable as Dec
+
+open Monoid M
+open Relation.Binary.EqReasoning setoid
+
+------------------------------------------------------------------------
+-- Monoid expressions
+
+-- There is one constructor for every operation, plus one for
+-- variables; there may be at most n variables.
+
+infixr 5 _⊕_
+
+data Expr (n : ℕ) : Set where
+ var : Fin n → Expr n
+ id : Expr n
+ _⊕_ : Expr n → Expr n → Expr n
+
+-- An environment contains one value for every variable.
+
+Env : ℕ → Set _
+Env n = Vec Carrier n
+
+-- The semantics of an expression is a function from an environment to
+-- a value.
+
+⟦_⟧ : ∀ {n} → Expr n → Env n → Carrier
+⟦ var x ⟧ ρ = lookup x ρ
+⟦ id ⟧ ρ = ε
+⟦ e₁ ⊕ e₂ ⟧ ρ = ⟦ e₁ ⟧ ρ ∙ ⟦ e₂ ⟧ ρ
+
+------------------------------------------------------------------------
+-- Normal forms
+
+-- A normal form is a list of variables.
+
+Normal : ℕ → Set
+Normal n = List (Fin n)
+
+-- The semantics of a normal form.
+
+⟦_⟧⇓ : ∀ {n} → Normal n → Env n → Carrier
+⟦ [] ⟧⇓ ρ = ε
+⟦ x ∷ nf ⟧⇓ ρ = lookup x ρ ∙ ⟦ nf ⟧⇓ ρ
+
+-- A normaliser.
+
+normalise : ∀ {n} → Expr n → Normal n
+normalise (var x) = x ∷ []
+normalise id = []
+normalise (e₁ ⊕ e₂) = normalise e₁ ++ normalise e₂
+
+-- The normaliser is homomorphic with respect to _++_/_∙_.
+
+homomorphic : ∀ {n} (nf₁ nf₂ : Normal n) (ρ : Env n) →
+ ⟦ nf₁ ++ nf₂ ⟧⇓ ρ ≈ (⟦ nf₁ ⟧⇓ ρ ∙ ⟦ nf₂ ⟧⇓ ρ)
+homomorphic [] nf₂ ρ = begin
+ ⟦ nf₂ ⟧⇓ ρ ≈⟨ sym $ proj₁ identity _ ⟩
+ ε ∙ ⟦ nf₂ ⟧⇓ ρ ∎
+homomorphic (x ∷ nf₁) nf₂ ρ = begin
+ lookup x ρ ∙ ⟦ nf₁ ++ nf₂ ⟧⇓ ρ ≈⟨ ∙-cong refl (homomorphic nf₁ nf₂ ρ) ⟩
+ lookup x ρ ∙ (⟦ nf₁ ⟧⇓ ρ ∙ ⟦ nf₂ ⟧⇓ ρ) ≈⟨ sym $ assoc _ _ _ ⟩
+ lookup x ρ ∙ ⟦ nf₁ ⟧⇓ ρ ∙ ⟦ nf₂ ⟧⇓ ρ ∎
+
+-- The normaliser preserves the semantics of the expression.
+
+normalise-correct :
+ ∀ {n} (e : Expr n) (ρ : Env n) → ⟦ normalise e ⟧⇓ ρ ≈ ⟦ e ⟧ ρ
+normalise-correct (var x) ρ = begin
+ lookup x ρ ∙ ε ≈⟨ proj₂ identity _ ⟩
+ lookup x ρ ∎
+normalise-correct id ρ = begin
+ ε ∎
+normalise-correct (e₁ ⊕ e₂) ρ = begin
+ ⟦ normalise e₁ ++ normalise e₂ ⟧⇓ ρ ≈⟨ homomorphic (normalise e₁) (normalise e₂) ρ ⟩
+ ⟦ normalise e₁ ⟧⇓ ρ ∙ ⟦ normalise e₂ ⟧⇓ ρ ≈⟨ ∙-cong (normalise-correct e₁ ρ) (normalise-correct e₂ ρ) ⟩
+ ⟦ e₁ ⟧ ρ ∙ ⟦ e₂ ⟧ ρ ∎
+
+------------------------------------------------------------------------
+-- "Tactics"
+
+open module R = Relation.Binary.Reflection
+ setoid var ⟦_⟧ (⟦_⟧⇓ ∘ normalise) normalise-correct
+ public using (solve; _⊜_)
+
+-- We can decide if two normal forms are /syntactically/ equal.
+
+infix 5 _≟_
+
+_≟_ : ∀ {n} (nf₁ nf₂ : Normal n) → Dec (nf₁ ≡ nf₂)
+nf₁ ≟ nf₂ = Dec.map′ Rel≡⇒≡ ≡⇒Rel≡ (decidable Fin._≟_ nf₁ nf₂)
+ where open Pointwise
+
+-- We can also give a sound, but not necessarily complete, procedure
+-- for determining if two expressions have the same semantics.
+
+prove′ : ∀ {n} (e₁ e₂ : Expr n) → Maybe (∀ ρ → ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ)
+prove′ e₁ e₂ =
+ Maybe.map lemma $ decToMaybe (normalise e₁ ≟ normalise e₂)
+ where
+ lemma : normalise e₁ ≡ normalise e₂ → ∀ ρ → ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ
+ lemma eq ρ =
+ R.prove ρ e₁ e₂ (begin
+ ⟦ normalise e₁ ⟧⇓ ρ ≡⟨ P.cong (λ e → ⟦ e ⟧⇓ ρ) eq ⟩
+ ⟦ normalise e₂ ⟧⇓ ρ ∎)
+
+-- This procedure can be combined with from-just.
+
+prove : ∀ n (es : Expr n × Expr n) →
+ From-just (∀ ρ → ⟦ proj₁ es ⟧ ρ ≈ ⟦ proj₂ es ⟧ ρ)
+ (uncurry prove′ es)
+prove _ = from-just ∘ uncurry prove′
diff --git a/src/Algebra/Morphism.agda b/src/Algebra/Morphism.agda
index 6da18bb..b17087e 100644
--- a/src/Algebra/Morphism.agda
+++ b/src/Algebra/Morphism.agda
@@ -9,7 +9,7 @@ module Algebra.Morphism where
open import Relation.Binary
open import Algebra
open import Algebra.FunctionProperties
-import Algebra.Props.Group as GroupP
+import Algebra.Properties.Group as GroupP
open import Function
open import Data.Product
open import Level
diff --git a/src/Algebra/Properties/AbelianGroup.agda b/src/Algebra/Properties/AbelianGroup.agda
new file mode 100644
index 0000000..244cf76
--- /dev/null
+++ b/src/Algebra/Properties/AbelianGroup.agda
@@ -0,0 +1,44 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Some derivable properties
+------------------------------------------------------------------------
+
+open import Algebra
+
+module Algebra.Properties.AbelianGroup
+ {g₁ g₂} (G : AbelianGroup g₁ g₂) where
+
+import Algebra.Properties.Group as GP
+open import Data.Product
+open import Function
+import Relation.Binary.EqReasoning as EqR
+
+open AbelianGroup G
+open EqR setoid
+
+open GP group public
+
+private
+ lemma : ∀ x y → x ∙ y ∙ x ⁻¹ ≈ y
+ lemma x y = begin
+ x ∙ y ∙ x ⁻¹ ≈⟨ comm _ _ ⟨ ∙-cong ⟩ refl ⟩
+ y ∙ x ∙ x ⁻¹ ≈⟨ assoc _ _ _ ⟩
+ y ∙ (x ∙ x ⁻¹) ≈⟨ refl ⟨ ∙-cong ⟩ proj₂ inverse _ ⟩
+ y ∙ ε ≈⟨ proj₂ identity _ ⟩
+ y ∎
+
+⁻¹-∙-comm : ∀ x y → x ⁻¹ ∙ y ⁻¹ ≈ (x ∙ y) ⁻¹
+⁻¹-∙-comm x y = begin
+ x ⁻¹ ∙ y ⁻¹ ≈⟨ comm _ _ ⟩
+ y ⁻¹ ∙ x ⁻¹ ≈⟨ sym $ lem ⟨ ∙-cong ⟩ refl ⟩
+ x ∙ (y ∙ (x ∙ y) ⁻¹ ∙ y ⁻¹) ∙ x ⁻¹ ≈⟨ lemma _ _ ⟩
+ y ∙ (x ∙ y) ⁻¹ ∙ y ⁻¹ ≈⟨ lemma _ _ ⟩
+ (x ∙ y) ⁻¹ ∎
+ where
+ lem = begin
+ x ∙ (y ∙ (x ∙ y) ⁻¹ ∙ y ⁻¹) ≈⟨ sym $ assoc _ _ _ ⟩
+ x ∙ (y ∙ (x ∙ y) ⁻¹) ∙ y ⁻¹ ≈⟨ sym $ assoc _ _ _ ⟨ ∙-cong ⟩ refl ⟩
+ x ∙ y ∙ (x ∙ y) ⁻¹ ∙ y ⁻¹ ≈⟨ proj₂ inverse _ ⟨ ∙-cong ⟩ refl ⟩
+ ε ∙ y ⁻¹ ≈⟨ proj₁ identity _ ⟩
+ y ⁻¹ ∎
diff --git a/src/Algebra/Properties/BooleanAlgebra.agda b/src/Algebra/Properties/BooleanAlgebra.agda
new file mode 100644
index 0000000..148eb72
--- /dev/null
+++ b/src/Algebra/Properties/BooleanAlgebra.agda
@@ -0,0 +1,570 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Some derivable properties
+------------------------------------------------------------------------
+
+open import Algebra
+
+module Algebra.Properties.BooleanAlgebra
+ {b₁ b₂} (B : BooleanAlgebra b₁ b₂)
+ where
+
+open BooleanAlgebra B
+import Algebra.Properties.DistributiveLattice
+private
+ open module DL = Algebra.Properties.DistributiveLattice
+ distributiveLattice public
+ hiding (replace-equality)
+open import Algebra.Structures
+import Algebra.FunctionProperties as P; open P _≈_
+import Relation.Binary.EqReasoning as EqR; open EqR setoid
+open import Relation.Binary
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence using (_⇔_; module Equivalence)
+open import Data.Product
+
+------------------------------------------------------------------------
+-- Some simple generalisations
+
+∨-complement : Inverse ⊤ ¬_ _∨_
+∨-complement = ∨-complementˡ , ∨-complementʳ
+ where
+ ∨-complementˡ : LeftInverse ⊤ ¬_ _∨_
+ ∨-complementˡ x = begin
+ ¬ x ∨ x ≈⟨ ∨-comm _ _ ⟩
+ x ∨ ¬ x ≈⟨ ∨-complementʳ _ ⟩
+ ⊤ ∎
+
+∧-complement : Inverse ⊥ ¬_ _∧_
+∧-complement = ∧-complementˡ , ∧-complementʳ
+ where
+ ∧-complementˡ : LeftInverse ⊥ ¬_ _∧_
+ ∧-complementˡ x = begin
+ ¬ x ∧ x ≈⟨ ∧-comm _ _ ⟩
+ x ∧ ¬ x ≈⟨ ∧-complementʳ _ ⟩
+ ⊥ ∎
+
+------------------------------------------------------------------------
+-- The dual construction is also a boolean algebra
+
+∧-∨-isBooleanAlgebra : IsBooleanAlgebra _≈_ _∧_ _∨_ ¬_ ⊥ ⊤
+∧-∨-isBooleanAlgebra = record
+ { isDistributiveLattice = ∧-∨-isDistributiveLattice
+ ; ∨-complementʳ = proj₂ ∧-complement
+ ; ∧-complementʳ = proj₂ ∨-complement
+ ; ¬-cong = ¬-cong
+ }
+
+∧-∨-booleanAlgebra : BooleanAlgebra _ _
+∧-∨-booleanAlgebra = record
+ { _∧_ = _∨_
+ ; _∨_ = _∧_
+ ; ⊤ = ⊥
+ ; ⊥ = ⊤
+ ; isBooleanAlgebra = ∧-∨-isBooleanAlgebra
+ }
+
+------------------------------------------------------------------------
+-- (∨, ∧, ⊥, ⊤) is a commutative semiring
+
+private
+
+ ∧-identity : Identity ⊤ _∧_
+ ∧-identity = (λ _ → ∧-comm _ _ ⟨ trans ⟩ x∧⊤=x _) , x∧⊤=x
+ where
+ x∧⊤=x : ∀ x → x ∧ ⊤ ≈ x
+ x∧⊤=x x = begin
+ x ∧ ⊤ ≈⟨ refl ⟨ ∧-cong ⟩ sym (proj₂ ∨-complement _) ⟩
+ x ∧ (x ∨ ¬ x) ≈⟨ proj₂ absorptive _ _ ⟩
+ x ∎
+
+ ∨-identity : Identity ⊥ _∨_
+ ∨-identity = (λ _ → ∨-comm _ _ ⟨ trans ⟩ x∨⊥=x _) , x∨⊥=x
+ where
+ x∨⊥=x : ∀ x → x ∨ ⊥ ≈ x
+ x∨⊥=x x = begin
+ x ∨ ⊥ ≈⟨ refl ⟨ ∨-cong ⟩ sym (proj₂ ∧-complement _) ⟩
+ x ∨ x ∧ ¬ x ≈⟨ proj₁ absorptive _ _ ⟩
+ x ∎
+
+ ∧-zero : Zero ⊥ _∧_
+ ∧-zero = (λ _ → ∧-comm _ _ ⟨ trans ⟩ x∧⊥=⊥ _) , x∧⊥=⊥
+ where
+ x∧⊥=⊥ : ∀ x → x ∧ ⊥ ≈ ⊥
+ x∧⊥=⊥ x = begin
+ x ∧ ⊥ ≈⟨ refl ⟨ ∧-cong ⟩ sym (proj₂ ∧-complement _) ⟩
+ x ∧ x ∧ ¬ x ≈⟨ sym $ ∧-assoc _ _ _ ⟩
+ (x ∧ x) ∧ ¬ x ≈⟨ ∧-idempotent _ ⟨ ∧-cong ⟩ refl ⟩
+ x ∧ ¬ x ≈⟨ proj₂ ∧-complement _ ⟩
+ ⊥ ∎
+
+∨-∧-isCommutativeSemiring : IsCommutativeSemiring _≈_ _∨_ _∧_ ⊥ ⊤
+∨-∧-isCommutativeSemiring = record
+ { +-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∨-assoc
+ ; ∙-cong = ∨-cong
+ }
+ ; identityˡ = proj₁ ∨-identity
+ ; comm = ∨-comm
+ }
+ ; *-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∧-assoc
+ ; ∙-cong = ∧-cong
+ }
+ ; identityˡ = proj₁ ∧-identity
+ ; comm = ∧-comm
+ }
+ ; distribʳ = proj₂ ∧-∨-distrib
+ ; zeroˡ = proj₁ ∧-zero
+ }
+
+∨-∧-commutativeSemiring : CommutativeSemiring _ _
+∨-∧-commutativeSemiring = record
+ { _+_ = _∨_
+ ; _*_ = _∧_
+ ; 0# = ⊥
+ ; 1# = ⊤
+ ; isCommutativeSemiring = ∨-∧-isCommutativeSemiring
+ }
+
+------------------------------------------------------------------------
+-- (∧, ∨, ⊤, ⊥) is a commutative semiring
+
+private
+
+ ∨-zero : Zero ⊤ _∨_
+ ∨-zero = (λ _ → ∨-comm _ _ ⟨ trans ⟩ x∨⊤=⊤ _) , x∨⊤=⊤
+ where
+ x∨⊤=⊤ : ∀ x → x ∨ ⊤ ≈ ⊤
+ x∨⊤=⊤ x = begin
+ x ∨ ⊤ ≈⟨ refl ⟨ ∨-cong ⟩ sym (proj₂ ∨-complement _) ⟩
+ x ∨ x ∨ ¬ x ≈⟨ sym $ ∨-assoc _ _ _ ⟩
+ (x ∨ x) ∨ ¬ x ≈⟨ ∨-idempotent _ ⟨ ∨-cong ⟩ refl ⟩
+ x ∨ ¬ x ≈⟨ proj₂ ∨-complement _ ⟩
+ ⊤ ∎
+
+∧-∨-isCommutativeSemiring : IsCommutativeSemiring _≈_ _∧_ _∨_ ⊤ ⊥
+∧-∨-isCommutativeSemiring = record
+ { +-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∧-assoc
+ ; ∙-cong = ∧-cong
+ }
+ ; identityˡ = proj₁ ∧-identity
+ ; comm = ∧-comm
+ }
+ ; *-isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∨-assoc
+ ; ∙-cong = ∨-cong
+ }
+ ; identityˡ = proj₁ ∨-identity
+ ; comm = ∨-comm
+ }
+ ; distribʳ = proj₂ ∨-∧-distrib
+ ; zeroˡ = proj₁ ∨-zero
+ }
+
+∧-∨-commutativeSemiring : CommutativeSemiring _ _
+∧-∨-commutativeSemiring =
+ record { isCommutativeSemiring = ∧-∨-isCommutativeSemiring }
+
+------------------------------------------------------------------------
+-- Some other properties
+
+-- I took the statement of this lemma (called Uniqueness of
+-- Complements) from some course notes, "Boolean Algebra", written
+-- by Gert Smolka.
+
+private
+ lemma : ∀ x y → x ∧ y ≈ ⊥ → x ∨ y ≈ ⊤ → ¬ x ≈ y
+ lemma x y x∧y=⊥ x∨y=⊤ = begin
+ ¬ x ≈⟨ sym $ proj₂ ∧-identity _ ⟩
+ ¬ x ∧ ⊤ ≈⟨ refl ⟨ ∧-cong ⟩ sym x∨y=⊤ ⟩
+ ¬ x ∧ (x ∨ y) ≈⟨ proj₁ ∧-∨-distrib _ _ _ ⟩
+ ¬ x ∧ x ∨ ¬ x ∧ y ≈⟨ proj₁ ∧-complement _ ⟨ ∨-cong ⟩ refl ⟩
+ ⊥ ∨ ¬ x ∧ y ≈⟨ sym x∧y=⊥ ⟨ ∨-cong ⟩ refl ⟩
+ x ∧ y ∨ ¬ x ∧ y ≈⟨ sym $ proj₂ ∧-∨-distrib _ _ _ ⟩
+ (x ∨ ¬ x) ∧ y ≈⟨ proj₂ ∨-complement _ ⟨ ∧-cong ⟩ refl ⟩
+ ⊤ ∧ y ≈⟨ proj₁ ∧-identity _ ⟩
+ y ∎
+
+¬⊥=⊤ : ¬ ⊥ ≈ ⊤
+¬⊥=⊤ = lemma ⊥ ⊤ (proj₂ ∧-identity _) (proj₂ ∨-zero _)
+
+¬⊤=⊥ : ¬ ⊤ ≈ ⊥
+¬⊤=⊥ = lemma ⊤ ⊥ (proj₂ ∧-zero _) (proj₂ ∨-identity _)
+
+¬-involutive : Involutive ¬_
+¬-involutive x = lemma (¬ x) x (proj₁ ∧-complement _) (proj₁ ∨-complement _)
+
+deMorgan₁ : ∀ x y → ¬ (x ∧ y) ≈ ¬ x ∨ ¬ y
+deMorgan₁ x y = lemma (x ∧ y) (¬ x ∨ ¬ y) lem₁ lem₂
+ where
+ lem₁ = begin
+ (x ∧ y) ∧ (¬ x ∨ ¬ y) ≈⟨ proj₁ ∧-∨-distrib _ _ _ ⟩
+ (x ∧ y) ∧ ¬ x ∨ (x ∧ y) ∧ ¬ y ≈⟨ (∧-comm _ _ ⟨ ∧-cong ⟩ refl) ⟨ ∨-cong ⟩ refl ⟩
+ (y ∧ x) ∧ ¬ x ∨ (x ∧ y) ∧ ¬ y ≈⟨ ∧-assoc _ _ _ ⟨ ∨-cong ⟩ ∧-assoc _ _ _ ⟩
+ y ∧ (x ∧ ¬ x) ∨ x ∧ (y ∧ ¬ y) ≈⟨ (refl ⟨ ∧-cong ⟩ proj₂ ∧-complement _) ⟨ ∨-cong ⟩
+ (refl ⟨ ∧-cong ⟩ proj₂ ∧-complement _) ⟩
+ (y ∧ ⊥) ∨ (x ∧ ⊥) ≈⟨ proj₂ ∧-zero _ ⟨ ∨-cong ⟩ proj₂ ∧-zero _ ⟩
+ ⊥ ∨ ⊥ ≈⟨ proj₂ ∨-identity _ ⟩
+ ⊥ ∎
+
+ lem₃ = begin
+ (x ∧ y) ∨ ¬ x ≈⟨ proj₂ ∨-∧-distrib _ _ _ ⟩
+ (x ∨ ¬ x) ∧ (y ∨ ¬ x) ≈⟨ proj₂ ∨-complement _ ⟨ ∧-cong ⟩ refl ⟩
+ ⊤ ∧ (y ∨ ¬ x) ≈⟨ proj₁ ∧-identity _ ⟩
+ y ∨ ¬ x ≈⟨ ∨-comm _ _ ⟩
+ ¬ x ∨ y ∎
+
+ lem₂ = begin
+ (x ∧ y) ∨ (¬ x ∨ ¬ y) ≈⟨ sym $ ∨-assoc _ _ _ ⟩
+ ((x ∧ y) ∨ ¬ x) ∨ ¬ y ≈⟨ lem₃ ⟨ ∨-cong ⟩ refl ⟩
+ (¬ x ∨ y) ∨ ¬ y ≈⟨ ∨-assoc _ _ _ ⟩
+ ¬ x ∨ (y ∨ ¬ y) ≈⟨ refl ⟨ ∨-cong ⟩ proj₂ ∨-complement _ ⟩
+ ¬ x ∨ ⊤ ≈⟨ proj₂ ∨-zero _ ⟩
+ ⊤ ∎
+
+deMorgan₂ : ∀ x y → ¬ (x ∨ y) ≈ ¬ x ∧ ¬ y
+deMorgan₂ x y = begin
+ ¬ (x ∨ y) ≈⟨ ¬-cong $ sym (¬-involutive _) ⟨ ∨-cong ⟩
+ sym (¬-involutive _) ⟩
+ ¬ (¬ ¬ x ∨ ¬ ¬ y) ≈⟨ ¬-cong $ sym $ deMorgan₁ _ _ ⟩
+ ¬ ¬ (¬ x ∧ ¬ y) ≈⟨ ¬-involutive _ ⟩
+ ¬ x ∧ ¬ y ∎
+
+-- One can replace the underlying equality with an equivalent one.
+
+replace-equality :
+ {_≈′_ : Rel Carrier b₂} →
+ (∀ {x y} → x ≈ y ⇔ x ≈′ y) → BooleanAlgebra _ _
+replace-equality {_≈′_} ≈⇔≈′ = record
+ { _≈_ = _≈′_
+ ; _∨_ = _∨_
+ ; _∧_ = _∧_
+ ; ¬_ = ¬_
+ ; ⊤ = ⊤
+ ; ⊥ = ⊥
+ ; isBooleanAlgebra = record
+ { isDistributiveLattice = DistributiveLattice.isDistributiveLattice
+ (DL.replace-equality ≈⇔≈′)
+ ; ∨-complementʳ = λ x → to ⟨$⟩ ∨-complementʳ x
+ ; ∧-complementʳ = λ x → to ⟨$⟩ ∧-complementʳ x
+ ; ¬-cong = λ i≈j → to ⟨$⟩ ¬-cong (from ⟨$⟩ i≈j)
+ }
+ } where open module E {x y} = Equivalence (≈⇔≈′ {x} {y})
+
+------------------------------------------------------------------------
+-- (⊕, ∧, id, ⊥, ⊤) is a commutative ring
+
+-- This construction is parameterised over the definition of xor.
+
+module XorRing
+ (xor : Op₂ Carrier)
+ (⊕-def : ∀ x y → xor x y ≈ (x ∨ y) ∧ ¬ (x ∧ y))
+ where
+
+ private
+ infixl 6 _⊕_
+
+ _⊕_ : Op₂ Carrier
+ _⊕_ = xor
+
+ private
+ helper : ∀ {x y u v} → x ≈ y → u ≈ v → x ∧ ¬ u ≈ y ∧ ¬ v
+ helper x≈y u≈v = x≈y ⟨ ∧-cong ⟩ ¬-cong u≈v
+
+ ⊕-¬-distribˡ : ∀ x y → ¬ (x ⊕ y) ≈ ¬ x ⊕ y
+ ⊕-¬-distribˡ x y = begin
+ ¬ (x ⊕ y) ≈⟨ ¬-cong $ ⊕-def _ _ ⟩
+ ¬ ((x ∨ y) ∧ (¬ (x ∧ y))) ≈⟨ ¬-cong (proj₂ ∧-∨-distrib _ _ _) ⟩
+ ¬ ((x ∧ ¬ (x ∧ y)) ∨ (y ∧ ¬ (x ∧ y))) ≈⟨ ¬-cong $
+ refl ⟨ ∨-cong ⟩
+ (refl ⟨ ∧-cong ⟩
+ ¬-cong (∧-comm _ _)) ⟩
+ ¬ ((x ∧ ¬ (x ∧ y)) ∨ (y ∧ ¬ (y ∧ x))) ≈⟨ ¬-cong $ lem _ _ ⟨ ∨-cong ⟩ lem _ _ ⟩
+ ¬ ((x ∧ ¬ y) ∨ (y ∧ ¬ x)) ≈⟨ deMorgan₂ _ _ ⟩
+ ¬ (x ∧ ¬ y) ∧ ¬ (y ∧ ¬ x) ≈⟨ deMorgan₁ _ _ ⟨ ∧-cong ⟩ refl ⟩
+ (¬ x ∨ (¬ ¬ y)) ∧ ¬ (y ∧ ¬ x) ≈⟨ helper (refl ⟨ ∨-cong ⟩ ¬-involutive _)
+ (∧-comm _ _) ⟩
+ (¬ x ∨ y) ∧ ¬ (¬ x ∧ y) ≈⟨ sym $ ⊕-def _ _ ⟩
+ ¬ x ⊕ y ∎
+ where
+ lem : ∀ x y → x ∧ ¬ (x ∧ y) ≈ x ∧ ¬ y
+ lem x y = begin
+ x ∧ ¬ (x ∧ y) ≈⟨ refl ⟨ ∧-cong ⟩ deMorgan₁ _ _ ⟩
+ x ∧ (¬ x ∨ ¬ y) ≈⟨ proj₁ ∧-∨-distrib _ _ _ ⟩
+ (x ∧ ¬ x) ∨ (x ∧ ¬ y) ≈⟨ proj₂ ∧-complement _ ⟨ ∨-cong ⟩ refl ⟩
+ ⊥ ∨ (x ∧ ¬ y) ≈⟨ proj₁ ∨-identity _ ⟩
+ x ∧ ¬ y ∎
+
+ private
+ ⊕-comm : Commutative _⊕_
+ ⊕-comm x y = begin
+ x ⊕ y ≈⟨ ⊕-def _ _ ⟩
+ (x ∨ y) ∧ ¬ (x ∧ y) ≈⟨ helper (∨-comm _ _) (∧-comm _ _) ⟩
+ (y ∨ x) ∧ ¬ (y ∧ x) ≈⟨ sym $ ⊕-def _ _ ⟩
+ y ⊕ x ∎
+
+ ⊕-¬-distribʳ : ∀ x y → ¬ (x ⊕ y) ≈ x ⊕ ¬ y
+ ⊕-¬-distribʳ x y = begin
+ ¬ (x ⊕ y) ≈⟨ ¬-cong $ ⊕-comm _ _ ⟩
+ ¬ (y ⊕ x) ≈⟨ ⊕-¬-distribˡ _ _ ⟩
+ ¬ y ⊕ x ≈⟨ ⊕-comm _ _ ⟩
+ x ⊕ ¬ y ∎
+
+ ⊕-annihilates-¬ : ∀ x y → x ⊕ y ≈ ¬ x ⊕ ¬ y
+ ⊕-annihilates-¬ x y = begin
+ x ⊕ y ≈⟨ sym $ ¬-involutive _ ⟩
+ ¬ ¬ (x ⊕ y) ≈⟨ ¬-cong $ ⊕-¬-distribˡ _ _ ⟩
+ ¬ (¬ x ⊕ y) ≈⟨ ⊕-¬-distribʳ _ _ ⟩
+ ¬ x ⊕ ¬ y ∎
+
+ private
+ ⊕-cong : _⊕_ Preserves₂ _≈_ ⟶ _≈_ ⟶ _≈_
+ ⊕-cong {x} {y} {u} {v} x≈y u≈v = begin
+ x ⊕ u ≈⟨ ⊕-def _ _ ⟩
+ (x ∨ u) ∧ ¬ (x ∧ u) ≈⟨ helper (x≈y ⟨ ∨-cong ⟩ u≈v)
+ (x≈y ⟨ ∧-cong ⟩ u≈v) ⟩
+ (y ∨ v) ∧ ¬ (y ∧ v) ≈⟨ sym $ ⊕-def _ _ ⟩
+ y ⊕ v ∎
+
+ ⊕-identity : Identity ⊥ _⊕_
+ ⊕-identity = ⊥⊕x=x , (λ _ → ⊕-comm _ _ ⟨ trans ⟩ ⊥⊕x=x _)
+ where
+ ⊥⊕x=x : ∀ x → ⊥ ⊕ x ≈ x
+ ⊥⊕x=x x = begin
+ ⊥ ⊕ x ≈⟨ ⊕-def _ _ ⟩
+ (⊥ ∨ x) ∧ ¬ (⊥ ∧ x) ≈⟨ helper (proj₁ ∨-identity _)
+ (proj₁ ∧-zero _) ⟩
+ x ∧ ¬ ⊥ ≈⟨ refl ⟨ ∧-cong ⟩ ¬⊥=⊤ ⟩
+ x ∧ ⊤ ≈⟨ proj₂ ∧-identity _ ⟩
+ x ∎
+
+ ⊕-inverse : Inverse ⊥ id _⊕_
+ ⊕-inverse = x⊕x=⊥ , (λ _ → ⊕-comm _ _ ⟨ trans ⟩ x⊕x=⊥ _)
+ where
+ x⊕x=⊥ : ∀ x → x ⊕ x ≈ ⊥
+ x⊕x=⊥ x = begin
+ x ⊕ x ≈⟨ ⊕-def _ _ ⟩
+ (x ∨ x) ∧ ¬ (x ∧ x) ≈⟨ helper (∨-idempotent _)
+ (∧-idempotent _) ⟩
+ x ∧ ¬ x ≈⟨ proj₂ ∧-complement _ ⟩
+ ⊥ ∎
+
+ distrib-∧-⊕ : _∧_ DistributesOver _⊕_
+ distrib-∧-⊕ = distˡ , distʳ
+ where
+ distˡ : _∧_ DistributesOverˡ _⊕_
+ distˡ x y z = begin
+ x ∧ (y ⊕ z) ≈⟨ refl ⟨ ∧-cong ⟩ ⊕-def _ _ ⟩
+ x ∧ ((y ∨ z) ∧ ¬ (y ∧ z)) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
+ (x ∧ (y ∨ z)) ∧ ¬ (y ∧ z) ≈⟨ refl ⟨ ∧-cong ⟩ deMorgan₁ _ _ ⟩
+ (x ∧ (y ∨ z)) ∧
+ (¬ y ∨ ¬ z) ≈⟨ sym $ proj₁ ∨-identity _ ⟩
+ ⊥ ∨
+ ((x ∧ (y ∨ z)) ∧
+ (¬ y ∨ ¬ z)) ≈⟨ lem₃ ⟨ ∨-cong ⟩ refl ⟩
+ ((x ∧ (y ∨ z)) ∧ ¬ x) ∨
+ ((x ∧ (y ∨ z)) ∧
+ (¬ y ∨ ¬ z)) ≈⟨ sym $ proj₁ ∧-∨-distrib _ _ _ ⟩
+ (x ∧ (y ∨ z)) ∧
+ (¬ x ∨ (¬ y ∨ ¬ z)) ≈⟨ refl ⟨ ∧-cong ⟩
+ (refl ⟨ ∨-cong ⟩ sym (deMorgan₁ _ _)) ⟩
+ (x ∧ (y ∨ z)) ∧
+ (¬ x ∨ ¬ (y ∧ z)) ≈⟨ refl ⟨ ∧-cong ⟩ sym (deMorgan₁ _ _) ⟩
+ (x ∧ (y ∨ z)) ∧
+ ¬ (x ∧ (y ∧ z)) ≈⟨ helper refl lem₁ ⟩
+ (x ∧ (y ∨ z)) ∧
+ ¬ ((x ∧ y) ∧ (x ∧ z)) ≈⟨ proj₁ ∧-∨-distrib _ _ _ ⟨ ∧-cong ⟩
+ refl ⟩
+ ((x ∧ y) ∨ (x ∧ z)) ∧
+ ¬ ((x ∧ y) ∧ (x ∧ z)) ≈⟨ sym $ ⊕-def _ _ ⟩
+ (x ∧ y) ⊕ (x ∧ z) ∎
+ where
+ lem₂ = begin
+ x ∧ (y ∧ z) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
+ (x ∧ y) ∧ z ≈⟨ ∧-comm _ _ ⟨ ∧-cong ⟩ refl ⟩
+ (y ∧ x) ∧ z ≈⟨ ∧-assoc _ _ _ ⟩
+ y ∧ (x ∧ z) ∎
+
+ lem₁ = begin
+ x ∧ (y ∧ z) ≈⟨ sym (∧-idempotent _) ⟨ ∧-cong ⟩ refl ⟩
+ (x ∧ x) ∧ (y ∧ z) ≈⟨ ∧-assoc _ _ _ ⟩
+ x ∧ (x ∧ (y ∧ z)) ≈⟨ refl ⟨ ∧-cong ⟩ lem₂ ⟩
+ x ∧ (y ∧ (x ∧ z)) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
+ (x ∧ y) ∧ (x ∧ z) ∎
+
+ lem₃ = begin
+ ⊥ ≈⟨ sym $ proj₂ ∧-zero _ ⟩
+ (y ∨ z) ∧ ⊥ ≈⟨ refl ⟨ ∧-cong ⟩ sym (proj₂ ∧-complement _) ⟩
+ (y ∨ z) ∧ (x ∧ ¬ x) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
+ ((y ∨ z) ∧ x) ∧ ¬ x ≈⟨ ∧-comm _ _ ⟨ ∧-cong ⟩ refl ⟩
+ (x ∧ (y ∨ z)) ∧ ¬ x ∎
+
+ distʳ : _∧_ DistributesOverʳ _⊕_
+ distʳ x y z = begin
+ (y ⊕ z) ∧ x ≈⟨ ∧-comm _ _ ⟩
+ x ∧ (y ⊕ z) ≈⟨ distˡ _ _ _ ⟩
+ (x ∧ y) ⊕ (x ∧ z) ≈⟨ ∧-comm _ _ ⟨ ⊕-cong ⟩ ∧-comm _ _ ⟩
+ (y ∧ x) ⊕ (z ∧ x) ∎
+
+ lemma₂ : ∀ x y u v →
+ (x ∧ y) ∨ (u ∧ v) ≈
+ ((x ∨ u) ∧ (y ∨ u)) ∧
+ ((x ∨ v) ∧ (y ∨ v))
+ lemma₂ x y u v = begin
+ (x ∧ y) ∨ (u ∧ v) ≈⟨ proj₁ ∨-∧-distrib _ _ _ ⟩
+ ((x ∧ y) ∨ u) ∧ ((x ∧ y) ∨ v) ≈⟨ proj₂ ∨-∧-distrib _ _ _
+ ⟨ ∧-cong ⟩
+ proj₂ ∨-∧-distrib _ _ _ ⟩
+ ((x ∨ u) ∧ (y ∨ u)) ∧
+ ((x ∨ v) ∧ (y ∨ v)) ∎
+
+ ⊕-assoc : Associative _⊕_
+ ⊕-assoc x y z = sym $ begin
+ x ⊕ (y ⊕ z) ≈⟨ refl ⟨ ⊕-cong ⟩ ⊕-def _ _ ⟩
+ x ⊕ ((y ∨ z) ∧ ¬ (y ∧ z)) ≈⟨ ⊕-def _ _ ⟩
+ (x ∨ ((y ∨ z) ∧ ¬ (y ∧ z))) ∧
+ ¬ (x ∧ ((y ∨ z) ∧ ¬ (y ∧ z))) ≈⟨ lem₃ ⟨ ∧-cong ⟩ lem₄ ⟩
+ (((x ∨ y) ∨ z) ∧ ((x ∨ ¬ y) ∨ ¬ z)) ∧
+ (((¬ x ∨ ¬ y) ∨ z) ∧ ((¬ x ∨ y) ∨ ¬ z)) ≈⟨ ∧-assoc _ _ _ ⟩
+ ((x ∨ y) ∨ z) ∧
+ (((x ∨ ¬ y) ∨ ¬ z) ∧
+ (((¬ x ∨ ¬ y) ∨ z) ∧ ((¬ x ∨ y) ∨ ¬ z))) ≈⟨ refl ⟨ ∧-cong ⟩ lem₅ ⟩
+ ((x ∨ y) ∨ z) ∧
+ (((¬ x ∨ ¬ y) ∨ z) ∧
+ (((x ∨ ¬ y) ∨ ¬ z) ∧ ((¬ x ∨ y) ∨ ¬ z))) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
+ (((x ∨ y) ∨ z) ∧ ((¬ x ∨ ¬ y) ∨ z)) ∧
+ (((x ∨ ¬ y) ∨ ¬ z) ∧ ((¬ x ∨ y) ∨ ¬ z)) ≈⟨ lem₁ ⟨ ∧-cong ⟩ lem₂ ⟩
+ (((x ∨ y) ∧ ¬ (x ∧ y)) ∨ z) ∧
+ ¬ (((x ∨ y) ∧ ¬ (x ∧ y)) ∧ z) ≈⟨ sym $ ⊕-def _ _ ⟩
+ ((x ∨ y) ∧ ¬ (x ∧ y)) ⊕ z ≈⟨ sym $ ⊕-def _ _ ⟨ ⊕-cong ⟩ refl ⟩
+ (x ⊕ y) ⊕ z ∎
+ where
+ lem₁ = begin
+ ((x ∨ y) ∨ z) ∧ ((¬ x ∨ ¬ y) ∨ z) ≈⟨ sym $ proj₂ ∨-∧-distrib _ _ _ ⟩
+ ((x ∨ y) ∧ (¬ x ∨ ¬ y)) ∨ z ≈⟨ (refl ⟨ ∧-cong ⟩ sym (deMorgan₁ _ _))
+ ⟨ ∨-cong ⟩ refl ⟩
+ ((x ∨ y) ∧ ¬ (x ∧ y)) ∨ z ∎
+
+ lem₂' = begin
+ (x ∨ ¬ y) ∧ (¬ x ∨ y) ≈⟨ sym $
+ proj₁ ∧-identity _
+ ⟨ ∧-cong ⟩
+ proj₂ ∧-identity _ ⟩
+ (⊤ ∧ (x ∨ ¬ y)) ∧ ((¬ x ∨ y) ∧ ⊤) ≈⟨ sym $
+ (proj₁ ∨-complement _ ⟨ ∧-cong ⟩ ∨-comm _ _)
+ ⟨ ∧-cong ⟩
+ (refl ⟨ ∧-cong ⟩ proj₁ ∨-complement _) ⟩
+ ((¬ x ∨ x) ∧ (¬ y ∨ x)) ∧
+ ((¬ x ∨ y) ∧ (¬ y ∨ y)) ≈⟨ sym $ lemma₂ _ _ _ _ ⟩
+ (¬ x ∧ ¬ y) ∨ (x ∧ y) ≈⟨ sym $ deMorgan₂ _ _ ⟨ ∨-cong ⟩ ¬-involutive _ ⟩
+ ¬ (x ∨ y) ∨ ¬ ¬ (x ∧ y) ≈⟨ sym (deMorgan₁ _ _) ⟩
+ ¬ ((x ∨ y) ∧ ¬ (x ∧ y)) ∎
+
+ lem₂ = begin
+ ((x ∨ ¬ y) ∨ ¬ z) ∧ ((¬ x ∨ y) ∨ ¬ z) ≈⟨ sym $ proj₂ ∨-∧-distrib _ _ _ ⟩
+ ((x ∨ ¬ y) ∧ (¬ x ∨ y)) ∨ ¬ z ≈⟨ lem₂' ⟨ ∨-cong ⟩ refl ⟩
+ ¬ ((x ∨ y) ∧ ¬ (x ∧ y)) ∨ ¬ z ≈⟨ sym $ deMorgan₁ _ _ ⟩
+ ¬ (((x ∨ y) ∧ ¬ (x ∧ y)) ∧ z) ∎
+
+ lem₃ = begin
+ x ∨ ((y ∨ z) ∧ ¬ (y ∧ z)) ≈⟨ refl ⟨ ∨-cong ⟩
+ (refl ⟨ ∧-cong ⟩ deMorgan₁ _ _) ⟩
+ x ∨ ((y ∨ z) ∧ (¬ y ∨ ¬ z)) ≈⟨ proj₁ ∨-∧-distrib _ _ _ ⟩
+ (x ∨ (y ∨ z)) ∧ (x ∨ (¬ y ∨ ¬ z)) ≈⟨ sym (∨-assoc _ _ _) ⟨ ∧-cong ⟩
+ sym (∨-assoc _ _ _) ⟩
+ ((x ∨ y) ∨ z) ∧ ((x ∨ ¬ y) ∨ ¬ z) ∎
+
+ lem₄' = begin
+ ¬ ((y ∨ z) ∧ ¬ (y ∧ z)) ≈⟨ deMorgan₁ _ _ ⟩
+ ¬ (y ∨ z) ∨ ¬ ¬ (y ∧ z) ≈⟨ deMorgan₂ _ _ ⟨ ∨-cong ⟩ ¬-involutive _ ⟩
+ (¬ y ∧ ¬ z) ∨ (y ∧ z) ≈⟨ lemma₂ _ _ _ _ ⟩
+ ((¬ y ∨ y) ∧ (¬ z ∨ y)) ∧
+ ((¬ y ∨ z) ∧ (¬ z ∨ z)) ≈⟨ (proj₁ ∨-complement _ ⟨ ∧-cong ⟩ ∨-comm _ _)
+ ⟨ ∧-cong ⟩
+ (refl ⟨ ∧-cong ⟩ proj₁ ∨-complement _) ⟩
+ (⊤ ∧ (y ∨ ¬ z)) ∧
+ ((¬ y ∨ z) ∧ ⊤) ≈⟨ proj₁ ∧-identity _ ⟨ ∧-cong ⟩
+ proj₂ ∧-identity _ ⟩
+ (y ∨ ¬ z) ∧ (¬ y ∨ z) ∎
+
+ lem₄ = begin
+ ¬ (x ∧ ((y ∨ z) ∧ ¬ (y ∧ z))) ≈⟨ deMorgan₁ _ _ ⟩
+ ¬ x ∨ ¬ ((y ∨ z) ∧ ¬ (y ∧ z)) ≈⟨ refl ⟨ ∨-cong ⟩ lem₄' ⟩
+ ¬ x ∨ ((y ∨ ¬ z) ∧ (¬ y ∨ z)) ≈⟨ proj₁ ∨-∧-distrib _ _ _ ⟩
+ (¬ x ∨ (y ∨ ¬ z)) ∧
+ (¬ x ∨ (¬ y ∨ z)) ≈⟨ sym (∨-assoc _ _ _) ⟨ ∧-cong ⟩
+ sym (∨-assoc _ _ _) ⟩
+ ((¬ x ∨ y) ∨ ¬ z) ∧
+ ((¬ x ∨ ¬ y) ∨ z) ≈⟨ ∧-comm _ _ ⟩
+ ((¬ x ∨ ¬ y) ∨ z) ∧
+ ((¬ x ∨ y) ∨ ¬ z) ∎
+
+ lem₅ = begin
+ ((x ∨ ¬ y) ∨ ¬ z) ∧
+ (((¬ x ∨ ¬ y) ∨ z) ∧ ((¬ x ∨ y) ∨ ¬ z)) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
+ (((x ∨ ¬ y) ∨ ¬ z) ∧ ((¬ x ∨ ¬ y) ∨ z)) ∧
+ ((¬ x ∨ y) ∨ ¬ z) ≈⟨ ∧-comm _ _ ⟨ ∧-cong ⟩ refl ⟩
+ (((¬ x ∨ ¬ y) ∨ z) ∧ ((x ∨ ¬ y) ∨ ¬ z)) ∧
+ ((¬ x ∨ y) ∨ ¬ z) ≈⟨ ∧-assoc _ _ _ ⟩
+ ((¬ x ∨ ¬ y) ∨ z) ∧
+ (((x ∨ ¬ y) ∨ ¬ z) ∧ ((¬ x ∨ y) ∨ ¬ z)) ∎
+
+ isCommutativeRing : IsCommutativeRing _≈_ _⊕_ _∧_ id ⊥ ⊤
+ isCommutativeRing = record
+ { isRing = record
+ { +-isAbelianGroup = record
+ { isGroup = record
+ { isMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ⊕-assoc
+ ; ∙-cong = ⊕-cong
+ }
+ ; identity = ⊕-identity
+ }
+ ; inverse = ⊕-inverse
+ ; ⁻¹-cong = id
+ }
+ ; comm = ⊕-comm
+ }
+ ; *-isMonoid = record
+ { isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∧-assoc
+ ; ∙-cong = ∧-cong
+ }
+ ; identity = ∧-identity
+ }
+ ; distrib = distrib-∧-⊕
+ }
+ ; *-comm = ∧-comm
+ }
+
+ commutativeRing : CommutativeRing _ _
+ commutativeRing = record
+ { _+_ = _⊕_
+ ; _*_ = _∧_
+ ; -_ = id
+ ; 0# = ⊥
+ ; 1# = ⊤
+ ; isCommutativeRing = isCommutativeRing
+ }
+
+infixl 6 _⊕_
+
+_⊕_ : Op₂ Carrier
+x ⊕ y = (x ∨ y) ∧ ¬ (x ∧ y)
+
+module DefaultXorRing = XorRing _⊕_ (λ _ _ → refl)
diff --git a/src/Algebra/Props/BooleanAlgebra/Expression.agda b/src/Algebra/Properties/BooleanAlgebra/Expression.agda
index 5a3814f..a94397e 100644
--- a/src/Algebra/Props/BooleanAlgebra/Expression.agda
+++ b/src/Algebra/Properties/BooleanAlgebra/Expression.agda
@@ -6,7 +6,7 @@
open import Algebra
-module Algebra.Props.BooleanAlgebra.Expression
+module Algebra.Properties.BooleanAlgebra.Expression
{b} (B : BooleanAlgebra b b)
where
diff --git a/src/Algebra/Properties/DistributiveLattice.agda b/src/Algebra/Properties/DistributiveLattice.agda
new file mode 100644
index 0000000..1b7f614
--- /dev/null
+++ b/src/Algebra/Properties/DistributiveLattice.agda
@@ -0,0 +1,85 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Some derivable properties
+------------------------------------------------------------------------
+
+open import Algebra
+
+module Algebra.Properties.DistributiveLattice
+ {dl₁ dl₂} (DL : DistributiveLattice dl₁ dl₂)
+ where
+
+open DistributiveLattice DL
+import Algebra.Properties.Lattice
+private
+ open module L = Algebra.Properties.Lattice lattice public
+ hiding (replace-equality)
+open import Algebra.Structures
+import Algebra.FunctionProperties as P; open P _≈_
+open import Relation.Binary
+import Relation.Binary.EqReasoning as EqR; open EqR setoid
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence using (_⇔_; module Equivalence)
+open import Data.Product
+
+∨-∧-distrib : _∨_ DistributesOver _∧_
+∨-∧-distrib = ∨-∧-distribˡ , ∨-∧-distribʳ
+ where
+ ∨-∧-distribˡ : _∨_ DistributesOverˡ _∧_
+ ∨-∧-distribˡ x y z = begin
+ x ∨ y ∧ z ≈⟨ ∨-comm _ _ ⟩
+ y ∧ z ∨ x ≈⟨ ∨-∧-distribʳ _ _ _ ⟩
+ (y ∨ x) ∧ (z ∨ x) ≈⟨ ∨-comm _ _ ⟨ ∧-cong ⟩ ∨-comm _ _ ⟩
+ (x ∨ y) ∧ (x ∨ z) ∎
+
+∧-∨-distrib : _∧_ DistributesOver _∨_
+∧-∨-distrib = ∧-∨-distribˡ , ∧-∨-distribʳ
+ where
+ ∧-∨-distribˡ : _∧_ DistributesOverˡ _∨_
+ ∧-∨-distribˡ x y z = begin
+ x ∧ (y ∨ z) ≈⟨ sym (proj₂ absorptive _ _) ⟨ ∧-cong ⟩ refl ⟩
+ (x ∧ (x ∨ y)) ∧ (y ∨ z) ≈⟨ (refl ⟨ ∧-cong ⟩ ∨-comm _ _) ⟨ ∧-cong ⟩ refl ⟩
+ (x ∧ (y ∨ x)) ∧ (y ∨ z) ≈⟨ ∧-assoc _ _ _ ⟩
+ x ∧ ((y ∨ x) ∧ (y ∨ z)) ≈⟨ refl ⟨ ∧-cong ⟩ sym (proj₁ ∨-∧-distrib _ _ _) ⟩
+ x ∧ (y ∨ x ∧ z) ≈⟨ sym (proj₁ absorptive _ _) ⟨ ∧-cong ⟩ refl ⟩
+ (x ∨ x ∧ z) ∧ (y ∨ x ∧ z) ≈⟨ sym $ proj₂ ∨-∧-distrib _ _ _ ⟩
+ x ∧ y ∨ x ∧ z ∎
+
+ ∧-∨-distribʳ : _∧_ DistributesOverʳ _∨_
+ ∧-∨-distribʳ x y z = begin
+ (y ∨ z) ∧ x ≈⟨ ∧-comm _ _ ⟩
+ x ∧ (y ∨ z) ≈⟨ ∧-∨-distribˡ _ _ _ ⟩
+ x ∧ y ∨ x ∧ z ≈⟨ ∧-comm _ _ ⟨ ∨-cong ⟩ ∧-comm _ _ ⟩
+ y ∧ x ∨ z ∧ x ∎
+
+-- The dual construction is also a distributive lattice.
+
+∧-∨-isDistributiveLattice : IsDistributiveLattice _≈_ _∧_ _∨_
+∧-∨-isDistributiveLattice = record
+ { isLattice = ∧-∨-isLattice
+ ; ∨-∧-distribʳ = proj₂ ∧-∨-distrib
+ }
+
+∧-∨-distributiveLattice : DistributiveLattice _ _
+∧-∨-distributiveLattice = record
+ { _∧_ = _∨_
+ ; _∨_ = _∧_
+ ; isDistributiveLattice = ∧-∨-isDistributiveLattice
+ }
+
+-- One can replace the underlying equality with an equivalent one.
+
+replace-equality :
+ {_≈′_ : Rel Carrier dl₂} →
+ (∀ {x y} → x ≈ y ⇔ x ≈′ y) → DistributiveLattice _ _
+replace-equality {_≈′_} ≈⇔≈′ = record
+ { _≈_ = _≈′_
+ ; _∧_ = _∧_
+ ; _∨_ = _∨_
+ ; isDistributiveLattice = record
+ { isLattice = Lattice.isLattice (L.replace-equality ≈⇔≈′)
+ ; ∨-∧-distribʳ = λ x y z → to ⟨$⟩ ∨-∧-distribʳ x y z
+ }
+ } where open module E {x y} = Equivalence (≈⇔≈′ {x} {y})
diff --git a/src/Algebra/Properties/Group.agda b/src/Algebra/Properties/Group.agda
new file mode 100644
index 0000000..93e2a9f
--- /dev/null
+++ b/src/Algebra/Properties/Group.agda
@@ -0,0 +1,70 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Some derivable properties
+------------------------------------------------------------------------
+
+open import Algebra
+
+module Algebra.Properties.Group {g₁ g₂} (G : Group g₁ g₂) where
+
+open Group G
+import Algebra.FunctionProperties as P; open P _≈_
+import Relation.Binary.EqReasoning as EqR; open EqR setoid
+open import Function
+open import Data.Product
+
+⁻¹-involutive : ∀ x → x ⁻¹ ⁻¹ ≈ x
+⁻¹-involutive x = begin
+ x ⁻¹ ⁻¹ ≈⟨ sym $ proj₂ identity _ ⟩
+ x ⁻¹ ⁻¹ ∙ ε ≈⟨ refl ⟨ ∙-cong ⟩ sym (proj₁ inverse _) ⟩
+ x ⁻¹ ⁻¹ ∙ (x ⁻¹ ∙ x) ≈⟨ sym $ assoc _ _ _ ⟩
+ x ⁻¹ ⁻¹ ∙ x ⁻¹ ∙ x ≈⟨ proj₁ inverse _ ⟨ ∙-cong ⟩ refl ⟩
+ ε ∙ x ≈⟨ proj₁ identity _ ⟩
+ x ∎
+
+private
+
+ left-helper : ∀ x y → x ≈ (x ∙ y) ∙ y ⁻¹
+ left-helper x y = begin
+ x ≈⟨ sym (proj₂ identity x) ⟩
+ x ∙ ε ≈⟨ refl ⟨ ∙-cong ⟩ sym (proj₂ inverse y) ⟩
+ x ∙ (y ∙ y ⁻¹) ≈⟨ sym (assoc x y (y ⁻¹)) ⟩
+ (x ∙ y) ∙ y ⁻¹ ∎
+
+ right-helper : ∀ x y → y ≈ x ⁻¹ ∙ (x ∙ y)
+ right-helper x y = begin
+ y ≈⟨ sym (proj₁ identity y) ⟩
+ ε ∙ y ≈⟨ sym (proj₁ inverse x) ⟨ ∙-cong ⟩ refl ⟩
+ (x ⁻¹ ∙ x) ∙ y ≈⟨ assoc (x ⁻¹) x y ⟩
+ x ⁻¹ ∙ (x ∙ y) ∎
+
+left-identity-unique : ∀ x y → x ∙ y ≈ y → x ≈ ε
+left-identity-unique x y eq = begin
+ x ≈⟨ left-helper x y ⟩
+ (x ∙ y) ∙ y ⁻¹ ≈⟨ eq ⟨ ∙-cong ⟩ refl ⟩
+ y ∙ y ⁻¹ ≈⟨ proj₂ inverse y ⟩
+ ε ∎
+
+right-identity-unique : ∀ x y → x ∙ y ≈ x → y ≈ ε
+right-identity-unique x y eq = begin
+ y ≈⟨ right-helper x y ⟩
+ x ⁻¹ ∙ (x ∙ y) ≈⟨ refl ⟨ ∙-cong ⟩ eq ⟩
+ x ⁻¹ ∙ x ≈⟨ proj₁ inverse x ⟩
+ ε ∎
+
+identity-unique : ∀ {x} → Identity x _∙_ → x ≈ ε
+identity-unique {x} id = left-identity-unique x x (proj₂ id x)
+
+left-inverse-unique : ∀ x y → x ∙ y ≈ ε → x ≈ y ⁻¹
+left-inverse-unique x y eq = begin
+ x ≈⟨ left-helper x y ⟩
+ (x ∙ y) ∙ y ⁻¹ ≈⟨ eq ⟨ ∙-cong ⟩ refl ⟩
+ ε ∙ y ⁻¹ ≈⟨ proj₁ identity (y ⁻¹) ⟩
+ y ⁻¹ ∎
+
+right-inverse-unique : ∀ x y → x ∙ y ≈ ε → y ≈ x ⁻¹
+right-inverse-unique x y eq = begin
+ y ≈⟨ sym (⁻¹-involutive y) ⟩
+ y ⁻¹ ⁻¹ ≈⟨ ⁻¹-cong (sym (left-inverse-unique x y eq)) ⟩
+ x ⁻¹ ∎
diff --git a/src/Algebra/Properties/Lattice.agda b/src/Algebra/Properties/Lattice.agda
new file mode 100644
index 0000000..13d0391
--- /dev/null
+++ b/src/Algebra/Properties/Lattice.agda
@@ -0,0 +1,106 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Some derivable properties
+------------------------------------------------------------------------
+
+open import Algebra
+
+module Algebra.Properties.Lattice {l₁ l₂} (L : Lattice l₁ l₂) where
+
+open Lattice L
+open import Algebra.Structures
+import Algebra.FunctionProperties as P; open P _≈_
+open import Relation.Binary
+import Relation.Binary.EqReasoning as EqR; open EqR setoid
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence using (_⇔_; module Equivalence)
+open import Data.Product
+
+∧-idempotent : Idempotent _∧_
+∧-idempotent x = begin
+ x ∧ x ≈⟨ refl ⟨ ∧-cong ⟩ sym (proj₁ absorptive _ _) ⟩
+ x ∧ (x ∨ x ∧ x) ≈⟨ proj₂ absorptive _ _ ⟩
+ x ∎
+
+∨-idempotent : Idempotent _∨_
+∨-idempotent x = begin
+ x ∨ x ≈⟨ refl ⟨ ∨-cong ⟩ sym (∧-idempotent _) ⟩
+ x ∨ x ∧ x ≈⟨ proj₁ absorptive _ _ ⟩
+ x ∎
+
+-- The dual construction is also a lattice.
+
+∧-∨-isLattice : IsLattice _≈_ _∧_ _∨_
+∧-∨-isLattice = record
+ { isEquivalence = isEquivalence
+ ; ∨-comm = ∧-comm
+ ; ∨-assoc = ∧-assoc
+ ; ∨-cong = ∧-cong
+ ; ∧-comm = ∨-comm
+ ; ∧-assoc = ∨-assoc
+ ; ∧-cong = ∨-cong
+ ; absorptive = swap absorptive
+ }
+
+∧-∨-lattice : Lattice _ _
+∧-∨-lattice = record
+ { _∧_ = _∨_
+ ; _∨_ = _∧_
+ ; isLattice = ∧-∨-isLattice
+ }
+
+-- Every lattice can be turned into a poset.
+
+poset : Poset _ _ _
+poset = record
+ { Carrier = Carrier
+ ; _≈_ = _≈_
+ ; _≤_ = λ x y → x ≈ x ∧ y
+ ; isPartialOrder = record
+ { isPreorder = record
+ { isEquivalence = isEquivalence
+ ; reflexive = λ {i} {j} i≈j → begin
+ i ≈⟨ sym $ ∧-idempotent _ ⟩
+ i ∧ i ≈⟨ ∧-cong refl i≈j ⟩
+ i ∧ j ∎
+ ; trans = λ {i} {j} {k} i≈i∧j j≈j∧k → begin
+ i ≈⟨ i≈i∧j ⟩
+ i ∧ j ≈⟨ ∧-cong refl j≈j∧k ⟩
+ i ∧ (j ∧ k) ≈⟨ sym (∧-assoc _ _ _) ⟩
+ (i ∧ j) ∧ k ≈⟨ ∧-cong (sym i≈i∧j) refl ⟩
+ i ∧ k ∎
+ }
+ ; antisym = λ {x} {y} x≈x∧y y≈y∧x → begin
+ x ≈⟨ x≈x∧y ⟩
+ x ∧ y ≈⟨ ∧-comm _ _ ⟩
+ y ∧ x ≈⟨ sym y≈y∧x ⟩
+ y ∎
+ }
+ }
+
+-- One can replace the underlying equality with an equivalent one.
+
+replace-equality : {_≈′_ : Rel Carrier l₂} →
+ (∀ {x y} → x ≈ y ⇔ x ≈′ y) → Lattice _ _
+replace-equality {_≈′_} ≈⇔≈′ = record
+ { _≈_ = _≈′_
+ ; _∧_ = _∧_
+ ; _∨_ = _∨_
+ ; isLattice = record
+ { isEquivalence = record
+ { refl = to ⟨$⟩ refl
+ ; sym = λ x≈y → to ⟨$⟩ sym (from ⟨$⟩ x≈y)
+ ; trans = λ x≈y y≈z → to ⟨$⟩ trans (from ⟨$⟩ x≈y) (from ⟨$⟩ y≈z)
+ }
+ ; ∨-comm = λ x y → to ⟨$⟩ ∨-comm x y
+ ; ∨-assoc = λ x y z → to ⟨$⟩ ∨-assoc x y z
+ ; ∨-cong = λ x≈y u≈v → to ⟨$⟩ ∨-cong (from ⟨$⟩ x≈y) (from ⟨$⟩ u≈v)
+ ; ∧-comm = λ x y → to ⟨$⟩ ∧-comm x y
+ ; ∧-assoc = λ x y z → to ⟨$⟩ ∧-assoc x y z
+ ; ∧-cong = λ x≈y u≈v → to ⟨$⟩ ∧-cong (from ⟨$⟩ x≈y) (from ⟨$⟩ u≈v)
+ ; absorptive = (λ x y → to ⟨$⟩ proj₁ absorptive x y)
+ , (λ x y → to ⟨$⟩ proj₂ absorptive x y)
+ }
+ } where open module E {x y} = Equivalence (≈⇔≈′ {x} {y})
diff --git a/src/Algebra/Properties/Ring.agda b/src/Algebra/Properties/Ring.agda
new file mode 100644
index 0000000..84ef44d
--- /dev/null
+++ b/src/Algebra/Properties/Ring.agda
@@ -0,0 +1,51 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Some derivable properties
+------------------------------------------------------------------------
+
+open import Algebra
+
+module Algebra.Properties.Ring {r₁ r₂} (R : Ring r₁ r₂) where
+
+import Algebra.Properties.AbelianGroup as AGP
+open import Data.Product
+open import Function
+import Relation.Binary.EqReasoning as EqR
+
+open Ring R
+open EqR setoid
+
+open AGP +-abelianGroup public
+ renaming ( ⁻¹-involutive to -‿involutive
+ ; left-identity-unique to +-left-identity-unique
+ ; right-identity-unique to +-right-identity-unique
+ ; identity-unique to +-identity-unique
+ ; left-inverse-unique to +-left-inverse-unique
+ ; right-inverse-unique to +-right-inverse-unique
+ ; ⁻¹-∙-comm to -‿+-comm
+ )
+
+-‿*-distribˡ : ∀ x y → - x * y ≈ - (x * y)
+-‿*-distribˡ x y = begin
+ - x * y ≈⟨ sym $ proj₂ +-identity _ ⟩
+ - x * y + 0# ≈⟨ refl ⟨ +-cong ⟩ sym (proj₂ -‿inverse _) ⟩
+ - x * y + (x * y + - (x * y)) ≈⟨ sym $ +-assoc _ _ _ ⟩
+ - x * y + x * y + - (x * y) ≈⟨ sym (proj₂ distrib _ _ _) ⟨ +-cong ⟩ refl ⟩
+ (- x + x) * y + - (x * y) ≈⟨ (proj₁ -‿inverse _ ⟨ *-cong ⟩ refl)
+ ⟨ +-cong ⟩
+ refl ⟩
+ 0# * y + - (x * y) ≈⟨ proj₁ zero _ ⟨ +-cong ⟩ refl ⟩
+ 0# + - (x * y) ≈⟨ proj₁ +-identity _ ⟩
+ - (x * y) ∎
+
+-‿*-distribʳ : ∀ x y → x * - y ≈ - (x * y)
+-‿*-distribʳ x y = begin
+ x * - y ≈⟨ sym $ proj₁ +-identity _ ⟩
+ 0# + x * - y ≈⟨ sym (proj₁ -‿inverse _) ⟨ +-cong ⟩ refl ⟩
+ - (x * y) + x * y + x * - y ≈⟨ +-assoc _ _ _ ⟩
+ - (x * y) + (x * y + x * - y) ≈⟨ refl ⟨ +-cong ⟩ sym (proj₁ distrib _ _ _) ⟩
+ - (x * y) + x * (y + - y) ≈⟨ refl ⟨ +-cong ⟩ (refl ⟨ *-cong ⟩ proj₂ -‿inverse _) ⟩
+ - (x * y) + x * 0# ≈⟨ refl ⟨ +-cong ⟩ proj₂ zero _ ⟩
+ - (x * y) + 0# ≈⟨ proj₂ +-identity _ ⟩
+ - (x * y) ∎
diff --git a/src/Algebra/Props/AbelianGroup.agda b/src/Algebra/Props/AbelianGroup.agda
index 87d62a0..601029d 100644
--- a/src/Algebra/Props/AbelianGroup.agda
+++ b/src/Algebra/Props/AbelianGroup.agda
@@ -1,7 +1,8 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Some derivable properties
+-- Compatibility module. Pending for removal. Use
+-- Algebra.Properties.AbelianGroup instead.
------------------------------------------------------------------------
open import Algebra
@@ -9,36 +10,4 @@ open import Algebra
module Algebra.Props.AbelianGroup
{g₁ g₂} (G : AbelianGroup g₁ g₂) where
-import Algebra.Props.Group as GP
-open import Data.Product
-open import Function
-import Relation.Binary.EqReasoning as EqR
-
-open AbelianGroup G
-open EqR setoid
-
-open GP group public
-
-private
- lemma : ∀ x y → x ∙ y ∙ x ⁻¹ ≈ y
- lemma x y = begin
- x ∙ y ∙ x ⁻¹ ≈⟨ comm _ _ ⟨ ∙-cong ⟩ refl ⟩
- y ∙ x ∙ x ⁻¹ ≈⟨ assoc _ _ _ ⟩
- y ∙ (x ∙ x ⁻¹) ≈⟨ refl ⟨ ∙-cong ⟩ proj₂ inverse _ ⟩
- y ∙ ε ≈⟨ proj₂ identity _ ⟩
- y ∎
-
-⁻¹-∙-comm : ∀ x y → x ⁻¹ ∙ y ⁻¹ ≈ (x ∙ y) ⁻¹
-⁻¹-∙-comm x y = begin
- x ⁻¹ ∙ y ⁻¹ ≈⟨ comm _ _ ⟩
- y ⁻¹ ∙ x ⁻¹ ≈⟨ sym $ lem ⟨ ∙-cong ⟩ refl ⟩
- x ∙ (y ∙ (x ∙ y) ⁻¹ ∙ y ⁻¹) ∙ x ⁻¹ ≈⟨ lemma _ _ ⟩
- y ∙ (x ∙ y) ⁻¹ ∙ y ⁻¹ ≈⟨ lemma _ _ ⟩
- (x ∙ y) ⁻¹ ∎
- where
- lem = begin
- x ∙ (y ∙ (x ∙ y) ⁻¹ ∙ y ⁻¹) ≈⟨ sym $ assoc _ _ _ ⟩
- x ∙ (y ∙ (x ∙ y) ⁻¹) ∙ y ⁻¹ ≈⟨ sym $ assoc _ _ _ ⟨ ∙-cong ⟩ refl ⟩
- x ∙ y ∙ (x ∙ y) ⁻¹ ∙ y ⁻¹ ≈⟨ proj₂ inverse _ ⟨ ∙-cong ⟩ refl ⟩
- ε ∙ y ⁻¹ ≈⟨ proj₁ identity _ ⟩
- y ⁻¹ ∎
+open import Algebra.Properties.AbelianGroup G public
diff --git a/src/Algebra/Props/BooleanAlgebra.agda b/src/Algebra/Props/BooleanAlgebra.agda
index de20cb2..4cf9a38 100644
--- a/src/Algebra/Props/BooleanAlgebra.agda
+++ b/src/Algebra/Props/BooleanAlgebra.agda
@@ -1,7 +1,8 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Some derivable properties
+-- Compatibility module. Pending for removal. Use
+-- Algebra.Properties.BooleanAlgebra instead.
------------------------------------------------------------------------
open import Algebra
@@ -10,561 +11,4 @@ module Algebra.Props.BooleanAlgebra
{b₁ b₂} (B : BooleanAlgebra b₁ b₂)
where
-open BooleanAlgebra B
-import Algebra.Props.DistributiveLattice
-private
- open module DL = Algebra.Props.DistributiveLattice
- distributiveLattice public
- hiding (replace-equality)
-open import Algebra.Structures
-import Algebra.FunctionProperties as P; open P _≈_
-import Relation.Binary.EqReasoning as EqR; open EqR setoid
-open import Relation.Binary
-open import Function
-open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence using (_⇔_; module Equivalence)
-open import Data.Product
-
-------------------------------------------------------------------------
--- Some simple generalisations
-
-∨-complement : Inverse ⊤ ¬_ _∨_
-∨-complement = ∨-complementˡ , ∨-complementʳ
- where
- ∨-complementˡ : LeftInverse ⊤ ¬_ _∨_
- ∨-complementˡ x = begin
- ¬ x ∨ x ≈⟨ ∨-comm _ _ ⟩
- x ∨ ¬ x ≈⟨ ∨-complementʳ _ ⟩
- ⊤ ∎
-
-∧-complement : Inverse ⊥ ¬_ _∧_
-∧-complement = ∧-complementˡ , ∧-complementʳ
- where
- ∧-complementˡ : LeftInverse ⊥ ¬_ _∧_
- ∧-complementˡ x = begin
- ¬ x ∧ x ≈⟨ ∧-comm _ _ ⟩
- x ∧ ¬ x ≈⟨ ∧-complementʳ _ ⟩
- ⊥ ∎
-
-------------------------------------------------------------------------
--- The dual construction is also a boolean algebra
-
-∧-∨-isBooleanAlgebra : IsBooleanAlgebra _≈_ _∧_ _∨_ ¬_ ⊥ ⊤
-∧-∨-isBooleanAlgebra = record
- { isDistributiveLattice = ∧-∨-isDistributiveLattice
- ; ∨-complementʳ = proj₂ ∧-complement
- ; ∧-complementʳ = proj₂ ∨-complement
- ; ¬-cong = ¬-cong
- }
-
-∧-∨-booleanAlgebra : BooleanAlgebra _ _
-∧-∨-booleanAlgebra = record
- { _∧_ = _∨_
- ; _∨_ = _∧_
- ; ⊤ = ⊥
- ; ⊥ = ⊤
- ; isBooleanAlgebra = ∧-∨-isBooleanAlgebra
- }
-
-------------------------------------------------------------------------
--- (∨, ∧, ⊥, ⊤) is a commutative semiring
-
-private
-
- ∧-identity : Identity ⊤ _∧_
- ∧-identity = (λ _ → ∧-comm _ _ ⟨ trans ⟩ x∧⊤=x _) , x∧⊤=x
- where
- x∧⊤=x : ∀ x → x ∧ ⊤ ≈ x
- x∧⊤=x x = begin
- x ∧ ⊤ ≈⟨ refl ⟨ ∧-cong ⟩ sym (proj₂ ∨-complement _) ⟩
- x ∧ (x ∨ ¬ x) ≈⟨ proj₂ absorptive _ _ ⟩
- x ∎
-
- ∨-identity : Identity ⊥ _∨_
- ∨-identity = (λ _ → ∨-comm _ _ ⟨ trans ⟩ x∨⊥=x _) , x∨⊥=x
- where
- x∨⊥=x : ∀ x → x ∨ ⊥ ≈ x
- x∨⊥=x x = begin
- x ∨ ⊥ ≈⟨ refl ⟨ ∨-cong ⟩ sym (proj₂ ∧-complement _) ⟩
- x ∨ x ∧ ¬ x ≈⟨ proj₁ absorptive _ _ ⟩
- x ∎
-
- ∧-zero : Zero ⊥ _∧_
- ∧-zero = (λ _ → ∧-comm _ _ ⟨ trans ⟩ x∧⊥=⊥ _) , x∧⊥=⊥
- where
- x∧⊥=⊥ : ∀ x → x ∧ ⊥ ≈ ⊥
- x∧⊥=⊥ x = begin
- x ∧ ⊥ ≈⟨ refl ⟨ ∧-cong ⟩ sym (proj₂ ∧-complement _) ⟩
- x ∧ x ∧ ¬ x ≈⟨ sym $ ∧-assoc _ _ _ ⟩
- (x ∧ x) ∧ ¬ x ≈⟨ ∧-idempotent _ ⟨ ∧-cong ⟩ refl ⟩
- x ∧ ¬ x ≈⟨ proj₂ ∧-complement _ ⟩
- ⊥ ∎
-
-∨-∧-isCommutativeSemiring : IsCommutativeSemiring _≈_ _∨_ _∧_ ⊥ ⊤
-∨-∧-isCommutativeSemiring = record
- { +-isCommutativeMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∨-assoc
- ; ∙-cong = ∨-cong
- }
- ; identityˡ = proj₁ ∨-identity
- ; comm = ∨-comm
- }
- ; *-isCommutativeMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∧-assoc
- ; ∙-cong = ∧-cong
- }
- ; identityˡ = proj₁ ∧-identity
- ; comm = ∧-comm
- }
- ; distribʳ = proj₂ ∧-∨-distrib
- ; zeroˡ = proj₁ ∧-zero
- }
-
-∨-∧-commutativeSemiring : CommutativeSemiring _ _
-∨-∧-commutativeSemiring = record
- { _+_ = _∨_
- ; _*_ = _∧_
- ; 0# = ⊥
- ; 1# = ⊤
- ; isCommutativeSemiring = ∨-∧-isCommutativeSemiring
- }
-
-------------------------------------------------------------------------
--- (∧, ∨, ⊤, ⊥) is a commutative semiring
-
-private
-
- ∨-zero : Zero ⊤ _∨_
- ∨-zero = (λ _ → ∨-comm _ _ ⟨ trans ⟩ x∨⊤=⊤ _) , x∨⊤=⊤
- where
- x∨⊤=⊤ : ∀ x → x ∨ ⊤ ≈ ⊤
- x∨⊤=⊤ x = begin
- x ∨ ⊤ ≈⟨ refl ⟨ ∨-cong ⟩ sym (proj₂ ∨-complement _) ⟩
- x ∨ x ∨ ¬ x ≈⟨ sym $ ∨-assoc _ _ _ ⟩
- (x ∨ x) ∨ ¬ x ≈⟨ ∨-idempotent _ ⟨ ∨-cong ⟩ refl ⟩
- x ∨ ¬ x ≈⟨ proj₂ ∨-complement _ ⟩
- ⊤ ∎
-
-∧-∨-isCommutativeSemiring : IsCommutativeSemiring _≈_ _∧_ _∨_ ⊤ ⊥
-∧-∨-isCommutativeSemiring = record
- { +-isCommutativeMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∧-assoc
- ; ∙-cong = ∧-cong
- }
- ; identityˡ = proj₁ ∧-identity
- ; comm = ∧-comm
- }
- ; *-isCommutativeMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∨-assoc
- ; ∙-cong = ∨-cong
- }
- ; identityˡ = proj₁ ∨-identity
- ; comm = ∨-comm
- }
- ; distribʳ = proj₂ ∨-∧-distrib
- ; zeroˡ = proj₁ ∨-zero
- }
-
-∧-∨-commutativeSemiring : CommutativeSemiring _ _
-∧-∨-commutativeSemiring =
- record { isCommutativeSemiring = ∧-∨-isCommutativeSemiring }
-
-------------------------------------------------------------------------
--- Some other properties
-
--- I took the statement of this lemma (called Uniqueness of
--- Complements) from some course notes, "Boolean Algebra", written
--- by Gert Smolka.
-
-private
- lemma : ∀ x y → x ∧ y ≈ ⊥ → x ∨ y ≈ ⊤ → ¬ x ≈ y
- lemma x y x∧y=⊥ x∨y=⊤ = begin
- ¬ x ≈⟨ sym $ proj₂ ∧-identity _ ⟩
- ¬ x ∧ ⊤ ≈⟨ refl ⟨ ∧-cong ⟩ sym x∨y=⊤ ⟩
- ¬ x ∧ (x ∨ y) ≈⟨ proj₁ ∧-∨-distrib _ _ _ ⟩
- ¬ x ∧ x ∨ ¬ x ∧ y ≈⟨ proj₁ ∧-complement _ ⟨ ∨-cong ⟩ refl ⟩
- ⊥ ∨ ¬ x ∧ y ≈⟨ sym x∧y=⊥ ⟨ ∨-cong ⟩ refl ⟩
- x ∧ y ∨ ¬ x ∧ y ≈⟨ sym $ proj₂ ∧-∨-distrib _ _ _ ⟩
- (x ∨ ¬ x) ∧ y ≈⟨ proj₂ ∨-complement _ ⟨ ∧-cong ⟩ refl ⟩
- ⊤ ∧ y ≈⟨ proj₁ ∧-identity _ ⟩
- y ∎
-
-¬⊥=⊤ : ¬ ⊥ ≈ ⊤
-¬⊥=⊤ = lemma ⊥ ⊤ (proj₂ ∧-identity _) (proj₂ ∨-zero _)
-
-¬⊤=⊥ : ¬ ⊤ ≈ ⊥
-¬⊤=⊥ = lemma ⊤ ⊥ (proj₂ ∧-zero _) (proj₂ ∨-identity _)
-
-¬-involutive : Involutive ¬_
-¬-involutive x = lemma (¬ x) x (proj₁ ∧-complement _) (proj₁ ∨-complement _)
-
-deMorgan₁ : ∀ x y → ¬ (x ∧ y) ≈ ¬ x ∨ ¬ y
-deMorgan₁ x y = lemma (x ∧ y) (¬ x ∨ ¬ y) lem₁ lem₂
- where
- lem₁ = begin
- (x ∧ y) ∧ (¬ x ∨ ¬ y) ≈⟨ proj₁ ∧-∨-distrib _ _ _ ⟩
- (x ∧ y) ∧ ¬ x ∨ (x ∧ y) ∧ ¬ y ≈⟨ (∧-comm _ _ ⟨ ∧-cong ⟩ refl) ⟨ ∨-cong ⟩ refl ⟩
- (y ∧ x) ∧ ¬ x ∨ (x ∧ y) ∧ ¬ y ≈⟨ ∧-assoc _ _ _ ⟨ ∨-cong ⟩ ∧-assoc _ _ _ ⟩
- y ∧ (x ∧ ¬ x) ∨ x ∧ (y ∧ ¬ y) ≈⟨ (refl ⟨ ∧-cong ⟩ proj₂ ∧-complement _) ⟨ ∨-cong ⟩
- (refl ⟨ ∧-cong ⟩ proj₂ ∧-complement _) ⟩
- (y ∧ ⊥) ∨ (x ∧ ⊥) ≈⟨ proj₂ ∧-zero _ ⟨ ∨-cong ⟩ proj₂ ∧-zero _ ⟩
- ⊥ ∨ ⊥ ≈⟨ proj₂ ∨-identity _ ⟩
- ⊥ ∎
-
- lem₃ = begin
- (x ∧ y) ∨ ¬ x ≈⟨ proj₂ ∨-∧-distrib _ _ _ ⟩
- (x ∨ ¬ x) ∧ (y ∨ ¬ x) ≈⟨ proj₂ ∨-complement _ ⟨ ∧-cong ⟩ refl ⟩
- ⊤ ∧ (y ∨ ¬ x) ≈⟨ proj₁ ∧-identity _ ⟩
- y ∨ ¬ x ≈⟨ ∨-comm _ _ ⟩
- ¬ x ∨ y ∎
-
- lem₂ = begin
- (x ∧ y) ∨ (¬ x ∨ ¬ y) ≈⟨ sym $ ∨-assoc _ _ _ ⟩
- ((x ∧ y) ∨ ¬ x) ∨ ¬ y ≈⟨ lem₃ ⟨ ∨-cong ⟩ refl ⟩
- (¬ x ∨ y) ∨ ¬ y ≈⟨ ∨-assoc _ _ _ ⟩
- ¬ x ∨ (y ∨ ¬ y) ≈⟨ refl ⟨ ∨-cong ⟩ proj₂ ∨-complement _ ⟩
- ¬ x ∨ ⊤ ≈⟨ proj₂ ∨-zero _ ⟩
- ⊤ ∎
-
-deMorgan₂ : ∀ x y → ¬ (x ∨ y) ≈ ¬ x ∧ ¬ y
-deMorgan₂ x y = begin
- ¬ (x ∨ y) ≈⟨ ¬-cong $ sym (¬-involutive _) ⟨ ∨-cong ⟩
- sym (¬-involutive _) ⟩
- ¬ (¬ ¬ x ∨ ¬ ¬ y) ≈⟨ ¬-cong $ sym $ deMorgan₁ _ _ ⟩
- ¬ ¬ (¬ x ∧ ¬ y) ≈⟨ ¬-involutive _ ⟩
- ¬ x ∧ ¬ y ∎
-
--- One can replace the underlying equality with an equivalent one.
-
-replace-equality :
- {_≈′_ : Rel Carrier b₂} →
- (∀ {x y} → x ≈ y ⇔ x ≈′ y) → BooleanAlgebra _ _
-replace-equality {_≈′_} ≈⇔≈′ = record
- { _≈_ = _≈′_
- ; _∨_ = _∨_
- ; _∧_ = _∧_
- ; ¬_ = ¬_
- ; ⊤ = ⊤
- ; ⊥ = ⊥
- ; isBooleanAlgebra = record
- { isDistributiveLattice = DistributiveLattice.isDistributiveLattice
- (DL.replace-equality ≈⇔≈′)
- ; ∨-complementʳ = λ x → to ⟨$⟩ ∨-complementʳ x
- ; ∧-complementʳ = λ x → to ⟨$⟩ ∧-complementʳ x
- ; ¬-cong = λ i≈j → to ⟨$⟩ ¬-cong (from ⟨$⟩ i≈j)
- }
- } where open module E {x y} = Equivalence (≈⇔≈′ {x} {y})
-
-------------------------------------------------------------------------
--- (⊕, ∧, id, ⊥, ⊤) is a commutative ring
-
--- This construction is parameterised over the definition of xor.
-
-module XorRing
- (xor : Op₂ Carrier)
- (⊕-def : ∀ x y → xor x y ≈ (x ∨ y) ∧ ¬ (x ∧ y))
- where
-
- private
- infixl 6 _⊕_
-
- _⊕_ : Op₂ Carrier
- _⊕_ = xor
-
- private
- helper : ∀ {x y u v} → x ≈ y → u ≈ v → x ∧ ¬ u ≈ y ∧ ¬ v
- helper x≈y u≈v = x≈y ⟨ ∧-cong ⟩ ¬-cong u≈v
-
- ⊕-¬-distribˡ : ∀ x y → ¬ (x ⊕ y) ≈ ¬ x ⊕ y
- ⊕-¬-distribˡ x y = begin
- ¬ (x ⊕ y) ≈⟨ ¬-cong $ ⊕-def _ _ ⟩
- ¬ ((x ∨ y) ∧ (¬ (x ∧ y))) ≈⟨ ¬-cong (proj₂ ∧-∨-distrib _ _ _) ⟩
- ¬ ((x ∧ ¬ (x ∧ y)) ∨ (y ∧ ¬ (x ∧ y))) ≈⟨ ¬-cong $
- refl ⟨ ∨-cong ⟩
- (refl ⟨ ∧-cong ⟩
- ¬-cong (∧-comm _ _)) ⟩
- ¬ ((x ∧ ¬ (x ∧ y)) ∨ (y ∧ ¬ (y ∧ x))) ≈⟨ ¬-cong $ lem _ _ ⟨ ∨-cong ⟩ lem _ _ ⟩
- ¬ ((x ∧ ¬ y) ∨ (y ∧ ¬ x)) ≈⟨ deMorgan₂ _ _ ⟩
- ¬ (x ∧ ¬ y) ∧ ¬ (y ∧ ¬ x) ≈⟨ deMorgan₁ _ _ ⟨ ∧-cong ⟩ refl ⟩
- (¬ x ∨ (¬ ¬ y)) ∧ ¬ (y ∧ ¬ x) ≈⟨ helper (refl ⟨ ∨-cong ⟩ ¬-involutive _)
- (∧-comm _ _) ⟩
- (¬ x ∨ y) ∧ ¬ (¬ x ∧ y) ≈⟨ sym $ ⊕-def _ _ ⟩
- ¬ x ⊕ y ∎
- where
- lem : ∀ x y → x ∧ ¬ (x ∧ y) ≈ x ∧ ¬ y
- lem x y = begin
- x ∧ ¬ (x ∧ y) ≈⟨ refl ⟨ ∧-cong ⟩ deMorgan₁ _ _ ⟩
- x ∧ (¬ x ∨ ¬ y) ≈⟨ proj₁ ∧-∨-distrib _ _ _ ⟩
- (x ∧ ¬ x) ∨ (x ∧ ¬ y) ≈⟨ proj₂ ∧-complement _ ⟨ ∨-cong ⟩ refl ⟩
- ⊥ ∨ (x ∧ ¬ y) ≈⟨ proj₁ ∨-identity _ ⟩
- x ∧ ¬ y ∎
-
- private
- ⊕-comm : Commutative _⊕_
- ⊕-comm x y = begin
- x ⊕ y ≈⟨ ⊕-def _ _ ⟩
- (x ∨ y) ∧ ¬ (x ∧ y) ≈⟨ helper (∨-comm _ _) (∧-comm _ _) ⟩
- (y ∨ x) ∧ ¬ (y ∧ x) ≈⟨ sym $ ⊕-def _ _ ⟩
- y ⊕ x ∎
-
- ⊕-¬-distribʳ : ∀ x y → ¬ (x ⊕ y) ≈ x ⊕ ¬ y
- ⊕-¬-distribʳ x y = begin
- ¬ (x ⊕ y) ≈⟨ ¬-cong $ ⊕-comm _ _ ⟩
- ¬ (y ⊕ x) ≈⟨ ⊕-¬-distribˡ _ _ ⟩
- ¬ y ⊕ x ≈⟨ ⊕-comm _ _ ⟩
- x ⊕ ¬ y ∎
-
- ⊕-annihilates-¬ : ∀ x y → x ⊕ y ≈ ¬ x ⊕ ¬ y
- ⊕-annihilates-¬ x y = begin
- x ⊕ y ≈⟨ sym $ ¬-involutive _ ⟩
- ¬ ¬ (x ⊕ y) ≈⟨ ¬-cong $ ⊕-¬-distribˡ _ _ ⟩
- ¬ (¬ x ⊕ y) ≈⟨ ⊕-¬-distribʳ _ _ ⟩
- ¬ x ⊕ ¬ y ∎
-
- private
- ⊕-cong : _⊕_ Preserves₂ _≈_ ⟶ _≈_ ⟶ _≈_
- ⊕-cong {x} {y} {u} {v} x≈y u≈v = begin
- x ⊕ u ≈⟨ ⊕-def _ _ ⟩
- (x ∨ u) ∧ ¬ (x ∧ u) ≈⟨ helper (x≈y ⟨ ∨-cong ⟩ u≈v)
- (x≈y ⟨ ∧-cong ⟩ u≈v) ⟩
- (y ∨ v) ∧ ¬ (y ∧ v) ≈⟨ sym $ ⊕-def _ _ ⟩
- y ⊕ v ∎
-
- ⊕-identity : Identity ⊥ _⊕_
- ⊕-identity = ⊥⊕x=x , (λ _ → ⊕-comm _ _ ⟨ trans ⟩ ⊥⊕x=x _)
- where
- ⊥⊕x=x : ∀ x → ⊥ ⊕ x ≈ x
- ⊥⊕x=x x = begin
- ⊥ ⊕ x ≈⟨ ⊕-def _ _ ⟩
- (⊥ ∨ x) ∧ ¬ (⊥ ∧ x) ≈⟨ helper (proj₁ ∨-identity _)
- (proj₁ ∧-zero _) ⟩
- x ∧ ¬ ⊥ ≈⟨ refl ⟨ ∧-cong ⟩ ¬⊥=⊤ ⟩
- x ∧ ⊤ ≈⟨ proj₂ ∧-identity _ ⟩
- x ∎
-
- ⊕-inverse : Inverse ⊥ id _⊕_
- ⊕-inverse = x⊕x=⊥ , (λ _ → ⊕-comm _ _ ⟨ trans ⟩ x⊕x=⊥ _)
- where
- x⊕x=⊥ : ∀ x → x ⊕ x ≈ ⊥
- x⊕x=⊥ x = begin
- x ⊕ x ≈⟨ ⊕-def _ _ ⟩
- (x ∨ x) ∧ ¬ (x ∧ x) ≈⟨ helper (∨-idempotent _)
- (∧-idempotent _) ⟩
- x ∧ ¬ x ≈⟨ proj₂ ∧-complement _ ⟩
- ⊥ ∎
-
- distrib-∧-⊕ : _∧_ DistributesOver _⊕_
- distrib-∧-⊕ = distˡ , distʳ
- where
- distˡ : _∧_ DistributesOverˡ _⊕_
- distˡ x y z = begin
- x ∧ (y ⊕ z) ≈⟨ refl ⟨ ∧-cong ⟩ ⊕-def _ _ ⟩
- x ∧ ((y ∨ z) ∧ ¬ (y ∧ z)) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
- (x ∧ (y ∨ z)) ∧ ¬ (y ∧ z) ≈⟨ refl ⟨ ∧-cong ⟩ deMorgan₁ _ _ ⟩
- (x ∧ (y ∨ z)) ∧
- (¬ y ∨ ¬ z) ≈⟨ sym $ proj₁ ∨-identity _ ⟩
- ⊥ ∨
- ((x ∧ (y ∨ z)) ∧
- (¬ y ∨ ¬ z)) ≈⟨ lem₃ ⟨ ∨-cong ⟩ refl ⟩
- ((x ∧ (y ∨ z)) ∧ ¬ x) ∨
- ((x ∧ (y ∨ z)) ∧
- (¬ y ∨ ¬ z)) ≈⟨ sym $ proj₁ ∧-∨-distrib _ _ _ ⟩
- (x ∧ (y ∨ z)) ∧
- (¬ x ∨ (¬ y ∨ ¬ z)) ≈⟨ refl ⟨ ∧-cong ⟩
- (refl ⟨ ∨-cong ⟩ sym (deMorgan₁ _ _)) ⟩
- (x ∧ (y ∨ z)) ∧
- (¬ x ∨ ¬ (y ∧ z)) ≈⟨ refl ⟨ ∧-cong ⟩ sym (deMorgan₁ _ _) ⟩
- (x ∧ (y ∨ z)) ∧
- ¬ (x ∧ (y ∧ z)) ≈⟨ helper refl lem₁ ⟩
- (x ∧ (y ∨ z)) ∧
- ¬ ((x ∧ y) ∧ (x ∧ z)) ≈⟨ proj₁ ∧-∨-distrib _ _ _ ⟨ ∧-cong ⟩
- refl ⟩
- ((x ∧ y) ∨ (x ∧ z)) ∧
- ¬ ((x ∧ y) ∧ (x ∧ z)) ≈⟨ sym $ ⊕-def _ _ ⟩
- (x ∧ y) ⊕ (x ∧ z) ∎
- where
- lem₂ = begin
- x ∧ (y ∧ z) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
- (x ∧ y) ∧ z ≈⟨ ∧-comm _ _ ⟨ ∧-cong ⟩ refl ⟩
- (y ∧ x) ∧ z ≈⟨ ∧-assoc _ _ _ ⟩
- y ∧ (x ∧ z) ∎
-
- lem₁ = begin
- x ∧ (y ∧ z) ≈⟨ sym (∧-idempotent _) ⟨ ∧-cong ⟩ refl ⟩
- (x ∧ x) ∧ (y ∧ z) ≈⟨ ∧-assoc _ _ _ ⟩
- x ∧ (x ∧ (y ∧ z)) ≈⟨ refl ⟨ ∧-cong ⟩ lem₂ ⟩
- x ∧ (y ∧ (x ∧ z)) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
- (x ∧ y) ∧ (x ∧ z) ∎
-
- lem₃ = begin
- ⊥ ≈⟨ sym $ proj₂ ∧-zero _ ⟩
- (y ∨ z) ∧ ⊥ ≈⟨ refl ⟨ ∧-cong ⟩ sym (proj₂ ∧-complement _) ⟩
- (y ∨ z) ∧ (x ∧ ¬ x) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
- ((y ∨ z) ∧ x) ∧ ¬ x ≈⟨ ∧-comm _ _ ⟨ ∧-cong ⟩ refl ⟩
- (x ∧ (y ∨ z)) ∧ ¬ x ∎
-
- distʳ : _∧_ DistributesOverʳ _⊕_
- distʳ x y z = begin
- (y ⊕ z) ∧ x ≈⟨ ∧-comm _ _ ⟩
- x ∧ (y ⊕ z) ≈⟨ distˡ _ _ _ ⟩
- (x ∧ y) ⊕ (x ∧ z) ≈⟨ ∧-comm _ _ ⟨ ⊕-cong ⟩ ∧-comm _ _ ⟩
- (y ∧ x) ⊕ (z ∧ x) ∎
-
- lemma₂ : ∀ x y u v →
- (x ∧ y) ∨ (u ∧ v) ≈
- ((x ∨ u) ∧ (y ∨ u)) ∧
- ((x ∨ v) ∧ (y ∨ v))
- lemma₂ x y u v = begin
- (x ∧ y) ∨ (u ∧ v) ≈⟨ proj₁ ∨-∧-distrib _ _ _ ⟩
- ((x ∧ y) ∨ u) ∧ ((x ∧ y) ∨ v) ≈⟨ proj₂ ∨-∧-distrib _ _ _
- ⟨ ∧-cong ⟩
- proj₂ ∨-∧-distrib _ _ _ ⟩
- ((x ∨ u) ∧ (y ∨ u)) ∧
- ((x ∨ v) ∧ (y ∨ v)) ∎
-
- ⊕-assoc : Associative _⊕_
- ⊕-assoc x y z = sym $ begin
- x ⊕ (y ⊕ z) ≈⟨ refl ⟨ ⊕-cong ⟩ ⊕-def _ _ ⟩
- x ⊕ ((y ∨ z) ∧ ¬ (y ∧ z)) ≈⟨ ⊕-def _ _ ⟩
- (x ∨ ((y ∨ z) ∧ ¬ (y ∧ z))) ∧
- ¬ (x ∧ ((y ∨ z) ∧ ¬ (y ∧ z))) ≈⟨ lem₃ ⟨ ∧-cong ⟩ lem₄ ⟩
- (((x ∨ y) ∨ z) ∧ ((x ∨ ¬ y) ∨ ¬ z)) ∧
- (((¬ x ∨ ¬ y) ∨ z) ∧ ((¬ x ∨ y) ∨ ¬ z)) ≈⟨ ∧-assoc _ _ _ ⟩
- ((x ∨ y) ∨ z) ∧
- (((x ∨ ¬ y) ∨ ¬ z) ∧
- (((¬ x ∨ ¬ y) ∨ z) ∧ ((¬ x ∨ y) ∨ ¬ z))) ≈⟨ refl ⟨ ∧-cong ⟩ lem₅ ⟩
- ((x ∨ y) ∨ z) ∧
- (((¬ x ∨ ¬ y) ∨ z) ∧
- (((x ∨ ¬ y) ∨ ¬ z) ∧ ((¬ x ∨ y) ∨ ¬ z))) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
- (((x ∨ y) ∨ z) ∧ ((¬ x ∨ ¬ y) ∨ z)) ∧
- (((x ∨ ¬ y) ∨ ¬ z) ∧ ((¬ x ∨ y) ∨ ¬ z)) ≈⟨ lem₁ ⟨ ∧-cong ⟩ lem₂ ⟩
- (((x ∨ y) ∧ ¬ (x ∧ y)) ∨ z) ∧
- ¬ (((x ∨ y) ∧ ¬ (x ∧ y)) ∧ z) ≈⟨ sym $ ⊕-def _ _ ⟩
- ((x ∨ y) ∧ ¬ (x ∧ y)) ⊕ z ≈⟨ sym $ ⊕-def _ _ ⟨ ⊕-cong ⟩ refl ⟩
- (x ⊕ y) ⊕ z ∎
- where
- lem₁ = begin
- ((x ∨ y) ∨ z) ∧ ((¬ x ∨ ¬ y) ∨ z) ≈⟨ sym $ proj₂ ∨-∧-distrib _ _ _ ⟩
- ((x ∨ y) ∧ (¬ x ∨ ¬ y)) ∨ z ≈⟨ (refl ⟨ ∧-cong ⟩ sym (deMorgan₁ _ _))
- ⟨ ∨-cong ⟩ refl ⟩
- ((x ∨ y) ∧ ¬ (x ∧ y)) ∨ z ∎
-
- lem₂' = begin
- (x ∨ ¬ y) ∧ (¬ x ∨ y) ≈⟨ sym $
- proj₁ ∧-identity _
- ⟨ ∧-cong ⟩
- proj₂ ∧-identity _ ⟩
- (⊤ ∧ (x ∨ ¬ y)) ∧ ((¬ x ∨ y) ∧ ⊤) ≈⟨ sym $
- (proj₁ ∨-complement _ ⟨ ∧-cong ⟩ ∨-comm _ _)
- ⟨ ∧-cong ⟩
- (refl ⟨ ∧-cong ⟩ proj₁ ∨-complement _) ⟩
- ((¬ x ∨ x) ∧ (¬ y ∨ x)) ∧
- ((¬ x ∨ y) ∧ (¬ y ∨ y)) ≈⟨ sym $ lemma₂ _ _ _ _ ⟩
- (¬ x ∧ ¬ y) ∨ (x ∧ y) ≈⟨ sym $ deMorgan₂ _ _ ⟨ ∨-cong ⟩ ¬-involutive _ ⟩
- ¬ (x ∨ y) ∨ ¬ ¬ (x ∧ y) ≈⟨ sym (deMorgan₁ _ _) ⟩
- ¬ ((x ∨ y) ∧ ¬ (x ∧ y)) ∎
-
- lem₂ = begin
- ((x ∨ ¬ y) ∨ ¬ z) ∧ ((¬ x ∨ y) ∨ ¬ z) ≈⟨ sym $ proj₂ ∨-∧-distrib _ _ _ ⟩
- ((x ∨ ¬ y) ∧ (¬ x ∨ y)) ∨ ¬ z ≈⟨ lem₂' ⟨ ∨-cong ⟩ refl ⟩
- ¬ ((x ∨ y) ∧ ¬ (x ∧ y)) ∨ ¬ z ≈⟨ sym $ deMorgan₁ _ _ ⟩
- ¬ (((x ∨ y) ∧ ¬ (x ∧ y)) ∧ z) ∎
-
- lem₃ = begin
- x ∨ ((y ∨ z) ∧ ¬ (y ∧ z)) ≈⟨ refl ⟨ ∨-cong ⟩
- (refl ⟨ ∧-cong ⟩ deMorgan₁ _ _) ⟩
- x ∨ ((y ∨ z) ∧ (¬ y ∨ ¬ z)) ≈⟨ proj₁ ∨-∧-distrib _ _ _ ⟩
- (x ∨ (y ∨ z)) ∧ (x ∨ (¬ y ∨ ¬ z)) ≈⟨ sym (∨-assoc _ _ _) ⟨ ∧-cong ⟩
- sym (∨-assoc _ _ _) ⟩
- ((x ∨ y) ∨ z) ∧ ((x ∨ ¬ y) ∨ ¬ z) ∎
-
- lem₄' = begin
- ¬ ((y ∨ z) ∧ ¬ (y ∧ z)) ≈⟨ deMorgan₁ _ _ ⟩
- ¬ (y ∨ z) ∨ ¬ ¬ (y ∧ z) ≈⟨ deMorgan₂ _ _ ⟨ ∨-cong ⟩ ¬-involutive _ ⟩
- (¬ y ∧ ¬ z) ∨ (y ∧ z) ≈⟨ lemma₂ _ _ _ _ ⟩
- ((¬ y ∨ y) ∧ (¬ z ∨ y)) ∧
- ((¬ y ∨ z) ∧ (¬ z ∨ z)) ≈⟨ (proj₁ ∨-complement _ ⟨ ∧-cong ⟩ ∨-comm _ _)
- ⟨ ∧-cong ⟩
- (refl ⟨ ∧-cong ⟩ proj₁ ∨-complement _) ⟩
- (⊤ ∧ (y ∨ ¬ z)) ∧
- ((¬ y ∨ z) ∧ ⊤) ≈⟨ proj₁ ∧-identity _ ⟨ ∧-cong ⟩
- proj₂ ∧-identity _ ⟩
- (y ∨ ¬ z) ∧ (¬ y ∨ z) ∎
-
- lem₄ = begin
- ¬ (x ∧ ((y ∨ z) ∧ ¬ (y ∧ z))) ≈⟨ deMorgan₁ _ _ ⟩
- ¬ x ∨ ¬ ((y ∨ z) ∧ ¬ (y ∧ z)) ≈⟨ refl ⟨ ∨-cong ⟩ lem₄' ⟩
- ¬ x ∨ ((y ∨ ¬ z) ∧ (¬ y ∨ z)) ≈⟨ proj₁ ∨-∧-distrib _ _ _ ⟩
- (¬ x ∨ (y ∨ ¬ z)) ∧
- (¬ x ∨ (¬ y ∨ z)) ≈⟨ sym (∨-assoc _ _ _) ⟨ ∧-cong ⟩
- sym (∨-assoc _ _ _) ⟩
- ((¬ x ∨ y) ∨ ¬ z) ∧
- ((¬ x ∨ ¬ y) ∨ z) ≈⟨ ∧-comm _ _ ⟩
- ((¬ x ∨ ¬ y) ∨ z) ∧
- ((¬ x ∨ y) ∨ ¬ z) ∎
-
- lem₅ = begin
- ((x ∨ ¬ y) ∨ ¬ z) ∧
- (((¬ x ∨ ¬ y) ∨ z) ∧ ((¬ x ∨ y) ∨ ¬ z)) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
- (((x ∨ ¬ y) ∨ ¬ z) ∧ ((¬ x ∨ ¬ y) ∨ z)) ∧
- ((¬ x ∨ y) ∨ ¬ z) ≈⟨ ∧-comm _ _ ⟨ ∧-cong ⟩ refl ⟩
- (((¬ x ∨ ¬ y) ∨ z) ∧ ((x ∨ ¬ y) ∨ ¬ z)) ∧
- ((¬ x ∨ y) ∨ ¬ z) ≈⟨ ∧-assoc _ _ _ ⟩
- ((¬ x ∨ ¬ y) ∨ z) ∧
- (((x ∨ ¬ y) ∨ ¬ z) ∧ ((¬ x ∨ y) ∨ ¬ z)) ∎
-
- isCommutativeRing : IsCommutativeRing _≈_ _⊕_ _∧_ id ⊥ ⊤
- isCommutativeRing = record
- { isRing = record
- { +-isAbelianGroup = record
- { isGroup = record
- { isMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ⊕-assoc
- ; ∙-cong = ⊕-cong
- }
- ; identity = ⊕-identity
- }
- ; inverse = ⊕-inverse
- ; ⁻¹-cong = id
- }
- ; comm = ⊕-comm
- }
- ; *-isMonoid = record
- { isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = ∧-assoc
- ; ∙-cong = ∧-cong
- }
- ; identity = ∧-identity
- }
- ; distrib = distrib-∧-⊕
- }
- ; *-comm = ∧-comm
- }
-
- commutativeRing : CommutativeRing _ _
- commutativeRing = record
- { _+_ = _⊕_
- ; _*_ = _∧_
- ; -_ = id
- ; 0# = ⊥
- ; 1# = ⊤
- ; isCommutativeRing = isCommutativeRing
- }
-
-infixl 6 _⊕_
-
-_⊕_ : Op₂ Carrier
-x ⊕ y = (x ∨ y) ∧ ¬ (x ∧ y)
-
-module DefaultXorRing = XorRing _⊕_ (λ _ _ → refl)
+open import Algebra.Properties.BooleanAlgebra B public
diff --git a/src/Algebra/Props/DistributiveLattice.agda b/src/Algebra/Props/DistributiveLattice.agda
index 37d8f98..0a1646e 100644
--- a/src/Algebra/Props/DistributiveLattice.agda
+++ b/src/Algebra/Props/DistributiveLattice.agda
@@ -1,7 +1,8 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Some derivable properties
+-- Compatibility module. Pending for removal. Use
+-- Algebra.Properties.DistributiveLattice instead.
------------------------------------------------------------------------
open import Algebra
@@ -10,76 +11,4 @@ module Algebra.Props.DistributiveLattice
{dl₁ dl₂} (DL : DistributiveLattice dl₁ dl₂)
where
-open DistributiveLattice DL
-import Algebra.Props.Lattice
-private
- open module L = Algebra.Props.Lattice lattice public
- hiding (replace-equality)
-open import Algebra.Structures
-import Algebra.FunctionProperties as P; open P _≈_
-open import Relation.Binary
-import Relation.Binary.EqReasoning as EqR; open EqR setoid
-open import Function
-open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence using (_⇔_; module Equivalence)
-open import Data.Product
-
-∨-∧-distrib : _∨_ DistributesOver _∧_
-∨-∧-distrib = ∨-∧-distribˡ , ∨-∧-distribʳ
- where
- ∨-∧-distribˡ : _∨_ DistributesOverˡ _∧_
- ∨-∧-distribˡ x y z = begin
- x ∨ y ∧ z ≈⟨ ∨-comm _ _ ⟩
- y ∧ z ∨ x ≈⟨ ∨-∧-distribʳ _ _ _ ⟩
- (y ∨ x) ∧ (z ∨ x) ≈⟨ ∨-comm _ _ ⟨ ∧-cong ⟩ ∨-comm _ _ ⟩
- (x ∨ y) ∧ (x ∨ z) ∎
-
-∧-∨-distrib : _∧_ DistributesOver _∨_
-∧-∨-distrib = ∧-∨-distribˡ , ∧-∨-distribʳ
- where
- ∧-∨-distribˡ : _∧_ DistributesOverˡ _∨_
- ∧-∨-distribˡ x y z = begin
- x ∧ (y ∨ z) ≈⟨ sym (proj₂ absorptive _ _) ⟨ ∧-cong ⟩ refl ⟩
- (x ∧ (x ∨ y)) ∧ (y ∨ z) ≈⟨ (refl ⟨ ∧-cong ⟩ ∨-comm _ _) ⟨ ∧-cong ⟩ refl ⟩
- (x ∧ (y ∨ x)) ∧ (y ∨ z) ≈⟨ ∧-assoc _ _ _ ⟩
- x ∧ ((y ∨ x) ∧ (y ∨ z)) ≈⟨ refl ⟨ ∧-cong ⟩ sym (proj₁ ∨-∧-distrib _ _ _) ⟩
- x ∧ (y ∨ x ∧ z) ≈⟨ sym (proj₁ absorptive _ _) ⟨ ∧-cong ⟩ refl ⟩
- (x ∨ x ∧ z) ∧ (y ∨ x ∧ z) ≈⟨ sym $ proj₂ ∨-∧-distrib _ _ _ ⟩
- x ∧ y ∨ x ∧ z ∎
-
- ∧-∨-distribʳ : _∧_ DistributesOverʳ _∨_
- ∧-∨-distribʳ x y z = begin
- (y ∨ z) ∧ x ≈⟨ ∧-comm _ _ ⟩
- x ∧ (y ∨ z) ≈⟨ ∧-∨-distribˡ _ _ _ ⟩
- x ∧ y ∨ x ∧ z ≈⟨ ∧-comm _ _ ⟨ ∨-cong ⟩ ∧-comm _ _ ⟩
- y ∧ x ∨ z ∧ x ∎
-
--- The dual construction is also a distributive lattice.
-
-∧-∨-isDistributiveLattice : IsDistributiveLattice _≈_ _∧_ _∨_
-∧-∨-isDistributiveLattice = record
- { isLattice = ∧-∨-isLattice
- ; ∨-∧-distribʳ = proj₂ ∧-∨-distrib
- }
-
-∧-∨-distributiveLattice : DistributiveLattice _ _
-∧-∨-distributiveLattice = record
- { _∧_ = _∨_
- ; _∨_ = _∧_
- ; isDistributiveLattice = ∧-∨-isDistributiveLattice
- }
-
--- One can replace the underlying equality with an equivalent one.
-
-replace-equality :
- {_≈′_ : Rel Carrier dl₂} →
- (∀ {x y} → x ≈ y ⇔ x ≈′ y) → DistributiveLattice _ _
-replace-equality {_≈′_} ≈⇔≈′ = record
- { _≈_ = _≈′_
- ; _∧_ = _∧_
- ; _∨_ = _∨_
- ; isDistributiveLattice = record
- { isLattice = Lattice.isLattice (L.replace-equality ≈⇔≈′)
- ; ∨-∧-distribʳ = λ x y z → to ⟨$⟩ ∨-∧-distribʳ x y z
- }
- } where open module E {x y} = Equivalence (≈⇔≈′ {x} {y})
+open import Algebra.Properties.DistributiveLattice DL public
diff --git a/src/Algebra/Props/Group.agda b/src/Algebra/Props/Group.agda
index 31283db..4c5bfb3 100644
--- a/src/Algebra/Props/Group.agda
+++ b/src/Algebra/Props/Group.agda
@@ -1,70 +1,12 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Some derivable properties
+-- Compatibility module. Pending for removal. Use
+-- Algebra.Properties.Group instead.
------------------------------------------------------------------------
open import Algebra
module Algebra.Props.Group {g₁ g₂} (G : Group g₁ g₂) where
-open Group G
-import Algebra.FunctionProperties as P; open P _≈_
-import Relation.Binary.EqReasoning as EqR; open EqR setoid
-open import Function
-open import Data.Product
-
-⁻¹-involutive : ∀ x → x ⁻¹ ⁻¹ ≈ x
-⁻¹-involutive x = begin
- x ⁻¹ ⁻¹ ≈⟨ sym $ proj₂ identity _ ⟩
- x ⁻¹ ⁻¹ ∙ ε ≈⟨ refl ⟨ ∙-cong ⟩ sym (proj₁ inverse _) ⟩
- x ⁻¹ ⁻¹ ∙ (x ⁻¹ ∙ x) ≈⟨ sym $ assoc _ _ _ ⟩
- x ⁻¹ ⁻¹ ∙ x ⁻¹ ∙ x ≈⟨ proj₁ inverse _ ⟨ ∙-cong ⟩ refl ⟩
- ε ∙ x ≈⟨ proj₁ identity _ ⟩
- x ∎
-
-private
-
- left-helper : ∀ x y → x ≈ (x ∙ y) ∙ y ⁻¹
- left-helper x y = begin
- x ≈⟨ sym (proj₂ identity x) ⟩
- x ∙ ε ≈⟨ refl ⟨ ∙-cong ⟩ sym (proj₂ inverse y) ⟩
- x ∙ (y ∙ y ⁻¹) ≈⟨ sym (assoc x y (y ⁻¹)) ⟩
- (x ∙ y) ∙ y ⁻¹ ∎
-
- right-helper : ∀ x y → y ≈ x ⁻¹ ∙ (x ∙ y)
- right-helper x y = begin
- y ≈⟨ sym (proj₁ identity y) ⟩
- ε ∙ y ≈⟨ sym (proj₁ inverse x) ⟨ ∙-cong ⟩ refl ⟩
- (x ⁻¹ ∙ x) ∙ y ≈⟨ assoc (x ⁻¹) x y ⟩
- x ⁻¹ ∙ (x ∙ y) ∎
-
-left-identity-unique : ∀ x y → x ∙ y ≈ y → x ≈ ε
-left-identity-unique x y eq = begin
- x ≈⟨ left-helper x y ⟩
- (x ∙ y) ∙ y ⁻¹ ≈⟨ eq ⟨ ∙-cong ⟩ refl ⟩
- y ∙ y ⁻¹ ≈⟨ proj₂ inverse y ⟩
- ε ∎
-
-right-identity-unique : ∀ x y → x ∙ y ≈ x → y ≈ ε
-right-identity-unique x y eq = begin
- y ≈⟨ right-helper x y ⟩
- x ⁻¹ ∙ (x ∙ y) ≈⟨ refl ⟨ ∙-cong ⟩ eq ⟩
- x ⁻¹ ∙ x ≈⟨ proj₁ inverse x ⟩
- ε ∎
-
-identity-unique : ∀ {x} → Identity x _∙_ → x ≈ ε
-identity-unique {x} id = left-identity-unique x x (proj₂ id x)
-
-left-inverse-unique : ∀ x y → x ∙ y ≈ ε → x ≈ y ⁻¹
-left-inverse-unique x y eq = begin
- x ≈⟨ left-helper x y ⟩
- (x ∙ y) ∙ y ⁻¹ ≈⟨ eq ⟨ ∙-cong ⟩ refl ⟩
- ε ∙ y ⁻¹ ≈⟨ proj₁ identity (y ⁻¹) ⟩
- y ⁻¹ ∎
-
-right-inverse-unique : ∀ x y → x ∙ y ≈ ε → y ≈ x ⁻¹
-right-inverse-unique x y eq = begin
- y ≈⟨ sym (⁻¹-involutive y) ⟩
- y ⁻¹ ⁻¹ ≈⟨ ⁻¹-cong (sym (left-inverse-unique x y eq)) ⟩
- x ⁻¹ ∎
+open import Algebra.Properties.Group G public
diff --git a/src/Algebra/Props/Lattice.agda b/src/Algebra/Props/Lattice.agda
index 4f9f142..2118664 100644
--- a/src/Algebra/Props/Lattice.agda
+++ b/src/Algebra/Props/Lattice.agda
@@ -1,106 +1,12 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Some derivable properties
+-- Compatibility module. Pending for removal. Use
+-- Algebra.Properties.Lattice instead.
------------------------------------------------------------------------
open import Algebra
module Algebra.Props.Lattice {l₁ l₂} (L : Lattice l₁ l₂) where
-open Lattice L
-open import Algebra.Structures
-import Algebra.FunctionProperties as P; open P _≈_
-open import Relation.Binary
-import Relation.Binary.EqReasoning as EqR; open EqR setoid
-open import Function
-open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence using (_⇔_; module Equivalence)
-open import Data.Product
-
-∧-idempotent : Idempotent _∧_
-∧-idempotent x = begin
- x ∧ x ≈⟨ refl ⟨ ∧-cong ⟩ sym (proj₁ absorptive _ _) ⟩
- x ∧ (x ∨ x ∧ x) ≈⟨ proj₂ absorptive _ _ ⟩
- x ∎
-
-∨-idempotent : Idempotent _∨_
-∨-idempotent x = begin
- x ∨ x ≈⟨ refl ⟨ ∨-cong ⟩ sym (∧-idempotent _) ⟩
- x ∨ x ∧ x ≈⟨ proj₁ absorptive _ _ ⟩
- x ∎
-
--- The dual construction is also a lattice.
-
-∧-∨-isLattice : IsLattice _≈_ _∧_ _∨_
-∧-∨-isLattice = record
- { isEquivalence = isEquivalence
- ; ∨-comm = ∧-comm
- ; ∨-assoc = ∧-assoc
- ; ∨-cong = ∧-cong
- ; ∧-comm = ∨-comm
- ; ∧-assoc = ∨-assoc
- ; ∧-cong = ∨-cong
- ; absorptive = swap absorptive
- }
-
-∧-∨-lattice : Lattice _ _
-∧-∨-lattice = record
- { _∧_ = _∨_
- ; _∨_ = _∧_
- ; isLattice = ∧-∨-isLattice
- }
-
--- Every lattice can be turned into a poset.
-
-poset : Poset _ _ _
-poset = record
- { Carrier = Carrier
- ; _≈_ = _≈_
- ; _≤_ = λ x y → x ≈ x ∧ y
- ; isPartialOrder = record
- { isPreorder = record
- { isEquivalence = isEquivalence
- ; reflexive = λ {i} {j} i≈j → begin
- i ≈⟨ sym $ ∧-idempotent _ ⟩
- i ∧ i ≈⟨ ∧-cong refl i≈j ⟩
- i ∧ j ∎
- ; trans = λ {i} {j} {k} i≈i∧j j≈j∧k → begin
- i ≈⟨ i≈i∧j ⟩
- i ∧ j ≈⟨ ∧-cong refl j≈j∧k ⟩
- i ∧ (j ∧ k) ≈⟨ sym (∧-assoc _ _ _) ⟩
- (i ∧ j) ∧ k ≈⟨ ∧-cong (sym i≈i∧j) refl ⟩
- i ∧ k ∎
- }
- ; antisym = λ {x} {y} x≈x∧y y≈y∧x → begin
- x ≈⟨ x≈x∧y ⟩
- x ∧ y ≈⟨ ∧-comm _ _ ⟩
- y ∧ x ≈⟨ sym y≈y∧x ⟩
- y ∎
- }
- }
-
--- One can replace the underlying equality with an equivalent one.
-
-replace-equality : {_≈′_ : Rel Carrier l₂} →
- (∀ {x y} → x ≈ y ⇔ x ≈′ y) → Lattice _ _
-replace-equality {_≈′_} ≈⇔≈′ = record
- { _≈_ = _≈′_
- ; _∧_ = _∧_
- ; _∨_ = _∨_
- ; isLattice = record
- { isEquivalence = record
- { refl = to ⟨$⟩ refl
- ; sym = λ x≈y → to ⟨$⟩ sym (from ⟨$⟩ x≈y)
- ; trans = λ x≈y y≈z → to ⟨$⟩ trans (from ⟨$⟩ x≈y) (from ⟨$⟩ y≈z)
- }
- ; ∨-comm = λ x y → to ⟨$⟩ ∨-comm x y
- ; ∨-assoc = λ x y z → to ⟨$⟩ ∨-assoc x y z
- ; ∨-cong = λ x≈y u≈v → to ⟨$⟩ ∨-cong (from ⟨$⟩ x≈y) (from ⟨$⟩ u≈v)
- ; ∧-comm = λ x y → to ⟨$⟩ ∧-comm x y
- ; ∧-assoc = λ x y z → to ⟨$⟩ ∧-assoc x y z
- ; ∧-cong = λ x≈y u≈v → to ⟨$⟩ ∧-cong (from ⟨$⟩ x≈y) (from ⟨$⟩ u≈v)
- ; absorptive = (λ x y → to ⟨$⟩ proj₁ absorptive x y)
- , (λ x y → to ⟨$⟩ proj₂ absorptive x y)
- }
- } where open module E {x y} = Equivalence (≈⇔≈′ {x} {y})
+open import Algebra.Properties.Lattice L public
diff --git a/src/Algebra/Props/Ring.agda b/src/Algebra/Props/Ring.agda
index b48862a..a321cd8 100644
--- a/src/Algebra/Props/Ring.agda
+++ b/src/Algebra/Props/Ring.agda
@@ -1,51 +1,12 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Some derivable properties
+-- Compatibility module. Pending for removal. Use
+-- Algebra.Properties.Ring instead.
------------------------------------------------------------------------
open import Algebra
module Algebra.Props.Ring {r₁ r₂} (R : Ring r₁ r₂) where
-import Algebra.Props.AbelianGroup as AGP
-open import Data.Product
-open import Function
-import Relation.Binary.EqReasoning as EqR
-
-open Ring R
-open EqR setoid
-
-open AGP +-abelianGroup public
- renaming ( ⁻¹-involutive to -‿involutive
- ; left-identity-unique to +-left-identity-unique
- ; right-identity-unique to +-right-identity-unique
- ; identity-unique to +-identity-unique
- ; left-inverse-unique to +-left-inverse-unique
- ; right-inverse-unique to +-right-inverse-unique
- ; ⁻¹-∙-comm to -‿+-comm
- )
-
--‿*-distribˡ : ∀ x y → - x * y ≈ - (x * y)
--‿*-distribˡ x y = begin
- - x * y ≈⟨ sym $ proj₂ +-identity _ ⟩
- - x * y + 0# ≈⟨ refl ⟨ +-cong ⟩ sym (proj₂ -‿inverse _) ⟩
- - x * y + (x * y + - (x * y)) ≈⟨ sym $ +-assoc _ _ _ ⟩
- - x * y + x * y + - (x * y) ≈⟨ sym (proj₂ distrib _ _ _) ⟨ +-cong ⟩ refl ⟩
- (- x + x) * y + - (x * y) ≈⟨ (proj₁ -‿inverse _ ⟨ *-cong ⟩ refl)
- ⟨ +-cong ⟩
- refl ⟩
- 0# * y + - (x * y) ≈⟨ proj₁ zero _ ⟨ +-cong ⟩ refl ⟩
- 0# + - (x * y) ≈⟨ proj₁ +-identity _ ⟩
- - (x * y) ∎
-
--‿*-distribʳ : ∀ x y → x * - y ≈ - (x * y)
--‿*-distribʳ x y = begin
- x * - y ≈⟨ sym $ proj₁ +-identity _ ⟩
- 0# + x * - y ≈⟨ sym (proj₁ -‿inverse _) ⟨ +-cong ⟩ refl ⟩
- - (x * y) + x * y + x * - y ≈⟨ +-assoc _ _ _ ⟩
- - (x * y) + (x * y + x * - y) ≈⟨ refl ⟨ +-cong ⟩ sym (proj₁ distrib _ _ _) ⟩
- - (x * y) + x * (y + - y) ≈⟨ refl ⟨ +-cong ⟩ (refl ⟨ *-cong ⟩ proj₂ -‿inverse _) ⟩
- - (x * y) + x * 0# ≈⟨ refl ⟨ +-cong ⟩ proj₂ zero _ ⟩
- - (x * y) + 0# ≈⟨ proj₂ +-identity _ ⟩
- - (x * y) ∎
+open import Algebra.Properties.Ring R public
diff --git a/src/Algebra/RingSolver/AlmostCommutativeRing.agda b/src/Algebra/RingSolver/AlmostCommutativeRing.agda
index a18b2c5..e388d8a 100644
--- a/src/Algebra/RingSolver/AlmostCommutativeRing.agda
+++ b/src/Algebra/RingSolver/AlmostCommutativeRing.agda
@@ -126,8 +126,8 @@ fromCommutativeRing CR = record
}
where
open CommutativeRing CR
- import Algebra.Props.Ring as R; open R ring
- import Algebra.Props.AbelianGroup as AG; open AG +-abelianGroup
+ import Algebra.Properties.Ring as R; open R ring
+ import Algebra.Properties.AbelianGroup as AG; open AG +-abelianGroup
-- Commutative semirings can be viewed as almost commutative rings by
-- using identity as the "almost negation".
diff --git a/src/Category/Applicative/Predicate.agda b/src/Category/Applicative/Predicate.agda
new file mode 100644
index 0000000..21a5308
--- /dev/null
+++ b/src/Category/Applicative/Predicate.agda
@@ -0,0 +1,48 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Applicative functors on indexed sets (predicates)
+------------------------------------------------------------------------
+
+-- Note that currently the applicative functor laws are not included
+-- here.
+
+module Category.Applicative.Predicate where
+
+open import Category.Functor.Predicate
+open import Data.Product
+open import Function
+open import Level
+open import Relation.Unary
+open import Relation.Unary.PredicateTransformer using (Pt)
+
+------------------------------------------------------------------------
+
+record RawPApplicative {i ℓ} {I : Set i} (F : Pt I ℓ) :
+ Set (i ⊔ suc ℓ) where
+ infixl 4 _⊛_ _<⊛_ _⊛>_
+ infix 4 _⊗_
+
+ field
+ pure : ∀ {P} → P ⊆ F P
+ _⊛_ : ∀ {P Q} → F (P ⇒ Q) ⊆ F P ⇒ F Q
+
+ rawPFunctor : RawPFunctor F
+ rawPFunctor = record
+ { _<$>_ = λ g x → pure g ⊛ x
+ }
+
+ private
+ open module RF = RawPFunctor rawPFunctor public
+
+ _<⊛_ : ∀ {P Q} → F P ⊆ const (∀ {j} → F Q j) ⇒ F P
+ x <⊛ y = const <$> x ⊛ y
+
+ _⊛>_ : ∀ {P Q} → const (∀ {i} → F P i) ⊆ F Q ⇒ F Q
+ x ⊛> y = flip const <$> x ⊛ y
+
+ _⊗_ : ∀ {P Q} → F P ⊆ F Q ⇒ F (P ∩ Q)
+ x ⊗ y = (_,_) <$> x ⊛ y
+
+ zipWith : ∀ {P Q R} → (P ⊆ Q ⇒ R) → F P ⊆ F Q ⇒ F R
+ zipWith f x y = f <$> x ⊛ y
diff --git a/src/Category/Functor/Predicate.agda b/src/Category/Functor/Predicate.agda
new file mode 100644
index 0000000..b829ca2
--- /dev/null
+++ b/src/Category/Functor/Predicate.agda
@@ -0,0 +1,25 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Functors on indexed sets (predicates)
+------------------------------------------------------------------------
+
+-- Note that currently the functor laws are not included here.
+
+module Category.Functor.Predicate where
+
+open import Function
+open import Level
+open import Relation.Unary
+open import Relation.Unary.PredicateTransformer using (PT)
+
+record RawPFunctor {i j ℓ₁ ℓ₂} {I : Set i} {J : Set j}
+ (F : PT I J ℓ₁ ℓ₂) : Set (i ⊔ j ⊔ suc ℓ₁ ⊔ suc ℓ₂)
+ where
+ infixl 4 _<$>_ _<$_
+
+ field
+ _<$>_ : ∀ {P Q} → P ⊆ Q → F P ⊆ F Q
+
+ _<$_ : ∀ {P Q} → (∀ {i} → P i) → F Q ⊆ F P
+ x <$ y = const x <$> y
diff --git a/src/Category/Monad/Indexed.agda b/src/Category/Monad/Indexed.agda
index d286f92..1619d9c 100644
--- a/src/Category/Monad/Indexed.agda
+++ b/src/Category/Monad/Indexed.agda
@@ -14,8 +14,8 @@ open import Level
record RawIMonad {i f} {I : Set i} (M : IFun I f) :
Set (i ⊔ suc f) where
- infixl 1 _>>=_ _>>_
- infixr 1 _=<<_
+ infixl 1 _>>=_ _>>_ _>=>_
+ infixr 1 _=<<_ _<=<_
field
return : ∀ {i A} → A → M i i A
@@ -27,6 +27,14 @@ record RawIMonad {i f} {I : Set i} (M : IFun I f) :
_=<<_ : ∀ {i j k A B} → (A → M j k B) → M i j A → M i k B
f =<< c = c >>= f
+ _>=>_ : ∀ {i j k a} {A : Set a} {B C} →
+ (A → M i j B) → (B → M j k C) → (A → M i k C)
+ f >=> g = _=<<_ g ∘ f
+
+ _<=<_ : ∀ {i j k B C a} {A : Set a} →
+ (B → M j k C) → (A → M i j B) → (A → M i k C)
+ g <=< f = f >=> g
+
join : ∀ {i j k A} → M i j (M j k A) → M i k A
join m = m >>= id
diff --git a/src/Category/Monad/Predicate.agda b/src/Category/Monad/Predicate.agda
new file mode 100644
index 0000000..bbc890c
--- /dev/null
+++ b/src/Category/Monad/Predicate.agda
@@ -0,0 +1,70 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Monads on indexed sets (predicates)
+------------------------------------------------------------------------
+
+-- Note that currently the monad laws are not included here.
+
+module Category.Monad.Predicate where
+
+open import Category.Applicative.Indexed
+open import Category.Monad
+open import Category.Monad.Indexed
+open import Data.Unit
+open import Data.Product
+open import Function
+open import Level
+open import Relation.Binary.PropositionalEquality
+open import Relation.Unary
+open import Relation.Unary.PredicateTransformer using (Pt)
+
+------------------------------------------------------------------------
+
+{_}×_ : ∀ {i a} {I : Set i} → I → Set a → Pred I _
+{ i }× A = λ j → i ≡ j × A
+
+{_} : ∀ {i} {I : Set i} → I → Pred I i
+{ i } = λ j → j ∈ { i }× ⊤
+
+------------------------------------------------------------------------
+
+record RawPMonad {i ℓ} {I : Set i} (M : Pt I (i ⊔ ℓ)) :
+ Set (suc i ⊔ suc ℓ) where
+
+ infixl 1 _?>=_ _?>_ _>?>_
+ infixr 1 _=<?_ _<?<_
+
+ -- ``Demonic'' operations (the opponent chooses the state).
+
+ field
+ return? : ∀ {P} → P ⊆ M P
+ _=<?_ : ∀ {P Q} → P ⊆ M Q → M P ⊆ M Q
+
+ _?>=_ : ∀ {P Q} → M P ⊆ const (P ⊆ M Q) ⇒ M Q
+ m ?>= f = f =<? m
+
+ _?>=′_ : ∀ {P Q} → M P ⊆ const (∀ j → {_ : P j} → j ∈ M Q) ⇒ M Q
+ m ?>=′ f = m ?>= λ {j} p → f j {p}
+
+ _?>_ : ∀ {P Q} → M P ⊆ const (∀ {j} → j ∈ M Q) ⇒ M Q
+ m₁ ?> m₂ = m₁ ?>= λ _ → m₂
+
+ join? : ∀ {P} → M (M P) ⊆ M P
+ join? m = m ?>= id
+
+ _>?>_ : {P Q R : _} → P ⊆ M Q → Q ⊆ M R → P ⊆ M R
+ f >?> g = _=<?_ g ∘ f
+
+ _<?<_ : ∀ {P Q R} → Q ⊆ M R → P ⊆ M Q → P ⊆ M R
+ g <?< f = f >?> g
+
+ -- ``Angelic'' operations (the player knows the state).
+
+ rawIMonad : RawIMonad (λ i j A → i ∈ M ({ j }× A))
+ rawIMonad = record
+ { return = λ x → return? (refl , x)
+ ; _>>=_ = λ m k → m ?>= λ { {._} (refl , x) → k x }
+ }
+
+ open RawIMonad rawIMonad public
diff --git a/src/Coinduction.agda b/src/Coinduction.agda
index 39d1e87..9587fbb 100644
--- a/src/Coinduction.agda
+++ b/src/Coinduction.agda
@@ -24,7 +24,8 @@ postulate
------------------------------------------------------------------------
-- Rec, a type which is analogous to the Rec type constructor used in
--- (the current version of) ΠΣ
+-- ΠΣ (see Altenkirch, Danielsson, Löh and Oury. ΠΣ: Dependent Types
+-- without the Sugar. FLOPS 2010, LNCS 6009.)
data Rec {a} (A : ∞ (Set a)) : Set a where
fold : (x : ♭ A) → Rec A
diff --git a/src/Data/AVL.agda b/src/Data/AVL.agda
index 80f1437..937e2bc 100644
--- a/src/Data/AVL.agda
+++ b/src/Data/AVL.agda
@@ -23,7 +23,7 @@ open import Data.Bool
import Data.DifferenceList as DiffList
open import Data.Empty
open import Data.List as List using (List)
-open import Data.Maybe
+open import Data.Maybe hiding (map)
open import Data.Nat hiding (_<_; compare; _⊔_)
open import Data.Product hiding (map)
open import Data.Unit
@@ -189,7 +189,7 @@ module Indexed where
-- 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 θ).
+ -- 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)
@@ -400,7 +400,7 @@ map : ({k : Key} → Value k → Value k) → Tree → Tree
map f (tree t) = tree $ Indexed.map f t
_∈?_ : Key → Tree → Bool
-k ∈? t = maybeToBool (lookup k t)
+k ∈? t = is-just (lookup k t)
headTail : Tree → Maybe (KV × Tree)
headTail (tree (Indexed.leaf _)) = nothing
diff --git a/src/Data/AVL/IndexedMap.agda b/src/Data/AVL/IndexedMap.agda
index b19258b..fe77087 100644
--- a/src/Data/AVL/IndexedMap.agda
+++ b/src/Data/AVL/IndexedMap.agda
@@ -15,7 +15,6 @@ module Data.AVL.IndexedMap
(isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_)
where
-open import Category.Functor
import Data.AVL
open import Data.Bool
open import Data.List as List using (List)
@@ -23,8 +22,6 @@ 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 (i ⊔ k ⊔ v)
@@ -68,10 +65,10 @@ _∈?_ : ∀ {i} → Key i → Map → Bool
_∈?_ k = AVL._∈?_ (, k)
headTail : Map → Maybe (KV × Map)
-headTail m = Prod.map toKV id <$> AVL.headTail m
+headTail m = Maybe.map (Prod.map toKV id) (AVL.headTail m)
initLast : Map → Maybe (Map × KV)
-initLast m = Prod.map id toKV <$> AVL.initLast m
+initLast m = Maybe.map (Prod.map id toKV) (AVL.initLast m)
fromList : List KV → Map
fromList = AVL.fromList ∘ List.map fromKV
diff --git a/src/Data/AVL/Sets.agda b/src/Data/AVL/Sets.agda
index bf74ae9..78a6335 100644
--- a/src/Data/AVL/Sets.agda
+++ b/src/Data/AVL/Sets.agda
@@ -12,7 +12,6 @@ module Data.AVL.Sets
(isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_)
where
-open import Category.Functor
import Data.AVL as AVL
open import Data.Bool
open import Data.List as List using (List)
@@ -22,8 +21,6 @@ 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
@@ -48,10 +45,10 @@ _∈?_ : Key → ⟨Set⟩ → Bool
_∈?_ = S._∈?_
headTail : ⟨Set⟩ → Maybe (Key × ⟨Set⟩)
-headTail s = Prod.map proj₁ id <$> S.headTail s
+headTail s = Maybe.map (Prod.map proj₁ id) (S.headTail s)
initLast : ⟨Set⟩ → Maybe (⟨Set⟩ × Key)
-initLast s = Prod.map id proj₁ <$> S.initLast s
+initLast s = Maybe.map (Prod.map id proj₁) (S.initLast s)
fromList : List Key → ⟨Set⟩
fromList = S.fromList ∘ List.map (λ k → (k , _))
diff --git a/src/Data/Bin.agda b/src/Data/Bin.agda
index 3b13db1..1808726 100644
--- a/src/Data/Bin.agda
+++ b/src/Data/Bin.agda
@@ -12,7 +12,7 @@ open Nat.≤-Reasoning
import Data.Nat.Properties as NP
open import Data.Digit
open import Data.Fin as Fin using (Fin; zero) renaming (suc to 1+_)
-open import Data.Fin.Props as FP using (_+′_)
+open import Data.Fin.Properties as FP using (_+′_)
open import Data.List as List hiding (downFrom)
open import Function
open import Data.Product
diff --git a/src/Data/Bool/Properties.agda b/src/Data/Bool/Properties.agda
index a38dbd4..e420c43 100644
--- a/src/Data/Bool/Properties.agda
+++ b/src/Data/Bool/Properties.agda
@@ -243,7 +243,7 @@ private
commutativeRing-xor-∧ : CommutativeRing _ _
commutativeRing-xor-∧ = commutativeRing
where
- import Algebra.Props.BooleanAlgebra as BA
+ import Algebra.Properties.BooleanAlgebra as BA
open BA booleanAlgebra
open XorRing _xor_ xor-is-ok
@@ -279,6 +279,10 @@ T-≡ : ∀ {b} → T b ⇔ b ≡ true
T-≡ {false} = equivalence (λ ()) (λ ())
T-≡ {true} = equivalence (const refl) (const _)
+T-not-≡ : ∀ {b} → T (not b) ⇔ b ≡ false
+T-not-≡ {false} = equivalence (const refl) (const _)
+T-not-≡ {true} = equivalence (λ ()) (λ ())
+
T-∧ : ∀ {b₁ b₂} → T (b₁ ∧ b₂) ⇔ (T b₁ × T b₂)
T-∧ {true} {true} = equivalence (const (_ , _)) (const _)
T-∧ {true} {false} = equivalence (λ ()) proj₂
@@ -292,3 +296,9 @@ T-∨ {false} {false} = equivalence inj₁ [ id , id ]
proof-irrelevance : ∀ {b} (p q : T b) → p ≡ q
proof-irrelevance {true} _ _ = refl
proof-irrelevance {false} () ()
+
+push-function-into-if :
+ ∀ {a b} {A : Set a} {B : Set b} (f : A → B) x {y z} →
+ f (if x then y else z) ≡ (if x then f y else f z)
+push-function-into-if _ true = P.refl
+push-function-into-if _ false = P.refl
diff --git a/src/Data/Char.agda b/src/Data/Char.agda
index e9fd661..c149af3 100644
--- a/src/Data/Char.agda
+++ b/src/Data/Char.agda
@@ -10,6 +10,7 @@ open import Data.Nat using (ℕ)
import Data.Nat.Properties as NatProp
open import Data.Bool using (Bool; true; false)
open import Relation.Nullary
+open import Relation.Nullary.Decidable
open import Relation.Binary
import Relation.Binary.On as On
open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
@@ -35,17 +36,37 @@ private
toNat : Char → ℕ
toNat = primCharToNat
-infix 4 _==_
-
-_==_ : Char → Char → Bool
-_==_ = primCharEquality
+-- Informative equality test.
_≟_ : Decidable {A = Char} _≡_
-s₁ ≟ s₂ with s₁ == s₂
+s₁ ≟ s₂ with primCharEquality s₁ s₂
... | true = yes trustMe
... | false = no whatever
where postulate whatever : _
+-- Boolean equality test.
+--
+-- Why is the definition _==_ = primCharEquality not used? One reason
+-- is that the present definition can sometimes improve type
+-- inference, at least with the version of Agda that is current at the
+-- time of writing: see unit-test below.
+
+infix 4 _==_
+
+_==_ : Char → Char → Bool
+c₁ == c₂ = ⌊ c₁ ≟ c₂ ⌋
+
+private
+
+ -- The following unit test does not type-check (at the time of
+ -- writing) if _==_ is replaced by primCharEquality.
+
+ data P : (Char → Bool) → Set where
+ p : (c : Char) → P (_==_ c)
+
+ unit-test : P (_==_ 'x')
+ unit-test = p _
+
setoid : Setoid _ _
setoid = PropEq.setoid Char
diff --git a/src/Data/Colist.agda b/src/Data/Colist.agda
index 8363c5c..3b73f12 100644
--- a/src/Data/Colist.agda
+++ b/src/Data/Colist.agda
@@ -8,26 +8,31 @@ module Data.Colist where
open import Category.Monad
open import Coinduction
-open import Data.Bool using (Bool; true; false)
-open import Data.Empty using (⊥)
-open import Data.Maybe using (Maybe; nothing; just)
-open import Data.Nat using (ℕ; zero; suc)
-open import Data.Conat using (Coℕ; zero; suc)
-open import Data.List using (List; []; _∷_)
-open import Data.List.NonEmpty using (List⁺; _∷_)
- renaming ([_] to [_]⁺)
+open import Data.Bool using (Bool; true; false)
open import Data.BoundedVec.Inefficient as BVec
using (BoundedVec; []; _∷_)
-open import Data.Product using (_,_)
-open import Data.Sum using (_⊎_; inj₁; inj₂)
+open import Data.Conat using (Coℕ; zero; suc)
+open import Data.Empty using (⊥)
+open import Data.Maybe using (Maybe; nothing; just; Is-just)
+open import Data.Nat using (ℕ; zero; suc; _≥′_; ≤′-refl; ≤′-step)
+open import Data.Nat.Properties using (s≤′s)
+open import Data.List using (List; []; _∷_)
+open import Data.List.NonEmpty using (List⁺; _∷_)
+open import Data.Product as Prod using (∃; _×_; _,_)
+open import Data.Sum as Sum using (_⊎_; inj₁; inj₂; [_,_]′)
open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Inverse as Inv using (_↔_; module Inverse)
open import Level using (_⊔_)
open import Relation.Binary
import Relation.Binary.InducedPreorders as Ind
-open import Relation.Binary.PropositionalEquality using (_≡_)
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
open import Relation.Nullary
open import Relation.Nullary.Negation
-open RawMonad (¬¬-Monad {p = Level.zero})
+
+module ¬¬Monad {p} where
+ open RawMonad (¬¬-Monad {p}) public
+open ¬¬Monad -- we don't want the RawMonad content to be opened publicly
------------------------------------------------------------------------
-- The type
@@ -90,15 +95,117 @@ _++_ : ∀ {a} {A : Set a} → Colist A → Colist A → Colist A
[] ++ ys = ys
(x ∷ xs) ++ ys = x ∷ ♯ (♭ xs ++ ys)
+-- Interleaves the two colists (until the shorter one, if any, has
+-- been exhausted).
+
+infixr 5 _⋎_
+
+_⋎_ : ∀ {a} {A : Set a} → Colist A → Colist A → Colist A
+[] ⋎ ys = ys
+(x ∷ xs) ⋎ ys = x ∷ ♯ (ys ⋎ ♭ xs)
+
concat : ∀ {a} {A : Set a} → Colist (List⁺ A) → Colist A
-concat [] = []
-concat ([ x ]⁺ ∷ xss) = x ∷ ♯ concat (♭ xss)
-concat ((x ∷ xs) ∷ xss) = x ∷ ♯ concat (xs ∷ xss)
+concat [] = []
+concat ((x ∷ []) ∷ xss) = x ∷ ♯ concat (♭ xss)
+concat ((x ∷ (y ∷ xs)) ∷ xss) = x ∷ ♯ concat ((y ∷ xs) ∷ xss)
[_] : ∀ {a} {A : Set a} → A → Colist A
[ x ] = x ∷ ♯ []
------------------------------------------------------------------------
+-- Any lemmas
+
+-- Any lemma for map.
+
+Any-map : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
+ {f : A → B} {xs} →
+ Any P (map f xs) ↔ Any (P ∘ f) xs
+Any-map {P = P} {f} = λ {xs} → record
+ { to = P.→-to-⟶ (to xs)
+ ; from = P.→-to-⟶ (from xs)
+ ; inverse-of = record
+ { left-inverse-of = from∘to xs
+ ; right-inverse-of = to∘from xs
+ }
+ }
+ where
+ to : ∀ xs → Any P (map f xs) → Any (P ∘ f) xs
+ to [] ()
+ to (x ∷ xs) (here px) = here px
+ to (x ∷ xs) (there p) = there (to (♭ xs) p)
+
+ from : ∀ xs → Any (P ∘ f) xs → Any P (map f xs)
+ from [] ()
+ from (x ∷ xs) (here px) = here px
+ from (x ∷ xs) (there p) = there (from (♭ xs) p)
+
+ from∘to : ∀ xs (p : Any P (map f xs)) → from xs (to xs p) ≡ p
+ from∘to [] ()
+ from∘to (x ∷ xs) (here px) = P.refl
+ from∘to (x ∷ xs) (there p) = P.cong there (from∘to (♭ xs) p)
+
+ to∘from : ∀ xs (p : Any (P ∘ f) xs) → to xs (from xs p) ≡ p
+ to∘from [] ()
+ to∘from (x ∷ xs) (here px) = P.refl
+ to∘from (x ∷ xs) (there p) = P.cong there (to∘from (♭ xs) p)
+
+-- Any lemma for _⋎_. This lemma implies that every member of xs or ys
+-- is a member of xs ⋎ ys, and vice versa.
+
+Any-⋎ : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys} →
+ Any P (xs ⋎ ys) ↔ (Any P xs ⊎ Any P ys)
+Any-⋎ {P = P} = λ xs → record
+ { to = P.→-to-⟶ (to xs)
+ ; from = P.→-to-⟶ (from xs)
+ ; inverse-of = record
+ { left-inverse-of = from∘to xs
+ ; right-inverse-of = to∘from xs
+ }
+ }
+ where
+ to : ∀ xs {ys} → Any P (xs ⋎ ys) → Any P xs ⊎ Any P ys
+ to [] p = inj₂ p
+ to (x ∷ xs) (here px) = inj₁ (here px)
+ to (x ∷ xs) (there p) = [ inj₂ , inj₁ ∘ there ]′ (to _ p)
+
+ mutual
+
+ from-left : ∀ {xs ys} → Any P xs → Any P (xs ⋎ ys)
+ from-left (here px) = here px
+ from-left {ys = ys} (there p) = there (from-right ys p)
+
+ from-right : ∀ xs {ys} → Any P ys → Any P (xs ⋎ ys)
+ from-right [] p = p
+ from-right (x ∷ xs) p = there (from-left p)
+
+ from : ∀ xs {ys} → Any P xs ⊎ Any P ys → Any P (xs ⋎ ys)
+ from xs = Sum.[ from-left , from-right xs ]
+
+ from∘to : ∀ xs {ys} (p : Any P (xs ⋎ ys)) → from xs (to xs p) ≡ p
+ from∘to [] p = P.refl
+ from∘to (x ∷ xs) (here px) = P.refl
+ from∘to (x ∷ xs) {ys = ys} (there p) with to ys p | from∘to ys p
+ from∘to (x ∷ xs) {ys = ys} (there .(from-left q)) | inj₁ q | P.refl = P.refl
+ from∘to (x ∷ xs) {ys = ys} (there .(from-right ys q)) | inj₂ q | P.refl = P.refl
+
+ mutual
+
+ to∘from-left : ∀ {xs ys} (p : Any P xs) →
+ to xs {ys = ys} (from-left p) ≡ inj₁ p
+ to∘from-left (here px) = P.refl
+ to∘from-left {ys = ys} (there p)
+ rewrite to∘from-right ys p = P.refl
+
+ to∘from-right : ∀ xs {ys} (p : Any P ys) →
+ to xs (from-right xs p) ≡ inj₂ p
+ to∘from-right [] p = P.refl
+ to∘from-right (x ∷ xs) {ys = ys} p
+ rewrite to∘from-left {xs = ys} {ys = ♭ xs} p = P.refl
+
+ to∘from : ∀ xs {ys} (p : Any P xs ⊎ Any P ys) → to xs (from xs p) ≡ p
+ to∘from xs = Sum.[ to∘from-left , to∘from-right xs ]
+
+------------------------------------------------------------------------
-- Equality
-- xs ≈ ys means that xs and ys are equal.
@@ -146,6 +253,83 @@ map-cong : ∀ {a b} {A : Set a} {B : Set b}
map-cong f [] = []
map-cong f (x ∷ xs≈) = f x ∷ ♯ map-cong f (♭ xs≈)
+-- Any respects pointwise implication (for the predicate) and equality
+-- (for the colist).
+
+Any-resp :
+ ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} {xs ys} →
+ (∀ {x} → P x → Q x) → xs ≈ ys → Any P xs → Any Q ys
+Any-resp f (x ∷ xs≈) (here px) = here (f px)
+Any-resp f (x ∷ xs≈) (there p) = there (Any-resp f (♭ xs≈) p)
+
+-- Any maps pointwise isomorphic predicates and equal colists to
+-- isomorphic types.
+
+Any-cong :
+ ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} {xs ys} →
+ (∀ {x} → P x ↔ Q x) → xs ≈ ys → Any P xs ↔ Any Q ys
+Any-cong {A = A} P↔Q xs≈ys = record
+ { to = P.→-to-⟶ (Any-resp (_⟨$⟩_ (Inverse.to P↔Q)) xs≈ys)
+ ; from = P.→-to-⟶ (Any-resp (_⟨$⟩_ (Inverse.from P↔Q)) (sym xs≈ys))
+ ; inverse-of = record
+ { left-inverse-of = resp∘resp P↔Q xs≈ys (sym xs≈ys)
+ ; right-inverse-of = resp∘resp (Inv.sym P↔Q) (sym xs≈ys) xs≈ys
+ }
+ }
+ where
+ open Setoid (setoid _) using (sym)
+
+ resp∘resp : ∀ {p q} {P : A → Set p} {Q : A → Set q} {xs ys}
+ (P↔Q : ∀ {x} → P x ↔ Q x)
+ (xs≈ys : xs ≈ ys) (ys≈xs : ys ≈ xs) (p : Any P xs) →
+ Any-resp (_⟨$⟩_ (Inverse.from P↔Q)) ys≈xs
+ (Any-resp (_⟨$⟩_ (Inverse.to P↔Q)) xs≈ys p) ≡ p
+ resp∘resp P↔Q (x ∷ xs≈) (.x ∷ ys≈) (here px) =
+ P.cong here (Inverse.left-inverse-of P↔Q px)
+ resp∘resp P↔Q (x ∷ xs≈) (.x ∷ ys≈) (there p) =
+ P.cong there (resp∘resp P↔Q (♭ xs≈) (♭ ys≈) p)
+
+------------------------------------------------------------------------
+-- Indices
+
+-- Converts Any proofs to indices into the colist. The index can also
+-- be seen as the size of the proof.
+
+index : ∀ {a p} {A : Set a} {P : A → Set p} {xs} → Any P xs → ℕ
+index (here px) = zero
+index (there p) = suc (index p)
+
+-- The position returned by index is guaranteed to be within bounds.
+
+lookup-index : ∀ {a p} {A : Set a} {P : A → Set p} {xs} (p : Any P xs) →
+ Is-just (lookup (index p) xs)
+lookup-index (here px) = just _
+lookup-index (there p) = lookup-index p
+
+-- Any-resp preserves the index.
+
+index-Any-resp :
+ ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q}
+ {f : ∀ {x} → P x → Q x} {xs ys}
+ (xs≈ys : xs ≈ ys) (p : Any P xs) →
+ index (Any-resp f xs≈ys p) ≡ index p
+index-Any-resp (x ∷ xs≈) (here px) = P.refl
+index-Any-resp (x ∷ xs≈) (there p) =
+ P.cong suc (index-Any-resp (♭ xs≈) p)
+
+-- The left-to-right direction of Any-⋎ returns a proof whose size is
+-- no larger than that of the input proof.
+
+index-Any-⋎ :
+ ∀ {a p} {A : Set a} {P : A → Set p} xs {ys} (p : Any P (xs ⋎ ys)) →
+ index p ≥′ [ index , index ]′ (Inverse.to (Any-⋎ xs) ⟨$⟩ p)
+index-Any-⋎ [] p = ≤′-refl
+index-Any-⋎ (x ∷ xs) (here px) = ≤′-refl
+index-Any-⋎ (x ∷ xs) {ys = ys} (there p)
+ with Inverse.to (Any-⋎ ys) ⟨$⟩ p | index-Any-⋎ ys p
+... | inj₁ q | q≤p = ≤′-step q≤p
+... | inj₂ q | q≤p = s≤′s q≤p
+
------------------------------------------------------------------------
-- Memberships, subsets, prefixes
@@ -171,6 +355,38 @@ data _⊑_ {a} {A : Set a} : Colist A → Colist A → Set a where
[] : ∀ {ys} → [] ⊑ ys
_∷_ : ∀ x {xs ys} (p : ∞ (♭ xs ⊑ ♭ ys)) → x ∷ xs ⊑ x ∷ ys
+-- Any can be expressed using _∈_ (and vice versa).
+
+Any-∈ : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
+ Any P xs ↔ ∃ λ x → x ∈ xs × P x
+Any-∈ {P = P} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ (λ { (x , x∈xs , p) → from x∈xs p })
+ ; inverse-of = record
+ { left-inverse-of = from∘to
+ ; right-inverse-of = λ { (x , x∈xs , p) → to∘from x∈xs p }
+ }
+ }
+ where
+ to : ∀ {xs} → Any P xs → ∃ λ x → x ∈ xs × P x
+ to (here p) = _ , here P.refl , p
+ to (there p) = Prod.map id (Prod.map there id) (to p)
+
+ from : ∀ {x xs} → x ∈ xs → P x → Any P xs
+ from (here P.refl) p = here p
+ from (there x∈xs) p = there (from x∈xs p)
+
+ to∘from : ∀ {x xs} (x∈xs : x ∈ xs) (p : P x) →
+ to (from x∈xs p) ≡ (x , x∈xs , p)
+ to∘from (here P.refl) p = P.refl
+ to∘from (there x∈xs) p =
+ P.cong (Prod.map id (Prod.map there id)) (to∘from x∈xs p)
+
+ from∘to : ∀ {xs} (p : Any P xs) →
+ let (x , x∈xs , px) = to p in from x∈xs px ≡ p
+ from∘to (here _) = P.refl
+ from∘to (there p) = P.cong there (from∘to p)
+
-- Prefixes are subsets.
⊑⇒⊆ : ∀ {a} → {A : Set a} → _⊑_ {A = A} ⇒ _⊆_
@@ -266,11 +482,9 @@ not-finite-is-infinite (x ∷ xs) hyp =
x ∷ ♯ not-finite-is-infinite (♭ xs) (hyp ∘ _∷_ x)
-- Colists are either finite or infinite (classically).
---
--- TODO: Make this definition universe polymorphic.
finite-or-infinite :
- {A : Set} (xs : Colist A) → ¬ ¬ (Finite xs ⊎ Infinite xs)
+ ∀ {a} {A : Set a} (xs : Colist A) → ¬ ¬ (Finite xs ⊎ Infinite xs)
finite-or-infinite xs = helper <$> excluded-middle
where
helper : Dec (Finite xs) → Finite xs ⊎ Infinite xs
diff --git a/src/Data/Colist/Infinite-merge.agda b/src/Data/Colist/Infinite-merge.agda
new file mode 100644
index 0000000..67cde1e
--- /dev/null
+++ b/src/Data/Colist/Infinite-merge.agda
@@ -0,0 +1,231 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Infinite merge operation for coinductive lists
+------------------------------------------------------------------------
+
+module Data.Colist.Infinite-merge where
+
+open import Coinduction
+open import Data.Colist as Colist hiding (_⋎_)
+open import Data.Nat
+open import Data.Nat.Properties
+open import Data.Product as Prod
+open import Data.Sum
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Inverse as Inv using (_↔_; module Inverse)
+import Function.Related as Related
+open import Function.Related.TypeIsomorphisms
+open import Induction.Nat
+import Induction.WellFounded as WF
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+open import Relation.Binary.Sum
+
+-- Some code that is used to work around Agda's syntactic guardedness
+-- checker.
+
+private
+
+ infixr 5 _∷_ _⋎_
+
+ data ColistP {a} (A : Set a) : Set a where
+ [] : ColistP A
+ _∷_ : A → ∞ (ColistP A) → ColistP A
+ _⋎_ : ColistP A → ColistP A → ColistP A
+
+ data ColistW {a} (A : Set a) : Set a where
+ [] : ColistW A
+ _∷_ : A → ColistP A → ColistW A
+
+ program : ∀ {a} {A : Set a} → Colist A → ColistP A
+ program [] = []
+ program (x ∷ xs) = x ∷ ♯ program (♭ xs)
+
+ mutual
+
+ _⋎W_ : ∀ {a} {A : Set a} → ColistW A → ColistP A → ColistW A
+ [] ⋎W ys = whnf ys
+ (x ∷ xs) ⋎W ys = x ∷ (ys ⋎ xs)
+
+ whnf : ∀ {a} {A : Set a} → ColistP A → ColistW A
+ whnf [] = []
+ whnf (x ∷ xs) = x ∷ ♭ xs
+ whnf (xs ⋎ ys) = whnf xs ⋎W ys
+
+ mutual
+
+ ⟦_⟧P : ∀ {a} {A : Set a} → ColistP A → Colist A
+ ⟦ xs ⟧P = ⟦ whnf xs ⟧W
+
+ ⟦_⟧W : ∀ {a} {A : Set a} → ColistW A → Colist A
+ ⟦ [] ⟧W = []
+ ⟦ x ∷ xs ⟧W = x ∷ ♯ ⟦ xs ⟧P
+
+ mutual
+
+ ⋎-homP : ∀ {a} {A : Set a} (xs : ColistP A) {ys} →
+ ⟦ xs ⋎ ys ⟧P ≈ ⟦ xs ⟧P Colist.⋎ ⟦ ys ⟧P
+ ⋎-homP xs = ⋎-homW (whnf xs) _
+
+ ⋎-homW : ∀ {a} {A : Set a} (xs : ColistW A) ys →
+ ⟦ xs ⋎W ys ⟧W ≈ ⟦ xs ⟧W Colist.⋎ ⟦ ys ⟧P
+ ⋎-homW (x ∷ xs) ys = x ∷ ♯ ⋎-homP ys
+ ⋎-homW [] ys = begin ⟦ ys ⟧P ∎
+ where open ≈-Reasoning
+
+ ⟦program⟧P : ∀ {a} {A : Set a} (xs : Colist A) →
+ ⟦ program xs ⟧P ≈ xs
+ ⟦program⟧P [] = []
+ ⟦program⟧P (x ∷ xs) = x ∷ ♯ ⟦program⟧P (♭ xs)
+
+ Any-⋎P : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys} →
+ Any P ⟦ program xs ⋎ ys ⟧P ↔ (Any P xs ⊎ Any P ⟦ ys ⟧P)
+ Any-⋎P {P = P} xs {ys} =
+ Any P ⟦ program xs ⋎ ys ⟧P ↔⟨ Any-cong Inv.id (⋎-homP (program xs)) ⟩
+ Any P (⟦ program xs ⟧P Colist.⋎ ⟦ ys ⟧P) ↔⟨ Any-⋎ _ ⟩
+ (Any P ⟦ program xs ⟧P ⊎ Any P ⟦ ys ⟧P) ↔⟨ Any-cong Inv.id (⟦program⟧P _) ⊎-cong (_ ∎) ⟩
+ (Any P xs ⊎ Any P ⟦ ys ⟧P) ∎
+ where open Related.EquationalReasoning
+
+ index-Any-⋎P :
+ ∀ {a p} {A : Set a} {P : A → Set p} xs {ys}
+ (p : Any P ⟦ program xs ⋎ ys ⟧P) →
+ index p ≥′ [ index , index ]′ (Inverse.to (Any-⋎P xs) ⟨$⟩ p)
+ index-Any-⋎P xs p
+ with Any-resp id (⋎-homW (whnf (program xs)) _) p
+ | index-Any-resp {f = id} (⋎-homW (whnf (program xs)) _) p
+ index-Any-⋎P xs p | q | q≡p
+ with Inverse.to (Any-⋎ ⟦ program xs ⟧P) ⟨$⟩ q
+ | index-Any-⋎ ⟦ program xs ⟧P q
+ index-Any-⋎P xs p | q | q≡p | inj₂ r | r≤q rewrite q≡p = r≤q
+ index-Any-⋎P xs p | q | q≡p | inj₁ r | r≤q
+ with Any-resp id (⟦program⟧P xs) r
+ | index-Any-resp {f = id} (⟦program⟧P xs) r
+ index-Any-⋎P xs p | q | q≡p | inj₁ r | r≤q | s | s≡r
+ rewrite s≡r | q≡p = r≤q
+
+-- Infinite variant of _⋎_.
+
+private
+
+ merge′ : ∀ {a} {A : Set a} → Colist (A × Colist A) → ColistP A
+ merge′ [] = []
+ merge′ ((x , xs) ∷ xss) = x ∷ ♯ (program xs ⋎ merge′ (♭ xss))
+
+merge : ∀ {a} {A : Set a} → Colist (A × Colist A) → Colist A
+merge xss = ⟦ merge′ xss ⟧P
+
+-- Any lemma for merge.
+
+Any-merge :
+ ∀ {a p} {A : Set a} {P : A → Set p} xss →
+ Any P (merge xss) ↔ Any (λ { (x , xs) → P x ⊎ Any P xs }) xss
+Any-merge {P = P} = λ xss → record
+ { to = P.→-to-⟶ (proj₁ ∘ to xss)
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = proj₂ ∘ to xss
+ ; right-inverse-of = λ p → begin
+ proj₁ (to xss (from p)) ≡⟨ from-injective _ _ (proj₂ (to xss (from p))) ⟩
+ p ∎
+ }
+ }
+ where
+ open P.≡-Reasoning
+
+ -- The from function.
+
+ Q = λ { (x , xs) → P x ⊎ Any P xs }
+
+ from : ∀ {xss} → Any Q xss → Any P (merge xss)
+ from (here (inj₁ p)) = here p
+ from (here (inj₂ p)) = there (Inverse.from (Any-⋎P _) ⟨$⟩ inj₁ p)
+ from (there {x = _ , xs} p) = there (Inverse.from (Any-⋎P xs) ⟨$⟩ inj₂ (from p))
+
+ -- Some lemmas.
+
+ drop-there :
+ ∀ {a p} {A : Set a} {P : A → Set p} {x xs} {p q : Any P _} →
+ _≡_ {A = Any P (x ∷ xs)} (there p) (there q) → p ≡ q
+ drop-there P.refl = P.refl
+
+ drop-inj₁ : ∀ {a b} {A : Set a} {B : Set b} {x y} →
+ _≡_ {A = A ⊎ B} (inj₁ x) (inj₁ y) → x ≡ y
+ drop-inj₁ P.refl = P.refl
+
+ drop-inj₂ : ∀ {a b} {A : Set a} {B : Set b} {x y} →
+ _≡_ {A = A ⊎ B} (inj₂ x) (inj₂ y) → x ≡ y
+ drop-inj₂ P.refl = P.refl
+
+ -- The from function is injective.
+
+ from-injective : ∀ {xss} (p₁ p₂ : Any Q xss) →
+ from p₁ ≡ from p₂ → p₁ ≡ p₂
+ from-injective (here (inj₁ p)) (here (inj₁ .p)) P.refl = P.refl
+ from-injective (here (inj₁ _)) (here (inj₂ _)) ()
+ from-injective (here (inj₂ _)) (here (inj₁ _)) ()
+ from-injective (here (inj₂ p₁)) (here (inj₂ p₂)) eq =
+ P.cong (here ∘ inj₂) $
+ drop-inj₁ $
+ Inverse.injective (Inv.sym (Any-⋎P _)) {x = inj₁ p₁} {y = inj₁ p₂} $
+ drop-there eq
+ from-injective (here (inj₁ _)) (there _) ()
+ from-injective (here (inj₂ p₁)) (there p₂) eq
+ with Inverse.injective (Inv.sym (Any-⋎P _))
+ {x = inj₁ p₁} {y = inj₂ (from p₂)}
+ (drop-there eq)
+ ... | ()
+ from-injective (there _) (here (inj₁ _)) ()
+ from-injective (there p₁) (here (inj₂ p₂)) eq
+ with Inverse.injective (Inv.sym (Any-⋎P _))
+ {x = inj₂ (from p₁)} {y = inj₁ p₂}
+ (drop-there eq)
+ ... | ()
+ from-injective (there {x = _ , xs} p₁) (there p₂) eq =
+ P.cong there $
+ from-injective p₁ p₂ $
+ drop-inj₂ $
+ Inverse.injective (Inv.sym (Any-⋎P xs))
+ {x = inj₂ (from p₁)} {y = inj₂ (from p₂)} $
+ drop-there eq
+
+ -- The to function (defined as a right inverse of from).
+
+ Input = ∃ λ xss → Any P (merge xss)
+
+ Pred : Input → Set _
+ Pred (xss , p) = ∃ λ (q : Any Q xss) → from q ≡ p
+
+ to : ∀ xss p → Pred (xss , p)
+ to = λ xss p →
+ WF.All.wfRec (WF.Inverse-image.well-founded size <-well-founded) _
+ Pred step (xss , p)
+ where
+ size : Input → ℕ
+ size (_ , p) = index p
+
+ step : ∀ p → WF.WfRec (_<′_ on size) Pred p → Pred p
+ step ([] , ()) rec
+ step ((x , xs) ∷ xss , here p) rec = here (inj₁ p) , P.refl
+ step ((x , xs) ∷ xss , there p) rec
+ with Inverse.to (Any-⋎P xs) ⟨$⟩ p
+ | Inverse.left-inverse-of (Any-⋎P xs) p
+ | index-Any-⋎P xs p
+ step ((x , xs) ∷ xss , there .(Inverse.from (Any-⋎P xs) ⟨$⟩ inj₁ q)) rec | inj₁ q | P.refl | _ = here (inj₂ q) , P.refl
+ step ((x , xs) ∷ xss , there .(Inverse.from (Any-⋎P xs) ⟨$⟩ inj₂ q)) rec | inj₂ q | P.refl | q≤p =
+ Prod.map there
+ (P.cong (there ∘ _⟨$⟩_ (Inverse.from (Any-⋎P xs)) ∘ inj₂))
+ (rec (♭ xss , q) (s≤′s q≤p))
+
+-- Every member of xss is a member of merge xss, and vice versa (with
+-- equal multiplicities).
+
+∈-merge : ∀ {a} {A : Set a} {y : A} xss →
+ y ∈ merge xss ↔ ∃₂ λ x xs → (x , xs) ∈ xss × (y ≡ x ⊎ y ∈ xs)
+∈-merge {y = y} xss =
+ y ∈ merge xss ↔⟨ Any-merge _ ⟩
+ Any (λ { (x , xs) → y ≡ x ⊎ y ∈ xs }) xss ↔⟨ Any-∈ ⟩
+ (∃ λ { (x , xs) → (x , xs) ∈ xss × (y ≡ x ⊎ y ∈ xs) }) ↔⟨ Σ-assoc ⟩
+ (∃₂ λ x xs → (x , xs) ∈ xss × (y ≡ x ⊎ y ∈ xs)) ∎
+ where open Related.EquationalReasoning
diff --git a/src/Data/Container/Any.agda b/src/Data/Container/Any.agda
index 64a6943..4f52fe2 100644
--- a/src/Data/Container/Any.agda
+++ b/src/Data/Container/Any.agda
@@ -67,7 +67,7 @@ cong {C = C} {P₁ = P₁} {P₂} {xs₁} {xs₂} P₁↔P₂ xs₁≈xs₂ =
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
+ ◈ = λ {_} {_} → flip ◇ in
◈ xs (◈ ys ∘ P) ↔ ◈ ys (◈ xs ∘ flip P)
swap {c} {C₁} {C₂} {P = P} {xs} {ys} =
◇ (λ x → ◇ (P x) ys) xs ↔⟨ ↔∈ C₁ ⟩
@@ -251,7 +251,7 @@ linear-identity {xs = xs} m {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′ ⟫⊸ ∘
+ join = λ {_} → ⟪ join′ ⟫⊸ ∘
_⟨$⟩_ (Inverse.from (Composition.correct C₁ C₂)) in
◇ P (join xss) ↔ ◇ (◇ P) xss
join↔◇ {C₁ = C₁} {C₂} P join xss =
diff --git a/src/Data/Container/FreeMonad.agda b/src/Data/Container/FreeMonad.agda
new file mode 100644
index 0000000..f3e8a25
--- /dev/null
+++ b/src/Data/Container/FreeMonad.agda
@@ -0,0 +1,55 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- The free monad construction on containers
+------------------------------------------------------------------------
+
+module Data.Container.FreeMonad where
+
+open import Level
+open import Function using (_∘_)
+open import Data.Empty using (⊥-elim)
+open import Data.Sum using (inj₁; inj₂)
+open import Data.Product
+open import Data.Container
+open import Data.Container.Combinator using (const; _⊎_)
+open import Data.W
+open import Category.Monad
+
+infixl 1 _⋆C_
+infix 1 _⋆_
+
+------------------------------------------------------------------------
+
+-- The free monad construction over a container and a set is, in
+-- universal algebra terminology, also known as the term algebra over a
+-- signature (a container) and a set (of variable symbols). The return
+-- of the free monad corresponds to variables and the bind operator
+-- corresponds to (parallel) substitution.
+
+-- A useful intuition is to think of containers describing single
+-- operations and the free monad construction over a container and a set
+-- describing a tree of operations as nodes and elements of the set as
+-- leafs. If one starts at the root, then any path will pass finitely
+-- many nodes (operations described by the container) and eventually end
+-- up in a leaf (element of the set) -- hence the Kleene star notation
+-- (the type can be read as a regular expression).
+
+_⋆C_ : ∀ {c} → Container c → Set c → Container c
+C ⋆C X = const X ⊎ C
+
+_⋆_ : ∀ {c} → Container c → Set c → Set c
+C ⋆ X = μ (C ⋆C X)
+
+do : ∀ {c} {C : Container c} {X} → ⟦ C ⟧ (C ⋆ X) → C ⋆ X
+do (s , k) = sup (inj₂ s) k
+
+rawMonad : ∀ {c} {C : Container c} → RawMonad (_⋆_ C)
+rawMonad = record { return = return; _>>=_ = _>>=_ }
+ where
+ return : ∀ {c} {C : Container c} {X} → X → C ⋆ X
+ return x = sup (inj₁ x) (⊥-elim ∘ lower)
+
+ _>>=_ : ∀ {c} {C : Container c} {X Y} → C ⋆ X → (X → C ⋆ Y) → C ⋆ Y
+ sup (inj₁ x) _ >>= k = k x
+ sup (inj₂ s) f >>= k = do (s , λ p → f p >>= k)
diff --git a/src/Data/Container/Indexed.agda b/src/Data/Container/Indexed.agda
new file mode 100644
index 0000000..cd1ab19
--- /dev/null
+++ b/src/Data/Container/Indexed.agda
@@ -0,0 +1,351 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Indexed containers aka interaction structures aka polynomial
+-- functors. The notation and presentation here is closest to that of
+-- Hancock and Hyvernat in "Programming interfaces and basic topology"
+-- (2006/9).
+--
+------------------------------------------------------------------------
+
+module Data.Container.Indexed where
+
+open import Level
+open import Function renaming (id to ⟨id⟩; _∘_ to _⟨∘⟩_)
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Inverse using (_↔_; module Inverse)
+open import Data.Product as Prod hiding (map)
+open import Relation.Unary using (Pred; _⊆_)
+open import Relation.Binary as B using (Preorder; module Preorder)
+open import Relation.Binary.PropositionalEquality as P using (_≡_; _≗_; refl)
+open import Relation.Binary.HeterogeneousEquality as H using (_≅_; refl)
+open import Relation.Binary.Indexed
+open import Data.W.Indexed
+open import Data.M.Indexed
+
+------------------------------------------------------------------------
+
+-- The type and its semantics ("extension").
+
+open import Data.Container.Indexed.Core public
+open Container public
+
+-- Abbreviation for the commonly used level one version of indexed
+-- containers.
+
+_▷_ : Set → Set → Set₁
+I ▷ O = Container I O zero zero
+
+-- The least and greatest fixpoint.
+
+μ ν : ∀ {o c r} {O : Set o} → Container O O c r → Pred O _
+μ = W
+ν = M
+
+-- Equality, parametrised on an underlying relation.
+
+Eq : ∀ {i o c r ℓ} {I : Set i} {O : Set o} (C : Container I O c r)
+ (X Y : Pred I ℓ) → REL X Y ℓ → REL (⟦ C ⟧ X) (⟦ C ⟧ Y) _
+Eq C _ _ _≈_ {o₁} {o₂} (c , k) (c′ , k′) =
+ o₁ ≡ o₂ × c ≅ c′ × (∀ r r′ → r ≅ r′ → k r ≈ k′ r′)
+
+private
+
+ -- Note that, if propositional equality were extensional, then Eq _≅_
+ -- and _≅_ would coincide.
+
+ Eq⇒≅ : ∀ {i o c r ℓ} {I : Set i} {O : Set o}
+ {C : Container I O c r} {X : Pred I ℓ} {o₁ o₂ : O}
+ {xs : ⟦ C ⟧ X o₁} {ys : ⟦ C ⟧ X o₂} → H.Extensionality r ℓ →
+ Eq C X X (λ x₁ x₂ → x₁ ≅ x₂) xs ys → xs ≅ ys
+ Eq⇒≅ {xs = c , k} {.c , k′} ext (refl , refl , k≈k′) =
+ H.cong (_,_ c) (ext (λ _ → refl) (λ r → k≈k′ r r refl))
+
+setoid : ∀ {i o c r} {I : Set i} {O : Set o} →
+ Container I O c r → Setoid I (c ⊔ r) _ → Setoid O _ _
+setoid C X = record
+ { Carrier = ⟦ C ⟧ X.Carrier
+ ; _≈_ = _≈_
+ ; isEquivalence = record
+ { refl = refl , refl , λ { r .r refl → X.refl }
+ ; sym = sym
+ ; trans = λ { {_} {i = xs} {ys} {zs} → trans {_} {i = xs} {ys} {zs} }
+ }
+ }
+ where
+ module X = Setoid X
+
+ _≈_ : Rel (⟦ C ⟧ X.Carrier) _
+ _≈_ = Eq C X.Carrier X.Carrier X._≈_
+
+ sym : Symmetric (⟦ C ⟧ X.Carrier) _≈_
+ sym {_} {._} {_ , _} {._ , _} (refl , refl , k) =
+ refl , refl , λ { r .r refl → X.sym (k r r refl) }
+
+ trans : Transitive (⟦ C ⟧ X.Carrier) _≈_
+ trans {._} {_} {._} {_ , _} {._ , _} {._ , _}
+ (refl , refl , k) (refl , refl , k′) =
+ refl , refl , λ { r .r refl → X.trans (k r r refl) (k′ r r refl) }
+
+------------------------------------------------------------------------
+-- Functoriality
+
+-- Indexed containers are functors.
+
+map : ∀ {i o c r ℓ₁ ℓ₂} {I : Set i} {O : Set o}
+ (C : Container I O c r) {X : Pred I ℓ₁} {Y : Pred I ℓ₂} →
+ X ⊆ Y → ⟦ C ⟧ X ⊆ ⟦ C ⟧ Y
+map _ f = Prod.map ⟨id⟩ (λ g → f ⟨∘⟩ g)
+
+module Map where
+
+ identity : ∀ {i o c r} {I : Set i} {O : Set o} (C : Container I O c r)
+ (X : Setoid I _ _) → let module X = Setoid X in
+ ∀ {o} {xs : ⟦ C ⟧ X.Carrier o} → Eq C X.Carrier X.Carrier
+ X._≈_ xs (map C {X.Carrier} ⟨id⟩ xs)
+ identity C X = Setoid.refl (setoid C X)
+
+ composition : ∀ {i o c r ℓ₁ ℓ₂} {I : Set i} {O : Set o}
+ (C : Container I O c r) {X : Pred I ℓ₁} {Y : Pred I ℓ₂}
+ (Z : Setoid I _ _) → let module Z = Setoid Z in
+ {f : Y ⊆ Z.Carrier} {g : X ⊆ Y} {o : O} {xs : ⟦ C ⟧ X o} →
+ Eq C Z.Carrier Z.Carrier Z._≈_
+ (map C {Y} f (map C {X} g xs))
+ (map C {X} (f ⟨∘⟩ g) xs)
+ composition C Z = Setoid.refl (setoid C Z)
+
+------------------------------------------------------------------------
+-- Container morphisms
+
+module _ {i₁ i₂ o₁ o₂}
+ {I₁ : Set i₁} {I₂ : Set i₂} {O₁ : Set o₁} {O₂ : Set o₂} where
+
+ -- General container morphism.
+
+ record ContainerMorphism {c₁ c₂ r₁ r₂ ℓ₁ ℓ₂}
+ (C₁ : Container I₁ O₁ c₁ r₁) (C₂ : Container I₂ O₂ c₂ r₂)
+ (f : I₁ → I₂) (g : O₁ → O₂)
+ (_∼_ : B.Rel I₂ ℓ₁) (_≈_ : B.REL (Set r₂) (Set r₁) ℓ₂)
+ (_·_ : ∀ {A B} → A ≈ B → A → B) :
+ Set (i₁ ⊔ i₂ ⊔ o₁ ⊔ o₂ ⊔ c₁ ⊔ c₂ ⊔ r₁ ⊔ r₂ ⊔ ℓ₁ ⊔ ℓ₂) where
+ field
+ command : Command C₁ ⊆ Command C₂ ⟨∘⟩ g
+ response : ∀ {o} {c₁ : Command C₁ o} →
+ Response C₂ (command c₁) ≈ Response C₁ c₁
+ coherent : ∀ {o} {c₁ : Command C₁ o} {r₂ : Response C₂ (command c₁)} →
+ f (next C₁ c₁ (response · r₂)) ∼ next C₂ (command c₁) r₂
+
+ open ContainerMorphism public
+
+ -- Plain container morphism.
+
+ _⇒[_/_]_ : ∀ {c₁ c₂ r₁ r₂} →
+ Container I₁ O₁ c₁ r₁ → (I₁ → I₂) → (O₁ → O₂) →
+ Container I₂ O₂ c₂ r₂ → Set _
+ C₁ ⇒[ f / g ] C₂ = ContainerMorphism C₁ C₂ f g _≡_ (λ R₂ R₁ → R₂ → R₁) _$_
+
+ -- Linear container morphism.
+
+ _⊸[_/_]_ : ∀ {c₁ c₂ r₁ r₂} →
+ Container I₁ O₁ c₁ r₁ → (I₁ → I₂) → (O₁ → O₂) →
+ Container I₂ O₂ c₂ r₂ → Set _
+ C₁ ⊸[ f / g ] C₂ = ContainerMorphism C₁ C₂ f g _≡_ _↔_
+ (λ r₂↔r₁ r₂ → Inverse.to r₂↔r₁ ⟨$⟩ r₂)
+
+ -- Cartesian container morphism.
+
+ _⇒C[_/_]_ : ∀ {c₁ c₂ r} →
+ Container I₁ O₁ c₁ r → (I₁ → I₂) → (O₁ → O₂) →
+ Container I₂ O₂ c₂ r → Set _
+ C₁ ⇒C[ f / g ] C₂ = ContainerMorphism C₁ C₂ f g _≡_ (λ R₂ R₁ → R₂ ≡ R₁)
+ (λ r₂≡r₁ r₂ → P.subst ⟨id⟩ r₂≡r₁ r₂)
+
+-- Degenerate cases where no reindexing is performed.
+
+module _ {i o c r} {I : Set i} {O : Set o} where
+
+ _⇒_ : B.Rel (Container I O c r) _
+ C₁ ⇒ C₂ = C₁ ⇒[ ⟨id⟩ / ⟨id⟩ ] C₂
+
+ _⊸_ : B.Rel (Container I O c r) _
+ C₁ ⊸ C₂ = C₁ ⊸[ ⟨id⟩ / ⟨id⟩ ] C₂
+
+ _⇒C_ : B.Rel (Container I O c r) _
+ C₁ ⇒C C₂ = C₁ ⇒C[ ⟨id⟩ / ⟨id⟩ ] C₂
+
+------------------------------------------------------------------------
+-- Plain morphisms
+
+-- Interpretation of _⇒_.
+
+⟪_⟫ : ∀ {i o c r ℓ} {I : Set i} {O : Set o} {C₁ C₂ : Container I O c r} →
+ C₁ ⇒ C₂ → (X : Pred I ℓ) → ⟦ C₁ ⟧ X ⊆ ⟦ C₂ ⟧ X
+⟪ m ⟫ X (c , k) = command m c , λ r₂ →
+ P.subst X (coherent m) (k (response m r₂))
+
+module PlainMorphism {i o c r} {I : Set i} {O : Set o} where
+
+ -- Naturality.
+
+ Natural : ∀ {ℓ} {C₁ C₂ : Container I O c r} →
+ ((X : Pred I ℓ) → ⟦ C₁ ⟧ X ⊆ ⟦ C₂ ⟧ X) → Set _
+ Natural {C₁ = C₁} {C₂} m =
+ ∀ {X} Y → let module Y = Setoid Y in (f : X ⊆ Y.Carrier) →
+ ∀ {o} (xs : ⟦ C₁ ⟧ X o) →
+ Eq C₂ Y.Carrier Y.Carrier Y._≈_
+ (m Y.Carrier $ map C₁ {X} f xs) (map C₂ {X} f $ m X xs)
+
+ -- Natural transformations.
+
+ NT : ∀ {ℓ} (C₁ C₂ : Container I O c r) → Set _
+ NT {ℓ} C₁ C₂ = ∃ λ (m : (X : Pred I ℓ) → ⟦ C₁ ⟧ X ⊆ ⟦ C₂ ⟧ X) →
+ Natural m
+
+ -- Container morphisms are natural.
+
+ natural : ∀ {ℓ} (C₁ C₂ : Container I O c r) (m : C₁ ⇒ C₂) → Natural {ℓ} ⟪ m ⟫
+ natural _ _ m {X} Y f _ = refl , refl , λ { r .r refl → lemma (coherent m) }
+ where
+ module Y = Setoid Y
+
+ lemma : ∀ {i j} (eq : i ≡ j) {x} →
+ P.subst Y.Carrier eq (f x) Y.≈ f (P.subst X eq x)
+ lemma refl = Y.refl
+
+ -- In fact, all natural functions of the right type are container
+ -- morphisms.
+
+ complete : ∀ {C₁ C₂ : Container I O c r} (nt : NT C₁ C₂) →
+ ∃ λ m → (X : Setoid I _ _) →
+ let module X = Setoid X in
+ ∀ {o} (xs : ⟦ C₁ ⟧ X.Carrier o) →
+ Eq C₂ X.Carrier X.Carrier X._≈_
+ (proj₁ nt X.Carrier xs) (⟪ m ⟫ X.Carrier {o} xs)
+ complete {C₁} {C₂} (nt , nat) = m , (λ X xs → nat X
+ (λ { (r , eq) → P.subst (Setoid.Carrier X) eq (proj₂ xs r) })
+ (proj₁ xs , (λ r → r , refl)))
+ where
+
+ m : C₁ ⇒ C₂
+ m = record
+ { command = λ c₁ → proj₁ (lemma c₁)
+ ; response = λ {_} {c₁} r₂ → proj₁ (proj₂ (lemma c₁) r₂)
+ ; coherent = λ {_} {c₁} {r₂} → proj₂ (proj₂ (lemma c₁) r₂)
+ }
+ where
+ lemma : ∀ {o} (c₁ : Command C₁ o) → Σ[ c₂ ∈ Command C₂ o ]
+ ((r₂ : Response C₂ c₂) → Σ[ r₁ ∈ Response C₁ c₁ ]
+ next C₁ c₁ r₁ ≡ next C₂ c₂ r₂)
+ lemma c₁ = nt (λ i → Σ[ r₁ ∈ Response C₁ c₁ ] next C₁ c₁ r₁ ≡ i)
+ (c₁ , λ r₁ → r₁ , refl)
+
+ -- Identity.
+
+ id : (C : Container I O c r) → C ⇒ C
+ id _ = record
+ { command = ⟨id⟩
+ ; response = ⟨id⟩
+ ; coherent = refl
+ }
+
+ -- Composition.
+
+ infixr 9 _∘_
+
+ _∘_ : {C₁ C₂ C₃ : Container I O c r} →
+ C₂ ⇒ C₃ → C₁ ⇒ C₂ → C₁ ⇒ C₃
+ f ∘ g = record
+ { command = command f ⟨∘⟩ command g
+ ; response = response g ⟨∘⟩ response f
+ ; coherent = coherent g ⟨ P.trans ⟩ coherent f
+ }
+
+ -- Identity and composition commute with ⟪_⟫.
+
+ id-correct : ∀ {ℓ} {C : Container I O c r} → ∀ {X : Pred I ℓ} {o} →
+ ⟪ id C ⟫ X {o} ≗ ⟨id⟩
+ id-correct _ = refl
+
+ ∘-correct : {C₁ C₂ C₃ : Container I O c r}
+ (f : C₂ ⇒ C₃) (g : C₁ ⇒ C₂) (X : Setoid I (c ⊔ r) _) →
+ let module X = Setoid X in
+ ∀ {o} {xs : ⟦ C₁ ⟧ X.Carrier o} →
+ Eq C₃ X.Carrier X.Carrier X._≈_
+ (⟪ f ∘ g ⟫ X.Carrier xs)
+ (⟪ f ⟫ X.Carrier (⟪ g ⟫ X.Carrier xs))
+ ∘-correct f g X = refl , refl , λ { r .r refl → lemma (coherent g)
+ (coherent f) }
+ where
+ module X = Setoid X
+
+ lemma : ∀ {i j k} (eq₁ : i ≡ j) (eq₂ : j ≡ k) {x} →
+ P.subst X.Carrier (P.trans eq₁ eq₂) x
+ X.≈
+ P.subst X.Carrier eq₂ (P.subst X.Carrier eq₁ x)
+ lemma refl refl = X.refl
+
+------------------------------------------------------------------------
+-- Linear container morphisms
+
+module LinearMorphism
+ {i o c r} {I : Set i} {O : Set o} {C₁ C₂ : Container I O c r}
+ (m : C₁ ⊸ C₂)
+ where
+
+ morphism : C₁ ⇒ C₂
+ morphism = record
+ { command = command m
+ ; response = _⟨$⟩_ (Inverse.to (response m))
+ ; coherent = coherent m
+ }
+
+ ⟪_⟫⊸ : ∀ {ℓ} (X : Pred I ℓ) → ⟦ C₁ ⟧ X ⊆ ⟦ C₂ ⟧ X
+ ⟪_⟫⊸ = ⟪ morphism ⟫
+
+open LinearMorphism public using (⟪_⟫⊸)
+
+------------------------------------------------------------------------
+-- Cartesian morphisms
+
+module CartesianMorphism
+ {i o c r} {I : Set i} {O : Set o} {C₁ C₂ : Container I O c r}
+ (m : C₁ ⇒C C₂)
+ where
+
+ morphism : C₁ ⇒ C₂
+ morphism = record
+ { command = command m
+ ; response = P.subst ⟨id⟩ (response m)
+ ; coherent = coherent m
+ }
+
+ ⟪_⟫C : ∀ {ℓ} (X : Pred I ℓ) → ⟦ C₁ ⟧ X ⊆ ⟦ C₂ ⟧ X
+ ⟪_⟫C = ⟪ morphism ⟫
+
+open CartesianMorphism public using (⟪_⟫C)
+
+------------------------------------------------------------------------
+-- All and any
+
+-- □ and ◇ are defined in the core module.
+
+module _ {i o c r ℓ₁ ℓ₂} {I : Set i} {O : Set o} (C : Container I O c r)
+ {X : Pred I ℓ₁} {P Q : Pred (Σ I X) ℓ₂} where
+
+ -- All.
+
+ □-map : P ⊆ Q → □ C P ⊆ □ C Q
+ □-map P⊆Q = _⟨∘⟩_ P⊆Q
+
+ -- Any.
+
+ ◇-map : P ⊆ Q → ◇ C P ⊆ ◇ C Q
+ ◇-map P⊆Q = Prod.map ⟨id⟩ P⊆Q
+
+-- Membership.
+
+infix 4 _∈_
+
+_∈_ : ∀ {i o c r ℓ} {I : Set i} {O : Set o}
+ {C : Container I O c r} {X : Pred I (i ⊔ ℓ)} → REL X (⟦ C ⟧ X) _
+_∈_ {C = C} {X} x xs = ◇ C {X = X} (_≅_ x) (, xs)
diff --git a/src/Data/Container/Indexed/Combinator.agda b/src/Data/Container/Indexed/Combinator.agda
new file mode 100644
index 0000000..426276e
--- /dev/null
+++ b/src/Data/Container/Indexed/Combinator.agda
@@ -0,0 +1,297 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Indexed container combinators
+------------------------------------------------------------------------
+
+module Data.Container.Indexed.Combinator where
+
+open import Level
+open import Data.Container.Indexed
+open import Data.Empty using (⊥; ⊥-elim)
+open import Data.Unit using (⊤)
+open import Data.Product as Prod hiding (Σ) renaming (_×_ to _⟨×⟩_)
+open import Data.Sum renaming (_⊎_ to _⟨⊎⟩_)
+open import Function as F hiding (id; const) renaming (_∘_ to _⟨∘⟩_)
+open import Function.Inverse using (_↔̇_)
+open import Relation.Unary using (Pred; _⊆_; _∪_; _∩_; ⋃; ⋂)
+ renaming (_⟨×⟩_ to _⟪×⟫_; _⟨⊙⟩_ to _⟪⊙⟫_; _⟨⊎⟩_ to _⟪⊎⟫_)
+open import Relation.Binary.PropositionalEquality as P
+ using (_≗_; refl)
+
+------------------------------------------------------------------------
+-- Combinators
+
+
+-- Identity.
+
+id : ∀ {o c r} {O : Set o} → Container O O c r
+id = F.const (Lift ⊤) ◃ (λ _ → Lift ⊤) / (λ {o} _ _ → o)
+
+-- Constant.
+
+const : ∀ {i o c r} {I : Set i} {O : Set o} →
+ Pred O c → Container I O c r
+const X = X ◃ (λ _ → Lift ⊥) / λ _ → ⊥-elim ⟨∘⟩ lower
+
+-- Duality.
+
+_^⊥ : ∀ {i o c r} {I : Set i} {O : Set o} →
+ Container I O c r → Container I O (c ⊔ r) c
+(C ◃ R / n) ^⊥ = record
+ { Command = λ o → (c : C o) → R c
+ ; Response = λ {o} _ → C o
+ ; next = λ f c → n c (f c)
+ }
+
+-- Strength.
+
+infixl 3 _⋊_
+
+_⋊_ : ∀ {i o c r z} {I : Set i} {O : Set o} (C : Container I O c r)
+ (Z : Set z) → Container (I ⟨×⟩ Z) (O ⟨×⟩ Z) c r
+C ◃ R / n ⋊ Z = C ⟨∘⟩ proj₁ ◃ R / λ {oz} c r → n c r , proj₂ oz
+
+infixr 3 _⋉_
+
+_⋉_ : ∀ {i o z c r} {I : Set i} {O : Set o} (Z : Set z)
+ (C : Container I O c r) → Container (Z ⟨×⟩ I) (Z ⟨×⟩ O) c r
+Z ⋉ C ◃ R / n = C ⟨∘⟩ proj₂ ◃ R / λ {zo} c r → proj₁ zo , n c r
+
+-- Composition.
+
+infixr 9 _∘_
+
+_∘_ : ∀ {i j k c r} {I : Set i} {J : Set j} {K : Set k} →
+ Container J K c r → Container I J c r → Container I K _ _
+C₁ ∘ C₂ = C ◃ R / n
+ where
+ C : ∀ k → Set _
+ C = ⟦ C₁ ⟧ (Command C₂)
+
+ R : ∀ {k} → ⟦ C₁ ⟧ (Command C₂) k → Set _
+ R (c , k) = ◇ C₁ {X = Command C₂} (Response C₂ ⟨∘⟩ proj₂) (_ , c , k)
+
+ n : ∀ {k} (c : ⟦ C₁ ⟧ (Command C₂) k) → R c → _
+ n (_ , f) (r₁ , r₂) = next C₂ (f r₁) r₂
+
+-- Product. (Note that, up to isomorphism, this is a special case of
+-- indexed product.)
+
+infixr 2 _×_
+
+_×_ : ∀ {i o c r} {I : Set i} {O : Set o} →
+ Container I O c r → Container I O c r → Container I O c r
+(C₁ ◃ R₁ / n₁) × (C₂ ◃ R₂ / n₂) = record
+ { Command = C₁ ∩ C₂
+ ; Response = R₁ ⟪⊙⟫ R₂
+ ; next = λ { (c₁ , c₂) → [ n₁ c₁ , n₂ c₂ ] }
+ }
+
+-- Indexed product.
+
+Π : ∀ {x i o c r} {X : Set x} {I : Set i} {O : Set o} →
+ (X → Container I O c r) → Container I O _ _
+Π {X = X} C = record
+ { Command = ⋂ X (Command ⟨∘⟩ C)
+ ; Response = ⋃[ x ∶ X ] λ c → Response (C x) (c x)
+ ; next = λ { c (x , r) → next (C x) (c x) r }
+ }
+
+-- Sum. (Note that, up to isomorphism, this is a special case of
+-- indexed sum.)
+
+infixr 1 _⊎_
+
+_⊎_ : ∀ {i o c r} {I : Set i} {O : Set o} →
+ Container I O c r → Container I O c r → Container I O _ _
+(C₁ ◃ R₁ / n₁) ⊎ (C₂ ◃ R₂ / n₂) = record
+ { Command = C₁ ∪ C₂
+ ; Response = R₁ ⟪⊎⟫ R₂
+ ; next = [ n₁ , n₂ ]
+ }
+
+-- Indexed sum.
+
+Σ : ∀ {x i o c r} {X : Set x} {I : Set i} {O : Set o} →
+ (X → Container I O c r) → Container I O _ r
+Σ {X = X} C = record
+ { Command = ⋃ X (Command ⟨∘⟩ C)
+ ; Response = λ { (x , c) → Response (C x) c }
+ ; next = λ { (x , c) r → next (C x) c r }
+ }
+
+-- Constant exponentiation. (Note that this is a special case of
+-- indexed product.)
+
+infix 0 const[_]⟶_
+
+const[_]⟶_ : ∀ {i o c r ℓ} {I : Set i} {O : Set o} →
+ Set ℓ → Container I O c r → Container I O _ _
+const[ X ]⟶ C = Π {X = X} (F.const C)
+
+------------------------------------------------------------------------
+-- Correctness proofs
+
+module Identity where
+
+ correct : ∀ {o ℓ c r} {O : Set o} {X : Pred O ℓ} →
+ ⟦ id {c = c}{r} ⟧ X ↔̇ F.id X
+ correct = record
+ { to = P.→-to-⟶ λ xs → proj₂ xs _
+ ; from = P.→-to-⟶ λ x → (_ , λ _ → x)
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+
+module Constant (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
+
+ correct : ∀ {i o ℓ₁ ℓ₂} {I : Set i} {O : Set o} (X : Pred O ℓ₁)
+ {Y : Pred O ℓ₂} → ⟦ const X ⟧ Y ↔̇ F.const X Y
+ correct X {Y} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { right-inverse-of = λ _ → refl
+ ; left-inverse-of =
+ λ xs → P.cong (_,_ (proj₁ xs)) (ext (⊥-elim ⟨∘⟩ lower))
+ }
+ }
+ where
+ to : ⟦ const X ⟧ Y ⊆ X
+ to = proj₁
+
+ from : X ⊆ ⟦ const X ⟧ Y
+ from = < F.id , F.const (⊥-elim ⟨∘⟩ lower) >
+
+module Duality where
+
+ correct : ∀ {i o c r ℓ} {I : Set i} {O : Set o}
+ (C : Container I O c r) (X : Pred I ℓ) →
+ ⟦ C ^⊥ ⟧ X ↔̇ (λ o → (c : Command C o) → ∃ λ r → X (next C c r))
+ correct C X = record
+ { to = P.→-to-⟶ λ { (f , g) → < f , g > }
+ ; from = P.→-to-⟶ λ f → proj₁ ⟨∘⟩ f , proj₂ ⟨∘⟩ f
+ ; inverse-of = record
+ { left-inverse-of = λ { (_ , _) → refl }
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+
+module Composition where
+
+ correct : ∀ {i j k ℓ c r} {I : Set i} {J : Set j} {K : Set k}
+ (C₁ : Container J K c r) (C₂ : Container I J c r) →
+ {X : Pred I ℓ} → ⟦ C₁ ∘ C₂ ⟧ X ↔̇ (⟦ C₁ ⟧ ⟨∘⟩ ⟦ C₂ ⟧) X
+ correct C₁ C₂ {X} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+ where
+ to : ⟦ C₁ ∘ C₂ ⟧ X ⊆ ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)
+ to ((c , f) , g) = (c , < f , curry g >)
+
+ from : ⟦ C₁ ⟧ (⟦ C₂ ⟧ X) ⊆ ⟦ C₁ ∘ C₂ ⟧ X
+ from (c , f) = ((c , proj₁ ⟨∘⟩ f) , uncurry (proj₂ ⟨∘⟩ f))
+
+module Product (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
+
+ correct : ∀ {i o c r} {I : Set i} {O : Set o}
+ (C₁ C₂ : Container I O c r) {X} →
+ ⟦ C₁ × C₂ ⟧ X ↔̇ (⟦ C₁ ⟧ X ∩ ⟦ C₂ ⟧ X)
+ correct C₁ C₂ {X} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = from∘to
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+ where
+ to : ⟦ C₁ × C₂ ⟧ X ⊆ ⟦ C₁ ⟧ X ∩ ⟦ C₂ ⟧ X
+ to ((c₁ , c₂) , k) = ((c₁ , k ⟨∘⟩ inj₁) , (c₂ , k ⟨∘⟩ inj₂))
+
+ from : ⟦ C₁ ⟧ X ∩ ⟦ C₂ ⟧ X ⊆ ⟦ C₁ × C₂ ⟧ X
+ from ((c₁ , k₁) , (c₂ , k₂)) = ((c₁ , c₂) , [ k₁ , k₂ ])
+
+ from∘to : from ⟨∘⟩ to ≗ F.id
+ from∘to (c , _) =
+ P.cong (_,_ c) (ext [ (λ _ → refl) , (λ _ → refl) ])
+
+module IndexedProduct where
+
+ correct : ∀ {x i o c r ℓ} {X : Set x} {I : Set i} {O : Set o}
+ (C : X → Container I O c r) {Y : Pred I ℓ} →
+ ⟦ Π C ⟧ Y ↔̇ ⋂[ x ∶ X ] ⟦ C x ⟧ Y
+ correct {X = X} C {Y} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+ where
+ to : ⟦ Π C ⟧ Y ⊆ ⋂[ x ∶ X ] ⟦ C x ⟧ Y
+ to (c , k) = λ x → (c x , λ r → k (x , r))
+
+ from : ⋂[ x ∶ X ] ⟦ C x ⟧ Y ⊆ ⟦ Π C ⟧ Y
+ from f = (proj₁ ⟨∘⟩ f , uncurry (proj₂ ⟨∘⟩ f))
+
+module Sum where
+
+ correct : ∀ {i o c r ℓ} {I : Set i} {O : Set o}
+ (C₁ C₂ : Container I O c r) {X : Pred I ℓ} →
+ ⟦ C₁ ⊎ C₂ ⟧ X ↔̇ (⟦ C₁ ⟧ X ∪ ⟦ C₂ ⟧ X)
+ correct C₁ C₂ {X} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = from∘to
+ ; right-inverse-of = [ (λ _ → refl) , (λ _ → refl) ]
+ }
+ }
+ where
+ to : ⟦ C₁ ⊎ C₂ ⟧ X ⊆ ⟦ C₁ ⟧ X ∪ ⟦ C₂ ⟧ X
+ to (inj₁ c₁ , k) = inj₁ (c₁ , k)
+ to (inj₂ c₂ , k) = inj₂ (c₂ , k)
+
+ from : ⟦ C₁ ⟧ X ∪ ⟦ C₂ ⟧ X ⊆ ⟦ C₁ ⊎ C₂ ⟧ X
+ from = [ Prod.map inj₁ F.id , Prod.map inj₂ F.id ]
+
+ from∘to : from ⟨∘⟩ to ≗ F.id
+ from∘to (inj₁ _ , _) = refl
+ from∘to (inj₂ _ , _) = refl
+
+module IndexedSum where
+
+ correct : ∀ {x i o c r ℓ} {X : Set x} {I : Set i} {O : Set o}
+ (C : X → Container I O c r) {Y : Pred I ℓ} →
+ ⟦ Σ C ⟧ Y ↔̇ ⋃[ x ∶ X ] ⟦ C x ⟧ Y
+ correct {X = X} C {Y} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl
+ ; right-inverse-of = λ _ → refl
+ }
+ }
+ where
+ to : ⟦ Σ C ⟧ Y ⊆ ⋃[ x ∶ X ] ⟦ C x ⟧ Y
+ to ((x , c) , k) = (x , (c , k))
+
+ from : ⋃[ x ∶ X ] ⟦ C x ⟧ Y ⊆ ⟦ Σ C ⟧ Y
+ from (x , (c , k)) = ((x , c) , k)
+
+module ConstantExponentiation where
+
+ correct : ∀ {x i o c r ℓ} {X : Set x} {I : Set i} {O : Set o}
+ (C : Container I O c r) {Y : Pred I ℓ} →
+ ⟦ const[ X ]⟶ C ⟧ Y ↔̇ (⋂ X (F.const (⟦ C ⟧ Y)))
+ correct C {Y} = IndexedProduct.correct (F.const C) {Y}
diff --git a/src/Data/Container/Indexed/Core.agda b/src/Data/Container/Indexed/Core.agda
new file mode 100644
index 0000000..e6b1585
--- /dev/null
+++ b/src/Data/Container/Indexed/Core.agda
@@ -0,0 +1,43 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Indexed containers core
+------------------------------------------------------------------------
+
+module Data.Container.Indexed.Core where
+
+open import Level
+open import Data.Product
+open import Relation.Unary
+
+infix 5 _◃_/_
+
+record Container {i o} (I : Set i) (O : Set o) (c r : Level) :
+ Set (i ⊔ o ⊔ suc c ⊔ suc r) where
+ constructor _◃_/_
+ field
+ Command : (o : O) → Set c
+ Response : ∀ {o} → Command o → Set r
+ next : ∀ {o} (c : Command o) → Response c → I
+
+-- The semantics ("extension") of an indexed container.
+
+⟦_⟧ : ∀ {i o c r ℓ} {I : Set i} {O : Set o} → Container I O c r →
+ Pred I ℓ → Pred O _
+⟦ C ◃ R / n ⟧ X o = Σ[ c ∈ C o ] ((r : R c) → X (n c r))
+
+------------------------------------------------------------------------
+-- All and any
+
+module _ {i o c r ℓ} {I : Set i} {O : Set o}
+ (C : Container I O c r) {X : Pred I ℓ} where
+
+ -- All.
+
+ □ : ∀ {ℓ′} → Pred (Σ I X) ℓ′ → Pred (Σ O (⟦ C ⟧ X)) _
+ □ P (_ , _ , k) = ∀ r → P (_ , k r)
+
+ -- Any.
+
+ ◇ : ∀ {ℓ′} → Pred (Σ I X) ℓ′ → Pred (Σ O (⟦ C ⟧ X)) _
+ ◇ P (_ , _ , k) = ∃ λ r → P (_ , k r)
diff --git a/src/Data/Container/Indexed/FreeMonad.agda b/src/Data/Container/Indexed/FreeMonad.agda
new file mode 100644
index 0000000..74051be
--- /dev/null
+++ b/src/Data/Container/Indexed/FreeMonad.agda
@@ -0,0 +1,58 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- The free monad construction on indexed containers
+------------------------------------------------------------------------
+
+module Data.Container.Indexed.FreeMonad where
+
+open import Level
+open import Function hiding (const)
+open import Category.Monad.Predicate
+open import Data.Container.Indexed hiding (_∈_)
+open import Data.Container.Indexed.Combinator hiding (id; _∘_)
+open import Data.Empty
+open import Data.Sum using (inj₁; inj₂)
+open import Data.Product
+open import Data.W.Indexed
+open import Relation.Unary
+open import Relation.Unary.PredicateTransformer
+
+------------------------------------------------------------------------
+
+infixl 9 _⋆C_
+infix 9 _⋆_
+
+_⋆C_ : ∀ {i o c r} {I : Set i} {O : Set o} →
+ Container I O c r → Pred O c → Container I O _ _
+C ⋆C X = const X ⊎ C
+
+_⋆_ : ∀ {ℓ} {O : Set ℓ} → Container O O ℓ ℓ → Pt O ℓ
+C ⋆ X = μ (C ⋆C X)
+
+pattern returnP x = (inj₁ x , _)
+pattern doP c k = (inj₂ c , k)
+
+do : ∀ {ℓ} {O : Set ℓ} {C : Container O O ℓ ℓ} {X} →
+ ⟦ C ⟧ (C ⋆ X) ⊆ C ⋆ X
+do (c , k) = sup (doP c k)
+
+rawPMonad : ∀ {ℓ} {O : Set ℓ} {C : Container O O ℓ ℓ} →
+ RawPMonad {ℓ = ℓ} (_⋆_ C)
+rawPMonad {C = C} = record
+ { return? = return
+ ; _=<?_ = _=<<_
+ }
+ where
+ return : ∀ {X} → X ⊆ C ⋆ X
+ return x = sup (inj₁ x , ⊥-elim ∘ lower)
+
+ _=<<_ : ∀ {X Y} → X ⊆ C ⋆ Y → C ⋆ X ⊆ C ⋆ Y
+ f =<< sup (returnP x) = f x
+ f =<< sup (doP c k) = do (c , λ r → f =<< k r)
+
+leaf : ∀ {ℓ} {O : Set ℓ} {C : Container O O ℓ ℓ} {X : Pred O ℓ} →
+ ⟦ C ⟧ X ⊆ C ⋆ X
+leaf (c , k) = do (c , return? ∘ k)
+ where
+ open RawPMonad rawPMonad
diff --git a/src/Data/Fin/Dec.agda b/src/Data/Fin/Dec.agda
index 1e7fd6f..6318eb8 100644
--- a/src/Data/Fin/Dec.agda
+++ b/src/Data/Fin/Dec.agda
@@ -14,7 +14,7 @@ open import Data.Vec.Equality as VecEq
using () renaming (module PropositionalEquality to PropVecEq)
open import Data.Fin
open import Data.Fin.Subset
-open import Data.Fin.Subset.Props
+open import Data.Fin.Subset.Properties
open import Data.Product as Prod
open import Data.Empty
open import Function
diff --git a/src/Data/Fin/Properties.agda b/src/Data/Fin/Properties.agda
new file mode 100644
index 0000000..c2a8b11
--- /dev/null
+++ b/src/Data/Fin/Properties.agda
@@ -0,0 +1,240 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties related to Fin, and operations making use of these
+-- properties (or other properties not available in Data.Fin)
+------------------------------------------------------------------------
+
+module Data.Fin.Properties where
+
+open import Algebra
+open import Data.Fin
+open import Data.Nat as N
+ using (ℕ; zero; suc; s≤s; z≤n; _∸_)
+ renaming (_≤_ to _ℕ≤_; _<_ to _ℕ<_; _+_ to _ℕ+_)
+import Data.Nat.Properties as N
+open import Data.Product
+open import Function
+open import Function.Equality as FunS using (_⟨$⟩_)
+open import Function.Injection using (_↣_)
+open import Algebra.FunctionProperties
+open import Relation.Nullary
+import Relation.Nullary.Decidable as Dec
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; refl; cong; subst)
+open import Category.Functor
+open import Category.Applicative
+
+------------------------------------------------------------------------
+-- Properties
+
+private
+ drop-suc : ∀ {o} {m n : Fin o} → Fin.suc m ≡ suc n → m ≡ n
+ drop-suc refl = refl
+
+preorder : ℕ → Preorder _ _ _
+preorder n = P.preorder (Fin n)
+
+setoid : ℕ → Setoid _ _
+setoid n = P.setoid (Fin n)
+
+strictTotalOrder : ℕ → StrictTotalOrder _ _ _
+strictTotalOrder n = record
+ { Carrier = Fin n
+ ; _≈_ = _≡_
+ ; _<_ = _<_
+ ; isStrictTotalOrder = record
+ { isEquivalence = P.isEquivalence
+ ; trans = N.<-trans
+ ; compare = cmp
+ ; <-resp-≈ = P.resp₂ _<_
+ }
+ }
+ where
+ cmp : ∀ {n} → Trichotomous _≡_ (_<_ {n})
+ cmp zero zero = tri≈ (λ()) refl (λ())
+ cmp zero (suc j) = tri< (s≤s z≤n) (λ()) (λ())
+ cmp (suc i) zero = tri> (λ()) (λ()) (s≤s z≤n)
+ cmp (suc i) (suc j) with cmp i j
+ ... | tri< lt ¬eq ¬gt = tri< (s≤s lt) (¬eq ∘ drop-suc) (¬gt ∘ N.≤-pred)
+ ... | tri> ¬lt ¬eq gt = tri> (¬lt ∘ N.≤-pred) (¬eq ∘ drop-suc) (s≤s gt)
+ ... | tri≈ ¬lt eq ¬gt = tri≈ (¬lt ∘ N.≤-pred) (cong suc eq) (¬gt ∘ N.≤-pred)
+
+decSetoid : ℕ → DecSetoid _ _
+decSetoid n = StrictTotalOrder.decSetoid (strictTotalOrder n)
+
+infix 4 _≟_
+
+_≟_ : {n : ℕ} → Decidable {A = Fin n} _≡_
+_≟_ {n} = DecSetoid._≟_ (decSetoid n)
+
+to-from : ∀ n → toℕ (fromℕ n) ≡ n
+to-from zero = refl
+to-from (suc n) = cong suc (to-from n)
+
+toℕ-injective : ∀ {n} {i j : Fin n} → toℕ i ≡ toℕ j → i ≡ j
+toℕ-injective {zero} {} {} _
+toℕ-injective {suc n} {zero} {zero} eq = refl
+toℕ-injective {suc n} {zero} {suc j} ()
+toℕ-injective {suc n} {suc i} {zero} ()
+toℕ-injective {suc n} {suc i} {suc j} eq =
+ cong suc (toℕ-injective (cong N.pred eq))
+
+bounded : ∀ {n} (i : Fin n) → toℕ i ℕ< n
+bounded zero = s≤s z≤n
+bounded (suc i) = s≤s (bounded i)
+
+prop-toℕ-≤ : ∀ {n} (x : Fin n) → toℕ x ℕ≤ N.pred n
+prop-toℕ-≤ zero = z≤n
+prop-toℕ-≤ (suc {n = zero} ())
+prop-toℕ-≤ (suc {n = suc n} i) = s≤s (prop-toℕ-≤ i)
+
+nℕ-ℕi≤n : ∀ n i → n ℕ-ℕ i ℕ≤ n
+nℕ-ℕi≤n n zero = begin n ∎
+ where open N.≤-Reasoning
+nℕ-ℕi≤n zero (suc ())
+nℕ-ℕi≤n (suc n) (suc i) = begin
+ n ℕ-ℕ i ≤⟨ nℕ-ℕi≤n n i ⟩
+ n ≤⟨ N.n≤1+n n ⟩
+ suc n ∎
+ where open N.≤-Reasoning
+
+inject-lemma : ∀ {n} {i : Fin n} (j : Fin′ i) →
+ toℕ (inject j) ≡ toℕ j
+inject-lemma {i = zero} ()
+inject-lemma {i = suc i} zero = refl
+inject-lemma {i = suc i} (suc j) = cong suc (inject-lemma j)
+
+inject+-lemma : ∀ {m} n (i : Fin m) → toℕ i ≡ toℕ (inject+ n i)
+inject+-lemma n zero = refl
+inject+-lemma n (suc i) = cong suc (inject+-lemma n i)
+
+inject₁-lemma : ∀ {m} (i : Fin m) → toℕ (inject₁ i) ≡ toℕ i
+inject₁-lemma zero = refl
+inject₁-lemma (suc i) = cong suc (inject₁-lemma i)
+
+inject≤-lemma : ∀ {m n} (i : Fin m) (le : m ℕ≤ n) →
+ toℕ (inject≤ i le) ≡ toℕ i
+inject≤-lemma zero (N.s≤s le) = refl
+inject≤-lemma (suc i) (N.s≤s le) = cong suc (inject≤-lemma i le)
+
+≺⇒<′ : _≺_ ⇒ N._<′_
+≺⇒<′ (n ≻toℕ i) = N.≤⇒≤′ (bounded i)
+
+<′⇒≺ : N._<′_ ⇒ _≺_
+<′⇒≺ {n} N.≤′-refl = subst (λ i → i ≺ suc n) (to-from n)
+ (suc n ≻toℕ fromℕ n)
+<′⇒≺ (N.≤′-step m≤′n) with <′⇒≺ m≤′n
+<′⇒≺ (N.≤′-step m≤′n) | n ≻toℕ i =
+ subst (λ i → i ≺ suc n) (inject₁-lemma i) (suc n ≻toℕ (inject₁ i))
+
+toℕ-raise : ∀ {m} n (i : Fin m) → toℕ (raise n i) ≡ n ℕ+ toℕ i
+toℕ-raise zero i = refl
+toℕ-raise (suc n) i = cong suc (toℕ-raise n i)
+
+fromℕ≤-toℕ : ∀ {m} (i : Fin m) (i<m : toℕ i ℕ< m) → fromℕ≤ i<m ≡ i
+fromℕ≤-toℕ zero (s≤s z≤n) = refl
+fromℕ≤-toℕ (suc i) (s≤s (s≤s m≤n)) = cong suc (fromℕ≤-toℕ i (s≤s m≤n))
+
+toℕ-fromℕ≤ : ∀ {m n} (m<n : m ℕ< n) → toℕ (fromℕ≤ m<n) ≡ m
+toℕ-fromℕ≤ (s≤s z≤n) = refl
+toℕ-fromℕ≤ (s≤s (s≤s m<n)) = cong suc (toℕ-fromℕ≤ (s≤s m<n))
+
+------------------------------------------------------------------------
+-- Operations
+
+infixl 6 _+′_
+
+_+′_ : ∀ {m n} (i : Fin m) (j : Fin n) → Fin (N.pred m ℕ+ n)
+i +′ j = inject≤ (i + j) (N._+-mono_ (prop-toℕ-≤ i) ≤-refl)
+ where open DecTotalOrder N.decTotalOrder renaming (refl to ≤-refl)
+
+-- reverse {n} "i" = "n ∸ 1 ∸ i".
+
+reverse : ∀ {n} → Fin n → Fin n
+reverse {zero} ()
+reverse {suc n} i = inject≤ (n ℕ- i) (N.n∸m≤n (toℕ i) (suc n))
+
+reverse-prop : ∀ {n} → (i : Fin n) → toℕ (reverse i) ≡ n ∸ suc (toℕ i)
+reverse-prop {zero} ()
+reverse-prop {suc n} i = begin
+ toℕ (inject≤ (n ℕ- i) _) ≡⟨ inject≤-lemma _ _ ⟩
+ toℕ (n ℕ- i) ≡⟨ toℕ‿ℕ- n i ⟩
+ n ∸ toℕ i ∎
+ where
+ open P.≡-Reasoning
+
+ toℕ‿ℕ- : ∀ n i → toℕ (n ℕ- i) ≡ n ∸ toℕ i
+ toℕ‿ℕ- n zero = to-from n
+ toℕ‿ℕ- zero (suc ())
+ toℕ‿ℕ- (suc n) (suc i) = toℕ‿ℕ- n i
+
+reverse-involutive : ∀ {n} → Involutive _≡_ reverse
+reverse-involutive {n} i = toℕ-injective (begin
+ toℕ (reverse (reverse i)) ≡⟨ reverse-prop _ ⟩
+ n ∸ suc (toℕ (reverse i)) ≡⟨ eq ⟩
+ toℕ i ∎)
+ where
+ open P.≡-Reasoning
+ open CommutativeSemiring N.commutativeSemiring using (+-comm)
+
+ lem₁ : ∀ m n → (m ℕ+ n) ∸ (m ℕ+ n ∸ m) ≡ m
+ lem₁ m n = begin
+ m ℕ+ n ∸ (m ℕ+ n ∸ m) ≡⟨ cong (λ ξ → m ℕ+ n ∸ (ξ ∸ m)) (+-comm m n) ⟩
+ m ℕ+ n ∸ (n ℕ+ m ∸ m) ≡⟨ cong (λ ξ → m ℕ+ n ∸ ξ) (N.m+n∸n≡m n m) ⟩
+ m ℕ+ n ∸ n ≡⟨ N.m+n∸n≡m m n ⟩
+ m ∎
+
+ lem₂ : ∀ n → (i : Fin n) → n ∸ suc (n ∸ suc (toℕ i)) ≡ toℕ i
+ lem₂ zero ()
+ lem₂ (suc n) i = begin
+ n ∸ (n ∸ toℕ i) ≡⟨ cong (λ ξ → ξ ∸ (ξ ∸ toℕ i)) i+j≡k ⟩
+ (toℕ i ℕ+ j) ∸ (toℕ i ℕ+ j ∸ toℕ i) ≡⟨ lem₁ (toℕ i) j ⟩
+ toℕ i ∎
+ where
+ decompose-n : ∃ λ j → n ≡ toℕ i ℕ+ j
+ decompose-n = n ∸ toℕ i , P.sym (N.m+n∸m≡n (prop-toℕ-≤ i))
+
+ j = proj₁ decompose-n
+ i+j≡k = proj₂ decompose-n
+
+ eq : n ∸ suc (toℕ (reverse i)) ≡ toℕ i
+ eq = begin
+ n ∸ suc (toℕ (reverse i)) ≡⟨ cong (λ ξ → n ∸ suc ξ) (reverse-prop i) ⟩
+ n ∸ suc (n ∸ suc (toℕ i)) ≡⟨ lem₂ n i ⟩
+ toℕ i ∎
+
+-- If there is an injection from a type to a finite set, then the type
+-- has decidable equality.
+
+eq? : ∀ {a n} {A : Set a} → A ↣ Fin n → Decidable {A = A} _≡_
+eq? inj = Dec.via-injection inj _≟_
+
+-- Quantification over finite sets commutes with applicative functors.
+
+sequence : ∀ {F n} {P : Fin n → Set} → RawApplicative F →
+ (∀ i → F (P i)) → F (∀ i → P i)
+sequence {F} RA = helper _ _
+ where
+ open RawApplicative RA
+
+ helper : ∀ n (P : Fin n → Set) → (∀ i → F (P i)) → F (∀ i → P i)
+ helper zero P ∀iPi = pure (λ())
+ helper (suc n) P ∀iPi =
+ combine <$> ∀iPi zero ⊛ helper n (λ n → P (suc n)) (∀iPi ∘ suc)
+ where
+ combine : P zero → (∀ i → P (suc i)) → ∀ i → P i
+ combine z s zero = z
+ combine z s (suc i) = s i
+
+private
+
+ -- Included just to show that sequence above has an inverse (under
+ -- 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 →
+ 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/Props.agda b/src/Data/Fin/Props.agda
index 1eae191..897f7bc 100644
--- a/src/Data/Fin/Props.agda
+++ b/src/Data/Fin/Props.agda
@@ -1,243 +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)
+-- Compatibility module. Pending for removal. Use Data.Fin.Properties
+-- instead.
------------------------------------------------------------------------
module Data.Fin.Props where
-open import Algebra
-open import Data.Fin
-open import Data.Nat as N
- using (ℕ; zero; suc; s≤s; z≤n; _∸_)
- renaming (_≤_ to _ℕ≤_; _<_ to _ℕ<_; _+_ to _ℕ+_)
-import Data.Nat.Properties as N
-open import Data.Product
-open import Function
-open import Function.Equality as FunS using (_⟨$⟩_)
-open import Function.Injection
- using (Injection; module Injection)
-open import Algebra.FunctionProperties
-open import Relation.Nullary
-open import Relation.Binary
-open import Relation.Binary.PropositionalEquality as P
- using (_≡_; refl; cong; subst)
-open import Category.Functor
-open import Category.Applicative
-
-------------------------------------------------------------------------
--- Properties
-
-private
- drop-suc : ∀ {o} {m n : Fin o} → Fin.suc m ≡ suc n → m ≡ n
- drop-suc refl = refl
-
-preorder : ℕ → Preorder _ _ _
-preorder n = P.preorder (Fin n)
-
-setoid : ℕ → Setoid _ _
-setoid n = P.setoid (Fin n)
-
-strictTotalOrder : ℕ → StrictTotalOrder _ _ _
-strictTotalOrder n = record
- { Carrier = Fin n
- ; _≈_ = _≡_
- ; _<_ = _<_
- ; isStrictTotalOrder = record
- { isEquivalence = P.isEquivalence
- ; trans = N.<-trans
- ; compare = cmp
- ; <-resp-≈ = P.resp₂ _<_
- }
- }
- where
- cmp : ∀ {n} → Trichotomous _≡_ (_<_ {n})
- cmp zero zero = tri≈ (λ()) refl (λ())
- cmp zero (suc j) = tri< (s≤s z≤n) (λ()) (λ())
- cmp (suc i) zero = tri> (λ()) (λ()) (s≤s z≤n)
- cmp (suc i) (suc j) with cmp i j
- ... | tri< lt ¬eq ¬gt = tri< (s≤s lt) (¬eq ∘ drop-suc) (¬gt ∘ N.≤-pred)
- ... | tri> ¬lt ¬eq gt = tri> (¬lt ∘ N.≤-pred) (¬eq ∘ drop-suc) (s≤s gt)
- ... | tri≈ ¬lt eq ¬gt = tri≈ (¬lt ∘ N.≤-pred) (cong suc eq) (¬gt ∘ N.≤-pred)
-
-decSetoid : ℕ → DecSetoid _ _
-decSetoid n = StrictTotalOrder.decSetoid (strictTotalOrder n)
-
-infix 4 _≟_
-
-_≟_ : {n : ℕ} → Decidable {A = Fin n} _≡_
-_≟_ {n} = DecSetoid._≟_ (decSetoid n)
-
-to-from : ∀ n → toℕ (fromℕ n) ≡ n
-to-from zero = refl
-to-from (suc n) = cong suc (to-from n)
-
-toℕ-injective : ∀ {n} {i j : Fin n} → toℕ i ≡ toℕ j → i ≡ j
-toℕ-injective {zero} {} {} _
-toℕ-injective {suc n} {zero} {zero} eq = refl
-toℕ-injective {suc n} {zero} {suc j} ()
-toℕ-injective {suc n} {suc i} {zero} ()
-toℕ-injective {suc n} {suc i} {suc j} eq =
- cong suc (toℕ-injective (cong N.pred eq))
-
-bounded : ∀ {n} (i : Fin n) → toℕ i ℕ< n
-bounded zero = s≤s z≤n
-bounded (suc i) = s≤s (bounded i)
-
-prop-toℕ-≤ : ∀ {n} (x : Fin n) → toℕ x ℕ≤ N.pred n
-prop-toℕ-≤ zero = z≤n
-prop-toℕ-≤ (suc {n = zero} ())
-prop-toℕ-≤ (suc {n = suc n} i) = s≤s (prop-toℕ-≤ i)
-
-nℕ-ℕi≤n : ∀ n i → n ℕ-ℕ i ℕ≤ n
-nℕ-ℕi≤n n zero = begin n ∎
- where open N.≤-Reasoning
-nℕ-ℕi≤n zero (suc ())
-nℕ-ℕi≤n (suc n) (suc i) = begin
- n ℕ-ℕ i ≤⟨ nℕ-ℕi≤n n i ⟩
- n ≤⟨ N.n≤1+n n ⟩
- suc n ∎
- where open N.≤-Reasoning
-
-inject-lemma : ∀ {n} {i : Fin n} (j : Fin′ i) →
- toℕ (inject j) ≡ toℕ j
-inject-lemma {i = zero} ()
-inject-lemma {i = suc i} zero = refl
-inject-lemma {i = suc i} (suc j) = cong suc (inject-lemma j)
-
-inject+-lemma : ∀ {m} n (i : Fin m) → toℕ i ≡ toℕ (inject+ n i)
-inject+-lemma n zero = refl
-inject+-lemma n (suc i) = cong suc (inject+-lemma n i)
-
-inject₁-lemma : ∀ {m} (i : Fin m) → toℕ (inject₁ i) ≡ toℕ i
-inject₁-lemma zero = refl
-inject₁-lemma (suc i) = cong suc (inject₁-lemma i)
-
-inject≤-lemma : ∀ {m n} (i : Fin m) (le : m ℕ≤ n) →
- toℕ (inject≤ i le) ≡ toℕ i
-inject≤-lemma zero (N.s≤s le) = refl
-inject≤-lemma (suc i) (N.s≤s le) = cong suc (inject≤-lemma i le)
-
-≺⇒<′ : _≺_ ⇒ N._<′_
-≺⇒<′ (n ≻toℕ i) = N.≤⇒≤′ (bounded i)
-
-<′⇒≺ : N._<′_ ⇒ _≺_
-<′⇒≺ {n} N.≤′-refl = subst (λ i → i ≺ suc n) (to-from n)
- (suc n ≻toℕ fromℕ n)
-<′⇒≺ (N.≤′-step m≤′n) with <′⇒≺ m≤′n
-<′⇒≺ (N.≤′-step m≤′n) | n ≻toℕ i =
- subst (λ i → i ≺ suc n) (inject₁-lemma i) (suc n ≻toℕ (inject₁ i))
-
-toℕ-raise : ∀ {m} n (i : Fin m) → toℕ (raise n i) ≡ n ℕ+ toℕ i
-toℕ-raise zero i = refl
-toℕ-raise (suc n) i = cong suc (toℕ-raise n i)
-
-fromℕ≤-toℕ : ∀ {m} (i : Fin m) (i<m : toℕ i ℕ< m) → fromℕ≤ i<m ≡ i
-fromℕ≤-toℕ zero (s≤s z≤n) = refl
-fromℕ≤-toℕ (suc i) (s≤s (s≤s m≤n)) = cong suc (fromℕ≤-toℕ i (s≤s m≤n))
-
-toℕ-fromℕ≤ : ∀ {m n} (m<n : m ℕ< n) → toℕ (fromℕ≤ m<n) ≡ m
-toℕ-fromℕ≤ (s≤s z≤n) = refl
-toℕ-fromℕ≤ (s≤s (s≤s m<n)) = cong suc (toℕ-fromℕ≤ (s≤s m<n))
-
-------------------------------------------------------------------------
--- Operations
-
-infixl 6 _+′_
-
-_+′_ : ∀ {m n} (i : Fin m) (j : Fin n) → Fin (N.pred m ℕ+ n)
-i +′ j = inject≤ (i + j) (N._+-mono_ (prop-toℕ-≤ i) ≤-refl)
- where open DecTotalOrder N.decTotalOrder renaming (refl to ≤-refl)
-
--- reverse {n} "i" = "n ∸ 1 ∸ i".
-
-reverse : ∀ {n} → Fin n → Fin n
-reverse {zero} ()
-reverse {suc n} i = inject≤ (n ℕ- i) (N.n∸m≤n (toℕ i) (suc n))
-
-reverse-prop : ∀ {n} → (i : Fin n) → toℕ (reverse i) ≡ n ∸ suc (toℕ i)
-reverse-prop {zero} ()
-reverse-prop {suc n} i = begin
- toℕ (inject≤ (n ℕ- i) _) ≡⟨ inject≤-lemma _ _ ⟩
- toℕ (n ℕ- i) ≡⟨ toℕ‿ℕ- n i ⟩
- n ∸ toℕ i ∎
- where
- open P.≡-Reasoning
-
- toℕ‿ℕ- : ∀ n i → toℕ (n ℕ- i) ≡ n ∸ toℕ i
- toℕ‿ℕ- n zero = to-from n
- toℕ‿ℕ- zero (suc ())
- toℕ‿ℕ- (suc n) (suc i) = toℕ‿ℕ- n i
-
-reverse-involutive : ∀ {n} → Involutive _≡_ reverse
-reverse-involutive {n} i = toℕ-injective (begin
- toℕ (reverse (reverse i)) ≡⟨ reverse-prop _ ⟩
- n ∸ suc (toℕ (reverse i)) ≡⟨ eq ⟩
- toℕ i ∎)
- where
- open P.≡-Reasoning
- open CommutativeSemiring N.commutativeSemiring using (+-comm)
-
- lem₁ : ∀ m n → (m ℕ+ n) ∸ (m ℕ+ n ∸ m) ≡ m
- lem₁ m n = begin
- m ℕ+ n ∸ (m ℕ+ n ∸ m) ≡⟨ cong (λ ξ → m ℕ+ n ∸ (ξ ∸ m)) (+-comm m n) ⟩
- m ℕ+ n ∸ (n ℕ+ m ∸ m) ≡⟨ cong (λ ξ → m ℕ+ n ∸ ξ) (N.m+n∸n≡m n m) ⟩
- m ℕ+ n ∸ n ≡⟨ N.m+n∸n≡m m n ⟩
- m ∎
-
- lem₂ : ∀ n → (i : Fin n) → n ∸ suc (n ∸ suc (toℕ i)) ≡ toℕ i
- lem₂ zero ()
- lem₂ (suc n) i = begin
- n ∸ (n ∸ toℕ i) ≡⟨ cong (λ ξ → ξ ∸ (ξ ∸ toℕ i)) i+j≡k ⟩
- (toℕ i ℕ+ j) ∸ (toℕ i ℕ+ j ∸ toℕ i) ≡⟨ lem₁ (toℕ i) j ⟩
- toℕ i ∎
- where
- decompose-n : ∃ λ j → n ≡ toℕ i ℕ+ j
- decompose-n = n ∸ toℕ i , P.sym (N.m+n∸m≡n (prop-toℕ-≤ i))
-
- j = proj₁ decompose-n
- i+j≡k = proj₂ decompose-n
-
- eq : n ∸ suc (toℕ (reverse i)) ≡ toℕ i
- eq = begin
- n ∸ suc (toℕ (reverse i)) ≡⟨ cong (λ ξ → n ∸ suc ξ) (reverse-prop i) ⟩
- n ∸ suc (n ∸ suc (toℕ i)) ≡⟨ lem₂ n i ⟩
- toℕ i ∎
-
--- If there is an injection from a set to a finite set, then equality
--- of the set can be decided.
-
-eq? : ∀ {s₁ s₂ n} {S : Setoid s₁ s₂} →
- Injection S (P.setoid (Fin n)) → Decidable (Setoid._≈_ S)
-eq? inj x y with to ⟨$⟩ x ≟ to ⟨$⟩ y where open Injection inj
-... | yes tox≡toy = yes (Injection.injective inj tox≡toy)
-... | no tox≢toy = no (tox≢toy ∘ FunS.cong (Injection.to inj))
-
--- Quantification over finite sets commutes with applicative functors.
-
-sequence : ∀ {F n} {P : Fin n → Set} → RawApplicative F →
- (∀ i → F (P i)) → F (∀ i → P i)
-sequence {F} RA = helper _ _
- where
- open RawApplicative RA
-
- helper : ∀ n (P : Fin n → Set) → (∀ i → F (P i)) → F (∀ i → P i)
- helper zero P ∀iPi = pure (λ())
- helper (suc n) P ∀iPi =
- combine <$> ∀iPi zero ⊛ helper n (λ n → P (suc n)) (∀iPi ∘ suc)
- where
- combine : P zero → (∀ i → P (suc i)) → ∀ i → P i
- combine z s zero = z
- combine z s (suc i) = s i
-
-private
-
- -- Included just to show that sequence above has an inverse (under
- -- 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 →
- F (∀ i → P i) → ∀ i → F (P i)
- sequence⁻¹ RF F∀iPi i = (λ f → f i) <$> F∀iPi
- where open RawFunctor RF
+open import Data.Fin.Properties public
diff --git a/src/Data/Fin/Subset.agda b/src/Data/Fin/Subset.agda
index 42a61ea..73578d7 100644
--- a/src/Data/Fin/Subset.agda
+++ b/src/Data/Fin/Subset.agda
@@ -7,8 +7,8 @@
module Data.Fin.Subset where
open import Algebra
-import Algebra.Props.BooleanAlgebra as BoolAlgProp
-import Algebra.Props.BooleanAlgebra.Expression as BAExpr
+import Algebra.Properties.BooleanAlgebra as BoolAlgProp
+import Algebra.Properties.BooleanAlgebra.Expression as BAExpr
import Data.Bool.Properties as BoolProp
open import Data.Fin
open import Data.List as List using (List)
diff --git a/src/Data/Fin/Subset/Properties.agda b/src/Data/Fin/Subset/Properties.agda
new file mode 100644
index 0000000..722a076
--- /dev/null
+++ b/src/Data/Fin/Subset/Properties.agda
@@ -0,0 +1,156 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Some properties about subsets
+------------------------------------------------------------------------
+
+module Data.Fin.Subset.Properties where
+
+open import Algebra
+import Algebra.Properties.BooleanAlgebra as BoolProp
+open import Data.Empty using (⊥-elim)
+open import Data.Fin using (Fin); open Data.Fin.Fin
+open import Data.Fin.Subset
+open import Data.Nat using (ℕ)
+open import Data.Product
+open import Data.Sum as Sum
+open import Data.Vec hiding (_∈_)
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence
+ using (_⇔_; equivalence; module Equivalence)
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+
+------------------------------------------------------------------------
+-- Constructor mangling
+
+drop-there : ∀ {s n x} {p : Subset n} → suc x ∈ s ∷ p → x ∈ p
+drop-there (there x∈p) = x∈p
+
+drop-∷-⊆ : ∀ {n s₁ s₂} {p₁ p₂ : Subset n} → s₁ ∷ p₁ ⊆ s₂ ∷ p₂ → p₁ ⊆ p₂
+drop-∷-⊆ p₁s₁⊆p₂s₂ x∈p₁ = drop-there $ p₁s₁⊆p₂s₂ (there x∈p₁)
+
+drop-∷-Empty : ∀ {n s} {p : Subset n} → Empty (s ∷ p) → Empty p
+drop-∷-Empty ¬∃∈ (x , x∈p) = ¬∃∈ (suc x , there x∈p)
+
+------------------------------------------------------------------------
+-- Properties involving ⊥
+
+∉⊥ : ∀ {n} {x : Fin n} → x ∉ ⊥
+∉⊥ (there p) = ∉⊥ p
+
+⊥⊆ : ∀ {n} {p : Subset n} → ⊥ ⊆ p
+⊥⊆ x∈⊥ with ∉⊥ x∈⊥
+... | ()
+
+Empty-unique : ∀ {n} {p : Subset n} →
+ Empty p → p ≡ ⊥
+Empty-unique {p = []} ¬∃∈ = P.refl
+Empty-unique {p = s ∷ p} ¬∃∈ with Empty-unique (drop-∷-Empty ¬∃∈)
+Empty-unique {p = outside ∷ .⊥} ¬∃∈ | P.refl = P.refl
+Empty-unique {p = inside ∷ .⊥} ¬∃∈ | P.refl =
+ ⊥-elim (¬∃∈ (zero , here))
+
+------------------------------------------------------------------------
+-- Properties involving ⊤
+
+∈⊤ : ∀ {n} {x : Fin n} → x ∈ ⊤
+∈⊤ {x = zero} = here
+∈⊤ {x = suc x} = there ∈⊤
+
+⊆⊤ : ∀ {n} {p : Subset n} → p ⊆ ⊤
+⊆⊤ = const ∈⊤
+
+------------------------------------------------------------------------
+-- A property involving ⁅_⁆
+
+x∈⁅y⁆⇔x≡y : ∀ {n} {x y : Fin n} → x ∈ ⁅ y ⁆ ⇔ x ≡ y
+x∈⁅y⁆⇔x≡y {x = x} {y} =
+ 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
+ to (suc y) (there p) = P.cong suc (to y p)
+ to zero here = P.refl
+ to zero (there p) with ∉⊥ p
+ ... | ()
+
+ x∈⁅x⁆ : ∀ {n} (x : Fin n) → x ∈ ⁅ x ⁆
+ x∈⁅x⁆ zero = here
+ x∈⁅x⁆ (suc x) = there (x∈⁅x⁆ x)
+
+------------------------------------------------------------------------
+-- A property involving _∪_
+
+∪⇔⊎ : ∀ {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 [] [] ()
+ to (inside ∷ p₁) (s₂ ∷ p₂) here = inj₁ here
+ to (outside ∷ p₁) (inside ∷ p₂) here = inj₂ here
+ to (s₁ ∷ p₁) (s₂ ∷ p₂) (there x∈p₁∪p₂) =
+ Sum.map there there (to p₁ p₂ x∈p₁∪p₂)
+
+ ⊆∪ˡ : ∀ {n p₁} (p₂ : Subset n) → p₁ ⊆ p₁ ∪ p₂
+ ⊆∪ˡ [] ()
+ ⊆∪ˡ (s ∷ p₂) here = here
+ ⊆∪ˡ (s ∷ p₂) (there x∈p₁) = there (⊆∪ˡ p₂ x∈p₁)
+
+ ⊆∪ʳ : ∀ {n} (p₁ p₂ : Subset n) → p₂ ⊆ p₁ ∪ p₂
+ ⊆∪ʳ p₁ p₂ rewrite BooleanAlgebra.∨-comm (booleanAlgebra _) p₁ p₂
+ = ⊆∪ˡ p₁
+
+ from : ∀ {n} {p₁ p₂ : Subset n} {x} → x ∈ p₁ ⊎ x ∈ p₂ → x ∈ p₁ ∪ p₂
+ from (inj₁ x∈p₁) = ⊆∪ˡ _ x∈p₁
+ from (inj₂ x∈p₂) = ⊆∪ʳ _ _ x∈p₂
+
+------------------------------------------------------------------------
+-- _⊆_ is a partial order
+
+-- The "natural poset" associated with the boolean algebra.
+
+module NaturalPoset where
+ private
+ open module BA {n} = BoolProp (booleanAlgebra n) public
+ using (poset)
+ open module Po {n} = Poset (poset {n = n}) public
+
+ -- _⊆_ is equivalent to the natural lattice order.
+
+ orders-equivalent : ∀ {n} {p₁ p₂ : Subset n} → p₁ ⊆ p₂ ⇔ p₁ ≤ p₂
+ orders-equivalent = equivalence (to _ _) (from _ _)
+ where
+ to : ∀ {n} (p₁ p₂ : Subset n) → p₁ ⊆ p₂ → p₁ ≤ p₂
+ to [] [] p₁⊆p₂ = P.refl
+ to (inside ∷ p₁) (_ ∷ p₂) p₁⊆p₂ with p₁⊆p₂ here
+ to (inside ∷ p₁) (.inside ∷ p₂) p₁⊆p₂ | here = P.cong (_∷_ inside) (to p₁ p₂ (drop-∷-⊆ p₁⊆p₂))
+ to (outside ∷ p₁) (_ ∷ p₂) p₁⊆p₂ = P.cong (_∷_ outside) (to p₁ p₂ (drop-∷-⊆ p₁⊆p₂))
+
+ from : ∀ {n} (p₁ p₂ : Subset n) → p₁ ≤ p₂ → p₁ ⊆ p₂
+ from [] [] p₁≤p₂ x = x
+ from (.inside ∷ _) (_ ∷ _) p₁≤p₂ here rewrite P.cong head p₁≤p₂ = here
+ from (_ ∷ p₁) (_ ∷ p₂) p₁≤p₂ (there xs[i]=x) =
+ there (from p₁ p₂ (P.cong tail p₁≤p₂) xs[i]=x)
+
+-- _⊆_ is a partial order.
+
+poset : ℕ → Poset _ _ _
+poset n = record
+ { Carrier = Subset n
+ ; _≈_ = _≡_
+ ; _≤_ = _⊆_
+ ; isPartialOrder = record
+ { isPreorder = record
+ { isEquivalence = P.isEquivalence
+ ; reflexive = λ i≡j → from ⟨$⟩ reflexive i≡j
+ ; trans = λ x⊆y y⊆z → from ⟨$⟩ trans (to ⟨$⟩ x⊆y) (to ⟨$⟩ y⊆z)
+ }
+ ; antisym = λ x⊆y y⊆x → antisym (to ⟨$⟩ x⊆y) (to ⟨$⟩ y⊆x)
+ }
+ }
+ where
+ open NaturalPoset
+ open module E {p₁ p₂} =
+ Equivalence (orders-equivalent {n = n} {p₁ = p₁} {p₂ = p₂})
diff --git a/src/Data/Fin/Subset/Props.agda b/src/Data/Fin/Subset/Props.agda
index 698f3e2..a24d1ee 100644
--- a/src/Data/Fin/Subset/Props.agda
+++ b/src/Data/Fin/Subset/Props.agda
@@ -1,156 +1,10 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Some properties about subsets
+-- Compatibility module. Pending for removal. Use
+-- Data.Fin.Subset.Properties instead.
------------------------------------------------------------------------
module Data.Fin.Subset.Props where
-open import Algebra
-import Algebra.Props.BooleanAlgebra as BoolProp
-open import Data.Empty using (⊥-elim)
-open import Data.Fin using (Fin); open Data.Fin.Fin
-open import Data.Fin.Subset
-open import Data.Nat using (ℕ)
-open import Data.Product
-open import Data.Sum as Sum
-open import Data.Vec hiding (_∈_)
-open import Function
-open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence
- using (_⇔_; equivalence; module Equivalence)
-open import Relation.Binary
-open import Relation.Binary.PropositionalEquality as P using (_≡_)
-
-------------------------------------------------------------------------
--- Constructor mangling
-
-drop-there : ∀ {s n x} {p : Subset n} → suc x ∈ s ∷ p → x ∈ p
-drop-there (there x∈p) = x∈p
-
-drop-∷-⊆ : ∀ {n s₁ s₂} {p₁ p₂ : Subset n} → s₁ ∷ p₁ ⊆ s₂ ∷ p₂ → p₁ ⊆ p₂
-drop-∷-⊆ p₁s₁⊆p₂s₂ x∈p₁ = drop-there $ p₁s₁⊆p₂s₂ (there x∈p₁)
-
-drop-∷-Empty : ∀ {n s} {p : Subset n} → Empty (s ∷ p) → Empty p
-drop-∷-Empty ¬∃∈ (x , x∈p) = ¬∃∈ (suc x , there x∈p)
-
-------------------------------------------------------------------------
--- Properties involving ⊥
-
-∉⊥ : ∀ {n} {x : Fin n} → x ∉ ⊥
-∉⊥ (there p) = ∉⊥ p
-
-⊥⊆ : ∀ {n} {p : Subset n} → ⊥ ⊆ p
-⊥⊆ x∈⊥ with ∉⊥ x∈⊥
-... | ()
-
-Empty-unique : ∀ {n} {p : Subset n} →
- Empty p → p ≡ ⊥
-Empty-unique {p = []} ¬∃∈ = P.refl
-Empty-unique {p = s ∷ p} ¬∃∈ with Empty-unique (drop-∷-Empty ¬∃∈)
-Empty-unique {p = outside ∷ .⊥} ¬∃∈ | P.refl = P.refl
-Empty-unique {p = inside ∷ .⊥} ¬∃∈ | P.refl =
- ⊥-elim (¬∃∈ (zero , here))
-
-------------------------------------------------------------------------
--- Properties involving ⊤
-
-∈⊤ : ∀ {n} {x : Fin n} → x ∈ ⊤
-∈⊤ {x = zero} = here
-∈⊤ {x = suc x} = there ∈⊤
-
-⊆⊤ : ∀ {n} {p : Subset n} → p ⊆ ⊤
-⊆⊤ = const ∈⊤
-
-------------------------------------------------------------------------
--- A property involving ⁅_⁆
-
-x∈⁅y⁆⇔x≡y : ∀ {n} {x y : Fin n} → x ∈ ⁅ y ⁆ ⇔ x ≡ y
-x∈⁅y⁆⇔x≡y {x = x} {y} =
- 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
- to (suc y) (there p) = P.cong suc (to y p)
- to zero here = P.refl
- to zero (there p) with ∉⊥ p
- ... | ()
-
- x∈⁅x⁆ : ∀ {n} (x : Fin n) → x ∈ ⁅ x ⁆
- x∈⁅x⁆ zero = here
- x∈⁅x⁆ (suc x) = there (x∈⁅x⁆ x)
-
-------------------------------------------------------------------------
--- A property involving _∪_
-
-∪⇔⊎ : ∀ {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 [] [] ()
- to (inside ∷ p₁) (s₂ ∷ p₂) here = inj₁ here
- to (outside ∷ p₁) (inside ∷ p₂) here = inj₂ here
- to (s₁ ∷ p₁) (s₂ ∷ p₂) (there x∈p₁∪p₂) =
- Sum.map there there (to p₁ p₂ x∈p₁∪p₂)
-
- ⊆∪ˡ : ∀ {n p₁} (p₂ : Subset n) → p₁ ⊆ p₁ ∪ p₂
- ⊆∪ˡ [] ()
- ⊆∪ˡ (s ∷ p₂) here = here
- ⊆∪ˡ (s ∷ p₂) (there x∈p₁) = there (⊆∪ˡ p₂ x∈p₁)
-
- ⊆∪ʳ : ∀ {n} (p₁ p₂ : Subset n) → p₂ ⊆ p₁ ∪ p₂
- ⊆∪ʳ p₁ p₂ rewrite BooleanAlgebra.∨-comm (booleanAlgebra _) p₁ p₂
- = ⊆∪ˡ p₁
-
- from : ∀ {n} {p₁ p₂ : Subset n} {x} → x ∈ p₁ ⊎ x ∈ p₂ → x ∈ p₁ ∪ p₂
- from (inj₁ x∈p₁) = ⊆∪ˡ _ x∈p₁
- from (inj₂ x∈p₂) = ⊆∪ʳ _ _ x∈p₂
-
-------------------------------------------------------------------------
--- _⊆_ is a partial order
-
--- The "natural poset" associated with the boolean algebra.
-
-module NaturalPoset where
- private
- open module BA {n} = BoolProp (booleanAlgebra n) public
- using (poset)
- open module Po {n} = Poset (poset {n = n}) public
-
- -- _⊆_ is equivalent to the natural lattice order.
-
- orders-equivalent : ∀ {n} {p₁ p₂ : Subset n} → p₁ ⊆ p₂ ⇔ p₁ ≤ p₂
- orders-equivalent = equivalence (to _ _) (from _ _)
- where
- to : ∀ {n} (p₁ p₂ : Subset n) → p₁ ⊆ p₂ → p₁ ≤ p₂
- to [] [] p₁⊆p₂ = P.refl
- to (inside ∷ p₁) (_ ∷ p₂) p₁⊆p₂ with p₁⊆p₂ here
- to (inside ∷ p₁) (.inside ∷ p₂) p₁⊆p₂ | here = P.cong (_∷_ inside) (to p₁ p₂ (drop-∷-⊆ p₁⊆p₂))
- to (outside ∷ p₁) (_ ∷ p₂) p₁⊆p₂ = P.cong (_∷_ outside) (to p₁ p₂ (drop-∷-⊆ p₁⊆p₂))
-
- from : ∀ {n} (p₁ p₂ : Subset n) → p₁ ≤ p₂ → p₁ ⊆ p₂
- from [] [] p₁≤p₂ x = x
- from (.inside ∷ _) (_ ∷ _) p₁≤p₂ here rewrite P.cong head p₁≤p₂ = here
- from (_ ∷ p₁) (_ ∷ p₂) p₁≤p₂ (there xs[i]=x) =
- there (from p₁ p₂ (P.cong tail p₁≤p₂) xs[i]=x)
-
--- _⊆_ is a partial order.
-
-poset : ℕ → Poset _ _ _
-poset n = record
- { Carrier = Subset n
- ; _≈_ = _≡_
- ; _≤_ = _⊆_
- ; isPartialOrder = record
- { isPreorder = record
- { isEquivalence = P.isEquivalence
- ; reflexive = λ i≡j → from ⟨$⟩ reflexive i≡j
- ; trans = λ x⊆y y⊆z → from ⟨$⟩ trans (to ⟨$⟩ x⊆y) (to ⟨$⟩ y⊆z)
- }
- ; antisym = λ x⊆y y⊆x → antisym (to ⟨$⟩ x⊆y) (to ⟨$⟩ y⊆x)
- }
- }
- where
- open NaturalPoset
- open module E {p₁ p₂} =
- Equivalence (orders-equivalent {n = n} {p₁ = p₁} {p₂ = p₂})
+open import Data.Fin.Subset.Properties public
diff --git a/src/Data/Graph/Acyclic.agda b/src/Data/Graph/Acyclic.agda
index 0bdd5fd..01b5dfb 100644
--- a/src/Data/Graph/Acyclic.agda
+++ b/src/Data/Graph/Acyclic.agda
@@ -14,9 +14,9 @@ open import Data.Nat as Nat using (ℕ; zero; suc; _<′_)
import Data.Nat.Properties as Nat
open import Data.Fin as Fin
using (Fin; Fin′; zero; suc; #_; toℕ) renaming (_ℕ-ℕ_ to _-_)
-open import Data.Fin.Props as FP using (_≟_)
+open import Data.Fin.Properties as FP using (_≟_)
open import Data.Product as Prod using (∃; _×_; _,_)
-open import Data.Maybe
+open import Data.Maybe using (Maybe; nothing; just)
open import Data.Empty
open import Data.Unit using (⊤; tt)
open import Data.Vec as Vec using (Vec; []; _∷_)
diff --git a/src/Data/Integer.agda b/src/Data/Integer.agda
index 9d94cac..f106a9f 100644
--- a/src/Data/Integer.agda
+++ b/src/Data/Integer.agda
@@ -91,6 +91,8 @@ signAbs i = PropEq.subst SignAbs (◃-left-inverse i) $
------------------------------------------------------------------------
-- Equality is decidable
+infix 4 _≟_
+
_≟_ : Decidable {A = ℤ} _≡_
i ≟ j with Sign._≟_ (sign i) (sign j) | ℕ._≟_ ∣ i ∣ ∣ j ∣
i ≟ j | yes sign-≡ | yes abs-≡ = yes (◃-cong sign-≡ abs-≡)
diff --git a/src/Data/Integer/Properties.agda b/src/Data/Integer/Properties.agda
index e2bbb3e..19cf2d4 100644
--- a/src/Data/Integer/Properties.agda
+++ b/src/Data/Integer/Properties.agda
@@ -9,13 +9,13 @@ module Data.Integer.Properties where
open import Algebra
import Algebra.FunctionProperties
import Algebra.Morphism as Morphism
-import Algebra.Props.AbelianGroup
+import Algebra.Properties.AbelianGroup
open import Algebra.Structures
open import Data.Integer hiding (suc; _≤?_)
import Data.Integer.Addition.Properties as Add
import Data.Integer.Multiplication.Properties as Mul
open import Data.Nat
- using (ℕ; suc; zero; _∸_; _≤?_; _<_; _≥_; _≱_; s≤s; z≤n)
+ using (ℕ; suc; zero; _∸_; _≤?_; _<_; _≥_; _≱_; s≤s; z≤n; ≤-pred)
renaming (_+_ to _ℕ+_; _*_ to _ℕ*_)
open import Data.Nat.Properties as ℕ using (_*-mono_)
open import Data.Product using (proj₁; proj₂; _,_)
@@ -118,7 +118,7 @@ private
; comm = +-comm
}
- open Algebra.Props.AbelianGroup
+ open Algebra.Properties.AbelianGroup
(record { isAbelianGroup = +-isAbelianGroup })
using () renaming (⁻¹-involutive to -‿involutive)
@@ -353,3 +353,31 @@ cancel-*-right i j .(s ◃ suc n) ≢0 eq | s ◂ suc n
| sign (s₂ ◃ suc n₂) | sign-◃ s₂ n₂
... | .(suc n₁) | refl | .s₁ | refl | .(suc n₂) | refl | .s₂ | refl =
SignProp.cancel-*-right s₁ s₂ (sign-cong eq)
+
+-- Multiplication with a positive number is right cancellative (for
+-- _≤_).
+
+cancel-*-+-right-≤ : ∀ m n o → m * + suc o ≤ n * + suc o → m ≤ n
+cancel-*-+-right-≤ (-[1+ m ]) (-[1+ n ]) o (-≤- n≤m) =
+ -≤- (≤-pred (ℕ.cancel-*-right-≤ (suc n) (suc m) o (s≤s n≤m)))
+cancel-*-+-right-≤ ℤ.-[1+ _ ] (+ _) _ _ = -≤+
+cancel-*-+-right-≤ (+ 0) ℤ.-[1+ _ ] _ ()
+cancel-*-+-right-≤ (+ suc _) ℤ.-[1+ _ ] _ ()
+cancel-*-+-right-≤ (+ 0) (+ 0) _ _ = +≤+ z≤n
+cancel-*-+-right-≤ (+ 0) (+ suc _) _ _ = +≤+ z≤n
+cancel-*-+-right-≤ (+ suc _) (+ 0) _ (+≤+ ())
+cancel-*-+-right-≤ (+ suc m) (+ suc n) o (+≤+ m≤n) =
+ +≤+ (ℕ.cancel-*-right-≤ (suc m) (suc n) o m≤n)
+
+-- Multiplication with a positive number is monotone.
+
+*-+-right-mono : ∀ n → (λ x → x * + suc n) Preserves _≤_ ⟶ _≤_
+*-+-right-mono _ (-≤+ {n = 0}) = -≤+
+*-+-right-mono _ (-≤+ {n = suc _}) = -≤+
+*-+-right-mono x (-≤- n≤m) =
+ -≤- (≤-pred (s≤s n≤m *-mono ≤-refl {x = suc x}))
+*-+-right-mono _ (+≤+ {m = 0} {n = 0} m≤n) = +≤+ m≤n
+*-+-right-mono _ (+≤+ {m = 0} {n = suc _} m≤n) = +≤+ z≤n
+*-+-right-mono _ (+≤+ {m = suc _} {n = 0} ())
+*-+-right-mono x (+≤+ {m = suc _} {n = suc _} m≤n) =
+ +≤+ (m≤n *-mono ≤-refl {x = suc x})
diff --git a/src/Data/List/All/Properties.agda b/src/Data/List/All/Properties.agda
index 03e08ba..9752b20 100644
--- a/src/Data/List/All/Properties.agda
+++ b/src/Data/List/All/Properties.agda
@@ -1,23 +1,27 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Properties relating All to various list functions
+-- Properties related to All
------------------------------------------------------------------------
module Data.List.All.Properties where
open import Data.Bool
open import Data.Bool.Properties
+open import Data.Empty
open import Data.List as List
import Data.List.Any as Any; open Any.Membership-≡
open import Data.List.All as All using (All; []; _∷_)
+open import Data.List.Any using (Any; here; there)
open import Data.Product as Prod
open import Function
open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence using (module Equivalence)
+open import Function.Equivalence using (_⇔_; module Equivalence)
open import Function.Inverse using (_↔_)
+open import Function.Surjection using (_↠_)
open import Relation.Binary.PropositionalEquality as P using (_≡_)
-open import Relation.Unary using () renaming (_⊆_ to _⋐_)
+open import Relation.Nullary
+open import Relation.Unary using (Decidable) renaming (_⊆_ to _⋐_)
-- Functions can be shifted between the predicate and the list.
@@ -100,3 +104,64 @@ private
; right-inverse-of = ++⁺∘++⁻ xs
}
}
+
+-- Three lemmas relating Any, All and negation.
+
+¬Any↠All¬ : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
+ ¬ Any P xs ↠ All (¬_ ∘ P) xs
+¬Any↠All¬ {P = P} = record
+ { to = P.→-to-⟶ (to _)
+ ; surjective = record
+ { from = P.→-to-⟶ from
+ ; right-inverse-of = to∘from
+ }
+ }
+ where
+ to : ∀ xs → ¬ Any P xs → All (¬_ ∘ P) xs
+ to [] ¬p = []
+ to (x ∷ xs) ¬p = ¬p ∘ here ∷ to xs (¬p ∘ there)
+
+ from : ∀ {xs} → All (¬_ ∘ P) xs → ¬ Any P xs
+ from [] ()
+ from (¬p ∷ _) (here p) = ¬p p
+ from (_ ∷ ¬p) (there p) = from ¬p p
+
+ to∘from : ∀ {xs} (¬p : All (¬_ ∘ P) xs) → to xs (from ¬p) ≡ ¬p
+ to∘from [] = P.refl
+ to∘from (¬p ∷ ¬ps) = P.cong₂ _∷_ P.refl (to∘from ¬ps)
+
+ -- If equality of functions were extensional, then the surjection
+ -- could be strengthened to a bijection.
+
+ from∘to : P.Extensionality _ _ →
+ ∀ xs → (¬p : ¬ Any P xs) → from (to xs ¬p) ≡ ¬p
+ from∘to ext [] ¬p = ext λ ()
+ from∘to ext (x ∷ xs) ¬p = ext λ
+ { (here p) → P.refl
+ ; (there p) → P.cong (λ f → f p) $ from∘to ext xs (¬p ∘ there)
+ }
+
+Any¬→¬All : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
+ Any (¬_ ∘ P) xs → ¬ All P xs
+Any¬→¬All (here ¬p) = ¬p ∘ All.head
+Any¬→¬All (there ¬p) = Any¬→¬All ¬p ∘ All.tail
+
+Any¬⇔¬All : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
+ Decidable P → Any (¬_ ∘ P) xs ⇔ ¬ All P xs
+Any¬⇔¬All {P = P} dec = record
+ { to = P.→-to-⟶ Any¬→¬All
+ ; from = P.→-to-⟶ (from _)
+ }
+ where
+ from : ∀ xs → ¬ All P xs → Any (¬_ ∘ P) xs
+ from [] ¬∀ = ⊥-elim (¬∀ [])
+ from (x ∷ xs) ¬∀ with dec x
+ ... | yes p = there (from xs (¬∀ ∘ _∷_ p))
+ ... | no ¬p = here ¬p
+
+ -- If equality of functions were extensional, then the logical
+ -- equivalence could be strengthened to a surjection.
+
+ to∘from : P.Extensionality _ _ →
+ ∀ {xs} (¬∀ : ¬ All P xs) → Any¬→¬All (from xs ¬∀) ≡ ¬∀
+ to∘from ext ¬∀ = ext (⊥-elim ∘ ¬∀)
diff --git a/src/Data/List/Any/Membership.agda b/src/Data/List/Any/Membership.agda
index 321dd53..4e948ef 100644
--- a/src/Data/List/Any/Membership.agda
+++ b/src/Data/List/Any/Membership.agda
@@ -32,7 +32,7 @@ open import Data.Sum as Sum
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as P
using (_≡_; refl; _≗_)
-import Relation.Binary.Props.DecTotalOrder as DTOProps
+import Relation.Binary.Properties.DecTotalOrder as DTOProperties
import Relation.Binary.Sigma.Pointwise as Σ
open import Relation.Unary using (_⟨×⟩_)
open import Relation.Nullary
@@ -221,7 +221,7 @@ finite {A = A} inj (x ∷ xs) ∈x∷xs = excluded-middle helper
module DTO = DecTotalOrder Nat.decTotalOrder
module STO = StrictTotalOrder
- (DTOProps.strictTotalOrder Nat.decTotalOrder)
+ (DTOProperties.strictTotalOrder Nat.decTotalOrder)
module CS = CommutativeSemiring NatProp.commutativeSemiring
not-x : ∀ {i} → ¬ (to ⟨$⟩ i ≡ x) → to ⟨$⟩ i ∈ xs
diff --git a/src/Data/List/Any/Properties.agda b/src/Data/List/Any/Properties.agda
index 7d9c9df..31c4f53 100644
--- a/src/Data/List/Any/Properties.agda
+++ b/src/Data/List/Any/Properties.agda
@@ -355,6 +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
diff --git a/src/Data/List/NonEmpty.agda b/src/Data/List/NonEmpty.agda
index 6668ed8..ad88455 100644
--- a/src/Data/List/NonEmpty.agda
+++ b/src/Data/List/NonEmpty.agda
@@ -6,88 +6,104 @@
module Data.List.NonEmpty where
-open import Data.Product hiding (map)
-open import Data.Nat
-open import Function
-open import Data.Vec as Vec using (Vec; []; _∷_)
-open import Data.List as List using (List; []; _∷_)
open import Category.Monad
-open import Relation.Binary.PropositionalEquality
+open import Data.Bool
+open import Data.Bool.Properties
+open import Data.List as List using (List; []; _∷_)
+open import Data.Maybe using (nothing; just)
+open import Data.Nat as Nat
+open import Data.Product using (_×_; ∃; proj₁; proj₂; _,_; ,_; uncurry)
+open import Data.Sum as Sum using (_⊎_; inj₁; inj₂)
+open import Data.Unit
+open import Data.Vec as Vec using (Vec; []; _∷_)
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence
+ using () renaming (module Equivalence to Eq)
+open import Relation.Binary.PropositionalEquality as P using (_≡_; refl)
+open import Relation.Nullary.Decidable using (⌊_⌋)
+
+------------------------------------------------------------------------
+-- Non-empty lists
+
+List⁺ : ∀ {a} → Set a → Set a
+List⁺ A = A × List A
+
+open Data.Product public using () renaming (_,_ to _∷_)
-infixr 5 _∷_ _∷ʳ_ _⁺++⁺_ _++⁺_ _⁺++_
+[_] : ∀ {a} {A : Set a} → A → List⁺ A
+[ x ] = x ∷ []
-data List⁺ {a} (A : Set a) : Set a where
- [_] : (x : A) → List⁺ A
- _∷_ : (x : A) (xs : List⁺ A) → List⁺ A
+infixr 5 _∷⁺_
-length_-1 : ∀ {a} {A : Set a} → List⁺ A → ℕ
-length [ x ] -1 = 0
-length x ∷ xs -1 = 1 + length xs -1
+_∷⁺_ : ∀ {a} {A : Set a} → A → List⁺ A → List⁺ A
+x ∷⁺ y ∷ xs = x ∷ y ∷ xs
+
+length : ∀ {a} {A : Set a} → List⁺ A → ℕ
+length (x ∷ xs) = suc (List.length xs)
------------------------------------------------------------------------
-- Conversion
+toList : ∀ {a} {A : Set a} → List⁺ A → List A
+toList (x ∷ xs) = x ∷ xs
+
fromVec : ∀ {n a} {A : Set a} → Vec A (suc n) → List⁺ A
-fromVec {zero} (x ∷ _) = [ x ]
-fromVec {suc n} (x ∷ xs) = x ∷ fromVec xs
+fromVec (x ∷ xs) = x ∷ Vec.toList xs
-toVec : ∀ {a} {A : Set a} (xs : List⁺ A) → Vec A (suc (length xs -1))
-toVec [ x ] = Vec.[_] x
-toVec (x ∷ xs) = x ∷ toVec xs
+toVec : ∀ {a} {A : Set a} (xs : List⁺ A) → Vec A (length xs)
+toVec (x ∷ xs) = x ∷ Vec.fromList xs
lift : ∀ {a b} {A : Set a} {B : Set b} →
(∀ {m} → Vec A (suc m) → ∃ λ n → Vec B (suc n)) →
List⁺ A → List⁺ B
lift f xs = fromVec (proj₂ (f (toVec xs)))
-fromList : ∀ {a} {A : Set a} → A → List A → List⁺ A
-fromList x xs = fromVec (Vec.fromList (x ∷ xs))
-
-toList : ∀ {a} {A : Set a} → List⁺ A → List A
-toList {a} = Vec.toList {a} ∘ toVec
-
------------------------------------------------------------------------
-- Other operations
head : ∀ {a} {A : Set a} → List⁺ A → A
-head {a} = Vec.head {a} ∘ toVec
+head (x ∷ xs) = x
tail : ∀ {a} {A : Set a} → List⁺ A → List A
-tail {a} = Vec.toList {a} ∘ Vec.tail {a} ∘ toVec
+tail (x ∷ xs) = xs
map : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → List⁺ A → List⁺ B
-map f = lift (λ xs → (, Vec.map f xs))
+map f (x ∷ xs) = (f x ∷ List.map f xs)
-- Right fold. Note that s is only applied to the last element (see
-- the examples below).
foldr : ∀ {a b} {A : Set a} {B : Set b} →
(A → B → B) → (A → B) → List⁺ A → B
-foldr c s [ x ] = s x
-foldr c s (x ∷ xs) = c x (foldr c s xs)
+foldr {A = A} {B} c s = uncurry foldr′
+ where
+ foldr′ : A → List A → B
+ foldr′ x [] = s x
+ foldr′ x (y ∷ xs) = c x (foldr′ y xs)
-- Left fold. Note that s is only applied to the first element (see
-- the examples below).
foldl : ∀ {a b} {A : Set a} {B : Set b} →
(B → A → B) → (A → B) → List⁺ A → B
-foldl c s [ x ] = s x
-foldl c s (x ∷ xs) = foldl c (c (s x)) xs
+foldl c s (x ∷ xs) = List.foldl c (s x) xs
-- Append (several variants).
+infixr 5 _⁺++⁺_ _++⁺_ _⁺++_
+
_⁺++⁺_ : ∀ {a} {A : Set a} → List⁺ A → List⁺ A → List⁺ A
-xs ⁺++⁺ ys = foldr _∷_ (λ x → x ∷ ys) xs
+(x ∷ xs) ⁺++⁺ (y ∷ ys) = x ∷ (xs List.++ y ∷ ys)
_⁺++_ : ∀ {a} {A : Set a} → List⁺ A → List A → List⁺ A
-xs ⁺++ ys = foldr _∷_ (λ x → fromList x ys) xs
+(x ∷ xs) ⁺++ ys = x ∷ (xs List.++ ys)
_++⁺_ : ∀ {a} {A : Set a} → List A → List⁺ A → List⁺ A
-xs ++⁺ ys = List.foldr _∷_ ys xs
+xs ++⁺ ys = List.foldr _∷⁺_ ys xs
concat : ∀ {a} {A : Set a} → List⁺ (List⁺ A) → List⁺ A
-concat [ xs ] = xs
-concat (xs ∷ xss) = xs ⁺++⁺ concat xss
+concat (xs ∷ xss) = xs ⁺++ List.concat (List.map toList xss)
monad : ∀ {f} → RawMonad (List⁺ {a = f})
monad = record
@@ -100,30 +116,74 @@ reverse = lift (,_ ∘′ Vec.reverse)
-- Snoc.
-_∷ʳ_ : ∀ {a} {A : Set a} → List⁺ A → A → List⁺ A
-xs ∷ʳ x = foldr _∷_ (λ y → y ∷ [ x ]) xs
+infixl 5 _∷ʳ_ _⁺∷ʳ_
+
+_∷ʳ_ : ∀ {a} {A : Set a} → List A → A → List⁺ A
+[] ∷ʳ y = [ y ]
+(x ∷ xs) ∷ʳ y = x ∷ (xs List.∷ʳ y)
+
+_⁺∷ʳ_ : ∀ {a} {A : Set a} → List⁺ A → A → List⁺ A
+xs ⁺∷ʳ x = toList xs ∷ʳ x
-- A snoc-view of non-empty lists.
infixl 5 _∷ʳ′_
data SnocView {a} {A : Set a} : List⁺ A → Set a where
- [_] : (x : A) → SnocView [ x ]
- _∷ʳ′_ : (xs : List⁺ A) (x : A) → SnocView (xs ∷ʳ x)
+ _∷ʳ′_ : (xs : List A) (x : A) → SnocView (xs ∷ʳ x)
snocView : ∀ {a} {A : Set a} (xs : List⁺ A) → SnocView xs
-snocView [ x ] = [ x ]
-snocView (x ∷ xs) with snocView xs
-snocView (x ∷ .([ y ])) | [ y ] = [ x ] ∷ʳ′ y
-snocView (x ∷ .(ys ∷ʳ y)) | ys ∷ʳ′ y = (x ∷ ys) ∷ʳ′ y
+snocView (x ∷ xs) with List.initLast xs
+snocView (x ∷ .[]) | [] = [] ∷ʳ′ x
+snocView (x ∷ .(xs List.∷ʳ y)) | xs List.∷ʳ' y = (x ∷ xs) ∷ʳ′ y
-- The last element in the list.
last : ∀ {a} {A : Set a} → List⁺ A → A
last xs with snocView xs
-last .([ y ]) | [ y ] = y
last .(ys ∷ʳ y) | ys ∷ʳ′ y = y
+-- Groups all contiguous elements for which the predicate returns the
+-- same result into lists.
+
+split : ∀ {a} {A : Set a}
+ (p : A → Bool) → List A →
+ List (List⁺ (∃ (T ∘ p)) ⊎ List⁺ (∃ (T ∘ not ∘ p)))
+split p [] = []
+split p (x ∷ xs) with p x | P.inspect p x | split p xs
+... | true | P.[ px≡t ] | inj₁ xs′ ∷ xss = inj₁ ((x , Eq.from T-≡ ⟨$⟩ px≡t) ∷⁺ xs′) ∷ xss
+... | true | P.[ px≡t ] | xss = inj₁ [ x , Eq.from T-≡ ⟨$⟩ px≡t ] ∷ xss
+... | false | P.[ px≡f ] | inj₂ xs′ ∷ xss = inj₂ ((x , Eq.from T-not-≡ ⟨$⟩ px≡f) ∷⁺ xs′) ∷ xss
+... | false | P.[ px≡f ] | xss = inj₂ [ x , Eq.from T-not-≡ ⟨$⟩ px≡f ] ∷ xss
+
+-- If we flatten the list returned by split, then we get the list we
+-- started with.
+
+flatten : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} →
+ List (List⁺ (∃ P) ⊎ List⁺ (∃ Q)) → List A
+flatten = List.concat ∘
+ List.map Sum.[ toList ∘ map proj₁ , toList ∘ map proj₁ ]
+
+flatten-split :
+ ∀ {a} {A : Set a}
+ (p : A → Bool) (xs : List A) → flatten (split p xs) ≡ xs
+flatten-split p [] = refl
+flatten-split p (x ∷ xs)
+ with p x | P.inspect p x | split p xs | flatten-split p xs
+... | true | P.[ _ ] | [] | hyp = P.cong (_∷_ x) hyp
+... | true | P.[ _ ] | inj₁ _ ∷ _ | hyp = P.cong (_∷_ x) hyp
+... | true | P.[ _ ] | inj₂ _ ∷ _ | hyp = P.cong (_∷_ x) hyp
+... | false | P.[ _ ] | [] | hyp = P.cong (_∷_ x) hyp
+... | false | P.[ _ ] | inj₁ _ ∷ _ | hyp = P.cong (_∷_ x) hyp
+... | false | P.[ _ ] | inj₂ _ ∷ _ | hyp = P.cong (_∷_ x) hyp
+
+-- Groups all contiguous elements /not/ satisfying the predicate into
+-- lists. Elements satisfying the predicate are dropped.
+
+wordsBy : ∀ {a} {A : Set a} → (A → Bool) → List A → List (List⁺ A)
+wordsBy p =
+ List.gfilter Sum.[ const nothing , just ∘′ map proj₁ ] ∘ split p
+
------------------------------------------------------------------------
-- Examples
@@ -138,39 +198,69 @@ private
(a b c : A)
where
- hd : head (a ∷ b ∷ [ c ]) ≡ a
+ hd : head (a ∷⁺ b ∷⁺ [ c ]) ≡ a
hd = refl
- tl : tail (a ∷ b ∷ [ c ]) ≡ b ∷ c ∷ []
+ tl : tail (a ∷⁺ b ∷⁺ [ c ]) ≡ b ∷ c ∷ []
tl = refl
- mp : map f (a ∷ b ∷ [ c ]) ≡ f a ∷ f b ∷ [ f c ]
+ mp : map f (a ∷⁺ b ∷⁺ [ c ]) ≡ f a ∷⁺ f b ∷⁺ [ f c ]
mp = refl
- right : foldr _⊕_ f (a ∷ b ∷ [ c ]) ≡ a ⊕ (b ⊕ f c)
+ right : foldr _⊕_ f (a ∷⁺ b ∷⁺ [ c ]) ≡ a ⊕ (b ⊕ f c)
right = refl
- left : foldl _⊗_ f (a ∷ b ∷ [ c ]) ≡ (f a ⊗ b) ⊗ c
+ left : foldl _⊗_ f (a ∷⁺ b ∷⁺ [ c ]) ≡ (f a ⊗ b) ⊗ c
left = refl
- ⁺app⁺ : (a ∷ b ∷ [ c ]) ⁺++⁺ (b ∷ [ c ]) ≡
- a ∷ b ∷ c ∷ b ∷ [ c ]
+ ⁺app⁺ : (a ∷⁺ b ∷⁺ [ c ]) ⁺++⁺ (b ∷⁺ [ c ]) ≡
+ a ∷⁺ b ∷⁺ c ∷⁺ b ∷⁺ [ c ]
⁺app⁺ = refl
- ⁺app : (a ∷ b ∷ [ c ]) ⁺++ (b ∷ c ∷ []) ≡
- a ∷ b ∷ c ∷ b ∷ [ c ]
+ ⁺app : (a ∷⁺ b ∷⁺ [ c ]) ⁺++ (b ∷ c ∷ []) ≡
+ a ∷⁺ b ∷⁺ c ∷⁺ b ∷⁺ [ c ]
⁺app = refl
- app⁺ : (a ∷ b ∷ c ∷ []) ++⁺ (b ∷ [ c ]) ≡
- a ∷ b ∷ c ∷ b ∷ [ c ]
+ app⁺ : (a ∷ b ∷ c ∷ []) ++⁺ (b ∷⁺ [ c ]) ≡
+ a ∷⁺ b ∷⁺ c ∷⁺ b ∷⁺ [ c ]
app⁺ = refl
- conc : concat ((a ∷ b ∷ [ c ]) ∷ [ b ∷ [ c ] ]) ≡
- a ∷ b ∷ c ∷ b ∷ [ c ]
+ conc : concat ((a ∷⁺ b ∷⁺ [ c ]) ∷⁺ [ b ∷⁺ [ c ] ]) ≡
+ a ∷⁺ b ∷⁺ c ∷⁺ b ∷⁺ [ c ]
conc = refl
- rev : reverse (a ∷ b ∷ [ c ]) ≡ c ∷ b ∷ [ a ]
+ rev : reverse (a ∷⁺ b ∷⁺ [ c ]) ≡ c ∷⁺ b ∷⁺ [ a ]
rev = refl
- snoc : (a ∷ b ∷ [ c ]) ∷ʳ a ≡ a ∷ b ∷ c ∷ [ a ]
+ snoc : (a ∷ b ∷ c ∷ []) ∷ʳ a ≡ a ∷⁺ b ∷⁺ c ∷⁺ [ a ]
snoc = refl
+
+ snoc⁺ : (a ∷⁺ b ∷⁺ [ c ]) ⁺∷ʳ a ≡ a ∷⁺ b ∷⁺ c ∷⁺ [ a ]
+ snoc⁺ = refl
+
+ split-true : split (const true) (a ∷ b ∷ c ∷ []) ≡
+ inj₁ ((a , tt) ∷⁺ (b , tt) ∷⁺ [ c , tt ]) ∷ []
+ split-true = refl
+
+ split-false : split (const false) (a ∷ b ∷ c ∷ []) ≡
+ inj₂ ((a , tt) ∷⁺ (b , tt) ∷⁺ [ c , tt ]) ∷ []
+ split-false = refl
+
+ split-≡1 :
+ split (λ n → ⌊ n Nat.≟ 1 ⌋) (1 ∷ 2 ∷ 3 ∷ 1 ∷ 1 ∷ 2 ∷ 1 ∷ []) ≡
+ inj₁ [ 1 , tt ] ∷ inj₂ ((2 , tt) ∷⁺ [ 3 , tt ]) ∷
+ inj₁ ((1 , tt) ∷⁺ [ 1 , tt ]) ∷ inj₂ [ 2 , tt ] ∷ inj₁ [ 1 , tt ] ∷
+ []
+ split-≡1 = refl
+
+ wordsBy-true : wordsBy (const true) (a ∷ b ∷ c ∷ []) ≡ []
+ wordsBy-true = refl
+
+ wordsBy-false : wordsBy (const false) (a ∷ b ∷ c ∷ []) ≡
+ (a ∷⁺ b ∷⁺ [ c ]) ∷ []
+ wordsBy-false = refl
+
+ wordsBy-≡1 :
+ wordsBy (λ n → ⌊ n Nat.≟ 1 ⌋) (1 ∷ 2 ∷ 3 ∷ 1 ∷ 1 ∷ 2 ∷ 1 ∷ []) ≡
+ (2 ∷⁺ [ 3 ]) ∷ [ 2 ] ∷ []
+ wordsBy-≡1 = refl
diff --git a/src/Data/List/NonEmpty/Properties.agda b/src/Data/List/NonEmpty/Properties.agda
index 812018f..2d1d6cc 100644
--- a/src/Data/List/NonEmpty/Properties.agda
+++ b/src/Data/List/NonEmpty/Properties.agda
@@ -6,17 +6,15 @@
module Data.List.NonEmpty.Properties where
-open import Algebra
open import Category.Monad
open import Data.List as List using (List; []; _∷_; _++_)
open import Data.List.NonEmpty as List⁺
-open import Data.Product
+open import Data.List.Properties
open import Function
open import Relation.Binary.PropositionalEquality
open ≡-Reasoning
private
- module LM {a} {A : Set a} = Monoid (List.monoid A)
open module LMo {a} =
RawMonad {f = a} List.monad
using () renaming (_>>=_ to _⋆>>=_)
@@ -25,33 +23,26 @@ private
η : ∀ {a} {A : Set a}
(xs : List⁺ A) → head xs ∷ tail xs ≡ List⁺.toList xs
-η [ x ] = refl
-η (x ∷ xs) = refl
+η _ = refl
toList-fromList : ∀ {a} {A : Set a} x (xs : List A) →
- x ∷ xs ≡ List⁺.toList (List⁺.fromList x xs)
-toList-fromList x [] = refl
-toList-fromList x (y ∷ xs) = cong (_∷_ x) (toList-fromList y xs)
+ x ∷ xs ≡ List⁺.toList (x ∷ xs)
+toList-fromList _ _ = refl
toList-⁺++ : ∀ {a} {A : Set a} (xs : List⁺ A) ys →
List⁺.toList xs ++ ys ≡
List⁺.toList (xs ⁺++ ys)
-toList-⁺++ [ x ] ys = toList-fromList x ys
-toList-⁺++ (x ∷ xs) ys = cong (_∷_ x) (toList-⁺++ xs ys)
+toList-⁺++ _ _ = refl
toList-⁺++⁺ : ∀ {a} {A : Set a} (xs ys : List⁺ A) →
List⁺.toList xs ++ List⁺.toList ys ≡
List⁺.toList (xs ⁺++⁺ ys)
-toList-⁺++⁺ [ x ] ys = refl
-toList-⁺++⁺ (x ∷ xs) ys = cong (_∷_ x) (toList-⁺++⁺ xs ys)
+toList-⁺++⁺ _ _ = refl
toList->>= : ∀ {ℓ} {A B : Set ℓ}
(f : A → List⁺ B) (xs : List⁺ A) →
(List⁺.toList xs ⋆>>= List⁺.toList ∘ f) ≡
(List⁺.toList (xs >>= f))
-toList->>= {ℓ} f [ x ] =
- proj₂ {a = ℓ} {b = ℓ} LM.identity _
toList->>= f (x ∷ xs) = begin
- List⁺.toList (f x) ++ (List⁺.toList xs ⋆>>= List⁺.toList ∘ f) ≡⟨ cong (_++_ (List⁺.toList (f x))) (toList->>= f xs) ⟩
- List⁺.toList (f x) ++ List⁺.toList (xs >>= f) ≡⟨ toList-⁺++⁺ (f x) (xs >>= f) ⟩
- List⁺.toList (f x ⁺++⁺ (xs >>= f)) ∎
+ List.concat (List.map (List⁺.toList ∘ f) (x ∷ xs)) ≡⟨ cong List.concat $ map-compose {g = List⁺.toList} {f = f} (x ∷ xs) ⟩
+ List.concat (List.map List⁺.toList (List.map f (x ∷ xs))) ∎
diff --git a/src/Data/List/Properties.agda b/src/Data/List/Properties.agda
index 1665830..2280e04 100644
--- a/src/Data/List/Properties.agda
+++ b/src/Data/List/Properties.agda
@@ -10,6 +10,7 @@
module Data.List.Properties where
open import Algebra
+import Algebra.Monoid-solver
open import Category.Monad
open import Data.Bool
open import Data.List as List
@@ -33,6 +34,9 @@ private
open module LMP {ℓ} = RawMonadPlus (List.monadPlus {ℓ = ℓ})
module LM {a} {A : Set a} = Monoid (List.monoid A)
+module List-solver {a} {A : Set a} =
+ Algebra.Monoid-solver (monoid A) renaming (id to nil)
+
∷-injective : ∀ {a} {A : Set a} {x y : A} {xs ys} →
x ∷ xs ≡ y List.∷ ys → x ≡ y × xs ≡ ys
∷-injective refl = (refl , refl)
@@ -300,6 +304,11 @@ length-++ : ∀ {a} {A : Set a} (xs : List A) {ys} →
length-++ [] = refl
length-++ (x ∷ xs) = P.cong suc (length-++ xs)
+length-replicate : ∀ {a} {A : Set a} n {x : A} →
+ length (replicate n x) ≡ n
+length-replicate zero = refl
+length-replicate (suc n) = P.cong suc (length-replicate n)
+
length-gfilter : ∀ {a b} {A : Set a} {B : Set b} (p : A → Maybe B) xs →
length (gfilter p xs) ≤ length xs
length-gfilter p [] = z≤n
diff --git a/src/Data/M/Indexed.agda b/src/Data/M/Indexed.agda
new file mode 100644
index 0000000..efd87f9
--- /dev/null
+++ b/src/Data/M/Indexed.agda
@@ -0,0 +1,42 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Indexed M-types (the dual of indexed W-types aka Petersson-Synek
+-- trees).
+------------------------------------------------------------------------
+
+module Data.M.Indexed where
+
+open import Level
+open import Coinduction
+open import Data.Product
+open import Data.Container.Indexed.Core
+open import Function
+open import Relation.Unary
+
+-- The family of indexed M-types.
+
+module _ {ℓ c r} {O : Set ℓ} (C : Container O O c r) where
+
+ data M (o : O) : Set (ℓ ⊔ c ⊔ r) where
+ inf : ⟦ C ⟧ (∞ ∘ M) o → M o
+
+ open Container C
+
+ -- Projections.
+
+ head : M ⊆ Command
+ head (inf (c , _)) = c
+
+ tail : ∀ {o} (m : M o) (r : Response (head m)) → M (next (head m) r)
+ tail (inf (_ , k)) r = ♭ (k r)
+
+ force : M ⊆ ⟦ C ⟧ M
+ force (inf (c , k)) = c , λ r → ♭ (k r)
+
+ -- Coiteration.
+
+ coit : ∀ {ℓ} {X : Pred O ℓ} → X ⊆ ⟦ C ⟧ X → X ⊆ M
+ coit ψ x = inf (proj₁ cs , λ r → ♯ coit ψ (proj₂ cs r))
+ where
+ cs = ψ x
diff --git a/src/Data/Maybe.agda b/src/Data/Maybe.agda
index 858c601..39dc158 100644
--- a/src/Data/Maybe.agda
+++ b/src/Data/Maybe.agda
@@ -16,16 +16,25 @@ open import Data.Maybe.Core public
------------------------------------------------------------------------
-- Some operations
-open import Data.Bool using (Bool; true; false)
+open import Data.Bool using (Bool; true; false; not)
open import Data.Unit using (⊤)
+open import Function
+open import Relation.Nullary
boolToMaybe : Bool → Maybe ⊤
boolToMaybe true = just _
boolToMaybe false = nothing
-maybeToBool : ∀ {a} {A : Set a} → Maybe A → Bool
-maybeToBool (just _) = true
-maybeToBool nothing = false
+is-just : ∀ {a} {A : Set a} → Maybe A → Bool
+is-just (just _) = true
+is-just nothing = false
+
+is-nothing : ∀ {a} {A : Set a} → Maybe A → Bool
+is-nothing = not ∘ is-just
+
+decToMaybe : ∀ {a} {A : Set a} → Dec A → Maybe A
+decToMaybe (yes x) = just x
+decToMaybe (no _) = nothing
-- A dependent eliminator.
@@ -53,14 +62,16 @@ from-just nothing = _
------------------------------------------------------------------------
-- Maybe monad
-open import Function
open import Category.Functor
open import Category.Monad
open import Category.Monad.Identity
+map : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → Maybe A → Maybe B
+map f = maybe (just ∘ f) nothing
+
functor : ∀ {f} → RawFunctor (Maybe {a = f})
functor = record
- { _<$>_ = λ f → maybe (just ∘ f) nothing
+ { _<$>_ = map
}
monadT : ∀ {f} {M : Set f → Set f} →
@@ -93,7 +104,6 @@ monadPlus {f} = record
------------------------------------------------------------------------
-- Equality
-open import Relation.Nullary
open import Relation.Binary as B
data Eq {a ℓ} {A : Set a} (_≈_ : Rel A ℓ) : Rel (Maybe A) (a ⊔ ℓ) where
@@ -149,35 +159,37 @@ decSetoid D = record
------------------------------------------------------------------------
-- Any and All
+open Data.Bool using (T)
open import Data.Empty using (⊥)
import Relation.Nullary.Decidable as Dec
open import Relation.Unary as U
-data Any {a} {A : Set a} (P : A → Set a) : Maybe A → Set a where
+data Any {a p} {A : Set a} (P : A → Set p) : Maybe A → Set (a ⊔ p) where
just : ∀ {x} (px : P x) → Any P (just x)
-data All {a} {A : Set a} (P : A → Set a) : Maybe A → Set a where
+data All {a p} {A : Set a} (P : A → Set p) : Maybe A → Set (a ⊔ p) where
just : ∀ {x} (px : P x) → All P (just x)
nothing : All P nothing
-IsJust : ∀ {a} {A : Set a} → Maybe A → Set a
-IsJust = Any (λ _ → Lift ⊤)
-
-IsNothing : ∀ {a} {A : Set a} → Maybe A → Set a
-IsNothing = All (λ _ → Lift ⊥)
+Is-just : ∀ {a} {A : Set a} → Maybe A → Set a
+Is-just = Any (λ _ → ⊤)
-private
+Is-nothing : ∀ {a} {A : Set a} → Maybe A → Set a
+Is-nothing = All (λ _ → ⊥)
- drop-just-Any : ∀ {A} {P : A → Set} {x} → Any P (just x) → P x
- drop-just-Any (just px) = px
+to-witness : ∀ {p} {P : Set p} {m : Maybe P} → Is-just m → P
+to-witness (just {x = p} _) = p
- drop-just-All : ∀ {A} {P : A → Set} {x} → All P (just x) → P x
- drop-just-All (just px) = px
+to-witness-T : ∀ {p} {P : Set p} (m : Maybe P) → T (is-just m) → P
+to-witness-T (just p) _ = p
+to-witness-T nothing ()
-anyDec : ∀ {A} {P : A → Set} → U.Decidable P → U.Decidable (Any P)
+anyDec : ∀ {a p} {A : Set a} {P : A → Set p} →
+ U.Decidable P → U.Decidable (Any P)
anyDec p nothing = no λ()
-anyDec p (just x) = Dec.map′ just drop-just-Any (p x)
+anyDec p (just x) = Dec.map′ just (λ { (Any.just px) → px }) (p x)
-allDec : ∀ {A} {P : A → Set} → U.Decidable P → U.Decidable (All P)
+allDec : ∀ {a p} {A : Set a} {P : A → Set p} →
+ U.Decidable P → U.Decidable (All P)
allDec p nothing = yes nothing
-allDec p (just x) = Dec.map′ just drop-just-All (p x)
+allDec p (just x) = Dec.map′ just (λ { (All.just px) → px }) (p x)
diff --git a/src/Data/Nat.agda b/src/Data/Nat.agda
index 5029913..6c2662d 100644
--- a/src/Data/Nat.agda
+++ b/src/Data/Nat.agda
@@ -8,12 +8,12 @@ module Data.Nat where
open import Function
open import Function.Equality as F using (_⟨$⟩_)
-open import Function.Injection
- using (Injection; module Injection)
+open import Function.Injection using (_↣_)
open import Data.Sum
open import Data.Empty
import Level
open import Relation.Nullary
+import Relation.Nullary.Decidable as Dec
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; refl)
@@ -28,9 +28,7 @@ data ℕ : Set where
zero : ℕ
suc : (n : ℕ) → ℕ
-{-# BUILTIN NATURAL ℕ #-}
-{-# BUILTIN ZERO zero #-}
-{-# BUILTIN SUC suc #-}
+{-# BUILTIN NATURAL ℕ #-}
infix 4 _≤_ _<_ _≥_ _>_ _≰_ _≮_ _≱_ _≯_
@@ -192,14 +190,11 @@ compare (suc .m) (suc .(suc m + k)) | less m k = less (suc m) k
compare (suc .m) (suc .m) | equal m = equal (suc m)
compare (suc .(suc m + k)) (suc .m) | greater m k = greater (suc m) k
--- If there is an injection from a set to ℕ, then equality of the set
--- can be decided.
+-- If there is an injection from a type to ℕ, then the type has
+-- decidable equality.
-eq? : ∀ {s₁ s₂} {S : Setoid s₁ s₂} →
- Injection S (PropEq.setoid ℕ) → Decidable (Setoid._≈_ S)
-eq? inj x y with to ⟨$⟩ x ≟ to ⟨$⟩ y where open Injection inj
-... | yes tox≡toy = yes (Injection.injective inj tox≡toy)
-... | no tox≢toy = no (tox≢toy ∘ F.cong (Injection.to inj))
+eq? : ∀ {a} {A : Set a} → A ↣ ℕ → Decidable {A = A} _≡_
+eq? inj = Dec.via-injection inj _≟_
------------------------------------------------------------------------
-- Some properties
diff --git a/src/Data/Nat/Coprimality.agda b/src/Data/Nat/Coprimality.agda
index 4b697ff..7734a59 100644
--- a/src/Data/Nat/Coprimality.agda
+++ b/src/Data/Nat/Coprimality.agda
@@ -8,7 +8,7 @@ module Data.Nat.Coprimality where
open import Data.Empty
open import Data.Fin using (toℕ; fromℕ≤)
-open import Data.Fin.Props as FinProp
+open import Data.Fin.Properties as FinProp
open import Data.Nat
open import Data.Nat.Primality
import Data.Nat.Properties as NatProp
diff --git a/src/Data/Nat/DivMod.agda b/src/Data/Nat/DivMod.agda
index fab8193..1ce9ce6 100644
--- a/src/Data/Nat/DivMod.agda
+++ b/src/Data/Nat/DivMod.agda
@@ -10,7 +10,7 @@ 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ℕ)
-import Data.Fin.Props as Fin
+import Data.Fin.Properties as Fin
open import Induction.Nat
open import Relation.Nullary.Decidable
open import Relation.Binary.PropositionalEquality
diff --git a/src/Data/Nat/Divisibility.agda b/src/Data/Nat/Divisibility.agda
index 07bffd8..59850d6 100644
--- a/src/Data/Nat/Divisibility.agda
+++ b/src/Data/Nat/Divisibility.agda
@@ -10,7 +10,7 @@ open import Data.Nat as Nat
open import Data.Nat.DivMod
import Data.Nat.Properties as NatProp
open import Data.Fin as Fin using (Fin; zero; suc)
-import Data.Fin.Props as FP
+import Data.Fin.Properties as FP
open NatProp.SemiringSolver
open import Algebra
private
diff --git a/src/Data/Nat/Properties.agda b/src/Data/Nat/Properties.agda
index 04506bf..ccd5f25 100644
--- a/src/Data/Nat/Properties.agda
+++ b/src/Data/Nat/Properties.agda
@@ -26,109 +26,11 @@ open import Data.Product
open import Data.Sum
------------------------------------------------------------------------
--- (ℕ, +, *, 0, 1) is a commutative semiring
-
-private
-
- +-assoc : Associative _+_
- +-assoc zero _ _ = refl
- +-assoc (suc m) n o = cong suc $ +-assoc m n o
-
- +-identity : Identity 0 _+_
- +-identity = (λ _ → refl) , n+0≡n
- where
- n+0≡n : RightIdentity 0 _+_
- n+0≡n zero = refl
- n+0≡n (suc n) = cong suc $ n+0≡n n
-
- m+1+n≡1+m+n : ∀ m n → m + suc n ≡ suc (m + n)
- m+1+n≡1+m+n zero n = refl
- m+1+n≡1+m+n (suc m) n = cong suc (m+1+n≡1+m+n m n)
- +-comm : Commutative _+_
- +-comm zero n = sym $ proj₂ +-identity n
- +-comm (suc m) n =
- begin
- suc m + n
- ≡⟨ refl ⟩
- suc (m + n)
- ≡⟨ cong suc (+-comm m n) ⟩
- suc (n + m)
- ≡⟨ sym (m+1+n≡1+m+n n m) ⟩
- n + suc m
- ∎
-
- m*1+n≡m+mn : ∀ m n → m * suc n ≡ m + m * n
- m*1+n≡m+mn zero n = refl
- m*1+n≡m+mn (suc m) n =
- begin
- suc m * suc n
- ≡⟨ refl ⟩
- suc n + m * suc n
- ≡⟨ cong (λ x → suc n + x) (m*1+n≡m+mn m n) ⟩
- suc n + (m + m * n)
- ≡⟨ refl ⟩
- suc (n + (m + m * n))
- ≡⟨ cong suc (sym $ +-assoc n m (m * n)) ⟩
- suc (n + m + m * n)
- ≡⟨ cong (λ x → suc (x + m * n)) (+-comm n m) ⟩
- suc (m + n + m * n)
- ≡⟨ cong suc (+-assoc m n (m * n)) ⟩
- suc (m + (n + m * n))
- ≡⟨ refl ⟩
- suc m + suc m * n
- ∎
-
- *-zero : Zero 0 _*_
- *-zero = (λ _ → refl) , n*0≡0
- where
- n*0≡0 : RightZero 0 _*_
- n*0≡0 zero = refl
- n*0≡0 (suc n) = n*0≡0 n
-
- *-comm : Commutative _*_
- *-comm zero n = sym $ proj₂ *-zero n
- *-comm (suc m) n =
- begin
- suc m * n
- ≡⟨ refl ⟩
- n + m * n
- ≡⟨ cong (λ x → n + x) (*-comm m n) ⟩
- n + n * m
- ≡⟨ sym (m*1+n≡m+mn n m) ⟩
- n * suc m
- ∎
-
- distribʳ-*-+ : _*_ DistributesOverʳ _+_
- distribʳ-*-+ m zero o = refl
- distribʳ-*-+ m (suc n) o =
- begin
- (suc n + o) * m
- ≡⟨ refl ⟩
- m + (n + o) * m
- ≡⟨ cong (_+_ m) $ distribʳ-*-+ m n o ⟩
- m + (n * m + o * m)
- ≡⟨ sym $ +-assoc m (n * m) (o * m) ⟩
- m + n * m + o * m
- ≡⟨ refl ⟩
- suc n * m + o * m
- ∎
-
- *-assoc : Associative _*_
- *-assoc zero n o = refl
- *-assoc (suc m) n o =
- begin
- (suc m * n) * o
- ≡⟨ refl ⟩
- (n + m * n) * o
- ≡⟨ distribʳ-*-+ o n (m * n) ⟩
- n * o + (m * n) * o
- ≡⟨ cong (λ x → n * o + x) $ *-assoc m n o ⟩
- n * o + m * (n * o)
- ≡⟨ refl ⟩
- suc m * (n * o)
- ∎
+-- basic lemmas about (ℕ, +, *, 0, 1):
+open import Data.Nat.Properties.Simple
+-- (ℕ, +, *, 0, 1) is a commutative semiring
isCommutativeSemiring : IsCommutativeSemiring _≡_ _+_ _*_ 0 1
isCommutativeSemiring = record
{ +-isCommutativeMonoid = record
@@ -137,7 +39,7 @@ isCommutativeSemiring = record
; assoc = +-assoc
; ∙-cong = cong₂ _+_
}
- ; identityˡ = proj₁ +-identity
+ ; identityˡ = λ _ → refl
; comm = +-comm
}
; *-isCommutativeMonoid = record
@@ -146,11 +48,11 @@ isCommutativeSemiring = record
; assoc = *-assoc
; ∙-cong = cong₂ _*_
}
- ; identityˡ = proj₂ +-identity
+ ; identityˡ = +-right-identity
; comm = *-comm
}
; distribʳ = distribʳ-*-+
- ; zeroˡ = proj₁ *-zero
+ ; zeroˡ = λ _ → refl
}
commutativeSemiring : CommutativeSemiring _ _
@@ -477,7 +379,7 @@ n∸n≡0 zero = refl
n∸n≡0 (suc n) = n∸n≡0 n
∸-+-assoc : ∀ m n o → (m ∸ n) ∸ o ≡ m ∸ (n + o)
-∸-+-assoc m n zero = cong (_∸_ m) (sym $ proj₂ +-identity n)
+∸-+-assoc m n zero = cong (_∸_ m) (sym $ +-right-identity n)
∸-+-assoc zero zero (suc o) = refl
∸-+-assoc zero (suc n) (suc o) = refl
∸-+-assoc (suc m) zero (suc o) = refl
@@ -486,7 +388,7 @@ n∸n≡0 (suc n) = n∸n≡0 n
+-∸-assoc : ∀ m {n o} → o ≤ n → (m + n) ∸ o ≡ m + (n ∸ o)
+-∸-assoc m (z≤n {n = n}) = begin m + n ∎
+-∸-assoc m (s≤s {m = o} {n = n} o≤n) = begin
- (m + suc n) ∸ suc o ≡⟨ cong (λ n → n ∸ suc o) (m+1+n≡1+m+n m n) ⟩
+ (m + suc n) ∸ suc o ≡⟨ cong (λ n → n ∸ suc o) (+-suc m n) ⟩
suc (m + n) ∸ suc o ≡⟨ refl ⟩
(m + n) ∸ o ≡⟨ +-∸-assoc m o≤n ⟩
m + (n ∸ o) ∎
@@ -495,7 +397,7 @@ m+n∸n≡m : ∀ m n → (m + n) ∸ n ≡ m
m+n∸n≡m m n = begin
(m + n) ∸ n ≡⟨ +-∸-assoc m (≤-refl {x = n}) ⟩
m + (n ∸ n) ≡⟨ cong (_+_ m) (n∸n≡0 n) ⟩
- m + 0 ≡⟨ proj₂ +-identity m ⟩
+ m + 0 ≡⟨ +-right-identity m ⟩
m ∎
m+n∸m≡n : ∀ {m n} → m ≤ n → m + (n ∸ m) ≡ n
@@ -540,9 +442,9 @@ i∸k∸j+j∸k≡i+j∸k (suc i) j zero = begin
i∸k∸j+j∸k≡i+j∸k (suc i) zero (suc k) = begin
i ∸ k + 0
- ≡⟨ proj₂ +-identity _ ⟩
+ ≡⟨ +-right-identity _ ⟩
i ∸ k
- ≡⟨ cong (λ x → x ∸ k) (sym (proj₂ +-identity _)) ⟩
+ ≡⟨ cong (λ x → x ∸ k) (sym (+-right-identity _)) ⟩
i + 0 ∸ k
i∸k∸j+j∸k≡i+j∸k (suc i) (suc j) (suc k) = begin
@@ -550,7 +452,7 @@ i∸k∸j+j∸k≡i+j∸k (suc i) (suc j) (suc k) = begin
≡⟨ i∸k∸j+j∸k≡i+j∸k (suc i) j k ⟩
suc i + j ∸ k
≡⟨ cong (λ x → x ∸ k)
- (sym (m+1+n≡1+m+n i j)) ⟩
+ (sym (+-suc i j)) ⟩
i + suc j ∸ k
diff --git a/src/Data/Nat/Properties/Simple.agda b/src/Data/Nat/Properties/Simple.agda
new file mode 100644
index 0000000..6c5cb93
--- /dev/null
+++ b/src/Data/Nat/Properties/Simple.agda
@@ -0,0 +1,110 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- A bunch of properties about natural number operations
+------------------------------------------------------------------------
+
+module Data.Nat.Properties.Simple where
+
+open import Data.Nat as Nat
+open import Function
+open import Relation.Binary.PropositionalEquality as PropEq
+ using (_≡_; _≢_; refl; sym; cong; cong₂)
+open PropEq.≡-Reasoning
+open import Data.Product
+open import Data.Sum
+
+------------------------------------------------------------------------
+
++-assoc : ∀ m n o → (m + n) + o ≡ m + (n + o)
++-assoc zero _ _ = refl
++-assoc (suc m) n o = cong suc $ +-assoc m n o
+
++-right-identity : ∀ n → n + 0 ≡ n
++-right-identity zero = refl
++-right-identity (suc n) = cong suc $ +-right-identity n
+
++-suc : ∀ m n → m + suc n ≡ suc (m + n)
++-suc zero n = refl
++-suc (suc m) n = cong suc (+-suc m n)
+
++-comm : ∀ m n → m + n ≡ n + m
++-comm zero n = sym $ +-right-identity n
++-comm (suc m) n =
+ begin
+ suc m + n
+ ≡⟨ refl ⟩
+ suc (m + n)
+ ≡⟨ cong suc (+-comm m n) ⟩
+ suc (n + m)
+ ≡⟨ sym (+-suc n m) ⟩
+ n + suc m
+ ∎
+
++-*-suc : ∀ m n → m * suc n ≡ m + m * n
++-*-suc zero n = refl
++-*-suc (suc m) n =
+ begin
+ suc m * suc n
+ ≡⟨ refl ⟩
+ suc n + m * suc n
+ ≡⟨ cong (λ x → suc n + x) (+-*-suc m n) ⟩
+ suc n + (m + m * n)
+ ≡⟨ refl ⟩
+ suc (n + (m + m * n))
+ ≡⟨ cong suc (sym $ +-assoc n m (m * n)) ⟩
+ suc (n + m + m * n)
+ ≡⟨ cong (λ x → suc (x + m * n)) (+-comm n m) ⟩
+ suc (m + n + m * n)
+ ≡⟨ cong suc (+-assoc m n (m * n)) ⟩
+ suc (m + (n + m * n))
+ ≡⟨ refl ⟩
+ suc m + suc m * n
+ ∎
+
+*-right-zero : ∀ n → n * 0 ≡ 0
+*-right-zero zero = refl
+*-right-zero (suc n) = *-right-zero n
+
+*-comm : ∀ m n → m * n ≡ n * m
+*-comm zero n = sym $ *-right-zero n
+*-comm (suc m) n =
+ begin
+ suc m * n
+ ≡⟨ refl ⟩
+ n + m * n
+ ≡⟨ cong (λ x → n + x) (*-comm m n) ⟩
+ n + n * m
+ ≡⟨ sym (+-*-suc n m) ⟩
+ n * suc m
+ ∎
+
+distribʳ-*-+ : ∀ m n o → (n + o) * m ≡ n * m + o * m
+distribʳ-*-+ m zero o = refl
+distribʳ-*-+ m (suc n) o =
+ begin
+ (suc n + o) * m
+ ≡⟨ refl ⟩
+ m + (n + o) * m
+ ≡⟨ cong (_+_ m) $ distribʳ-*-+ m n o ⟩
+ m + (n * m + o * m)
+ ≡⟨ sym $ +-assoc m (n * m) (o * m) ⟩
+ m + n * m + o * m
+ ≡⟨ refl ⟩
+ suc n * m + o * m
+ ∎
+
+*-assoc : ∀ m n o → (m * n) * o ≡ m * (n * o)
+*-assoc zero n o = refl
+*-assoc (suc m) n o =
+ begin
+ (suc m * n) * o
+ ≡⟨ refl ⟩
+ (n + m * n) * o
+ ≡⟨ distribʳ-*-+ o n (m * n) ⟩
+ n * o + (m * n) * o
+ ≡⟨ cong (λ x → n * o + x) $ *-assoc m n o ⟩
+ n * o + m * (n * o)
+ ≡⟨ refl ⟩
+ suc m * (n * o)
+ ∎
diff --git a/src/Data/Rational.agda b/src/Data/Rational.agda
index 0d49196..66be57e 100644
--- a/src/Data/Rational.agda
+++ b/src/Data/Rational.agda
@@ -6,19 +6,23 @@
module Data.Rational where
+import Algebra
import Data.Bool.Properties as Bool
open import Function
-open import Data.Integer hiding (suc) renaming (_*_ to _ℤ*_)
+open import Data.Integer as ℤ using (ℤ; ∣_∣; +_; -_)
open import Data.Integer.Divisibility as ℤDiv using (Coprime)
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 Data.Nat as ℕ using (ℕ; zero; suc)
+open import Data.Sum
import Level
open import Relation.Nullary.Decidable
+open import Relation.Nullary
open import Relation.Binary
-open import Relation.Binary.PropositionalEquality as PropEq
-open ≡-Reasoning
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; refl; cong; cong₂)
+open P.≡-Reasoning
------------------------------------------------------------------------
-- The definition
@@ -72,9 +76,9 @@ private
infix 4 _≃_
_≃_ : Rel ℚ Level.zero
-p ≃ q = P.numerator ℤ* Q.denominator ≡
- Q.numerator ℤ* P.denominator
- where module P = ℚ p; module Q = ℚ q
+p ≃ q = numerator p ℤ.* denominator q ≡
+ numerator q ℤ.* denominator p
+ where open ℚ
-- _≃_ coincides with propositional equality.
@@ -82,13 +86,14 @@ p ≃ q = P.numerator ℤ* Q.denominator ≡
≡⇒≃ refl = refl
≃⇒≡ : _≃_ ⇒ _≡_
-≃⇒≡ {p} {q} = helper P.numerator P.denominator-1 P.isCoprime
- Q.numerator Q.denominator-1 Q.isCoprime
+≃⇒≡ {i = p} {j = q} =
+ helper (numerator p) (denominator-1 p) (isCoprime p)
+ (numerator q) (denominator-1 q) (isCoprime q)
where
- module P = ℚ p; module Q = ℚ q
+ open ℚ
helper : ∀ n₁ d₁ c₁ n₂ d₂ c₂ →
- n₁ ℤ* + suc d₂ ≡ n₂ ℤ* + suc d₁ →
+ n₁ ℤ.* + suc d₂ ≡ n₂ ℤ.* + suc d₁ →
(n₁ ÷ suc d₁) {c₁} ≡ (n₂ ÷ suc d₂) {c₂}
helper n₁ d₁ c₁ n₂ d₂ c₂ eq
with Poset.antisym ℕDiv.poset 1+d₁∣1+d₂ 1+d₂∣1+d₁
@@ -97,19 +102,116 @@ p ≃ q = P.numerator ℤ* Q.denominator ≡
1+d₁∣1+d₂ = ℤDiv.coprime-divisor (+ suc d₁) n₁ (+ suc d₂)
(C.sym $ toWitness c₁) $
ℕDiv.divides ∣ n₂ ∣ (begin
- ∣ n₁ ℤ* + suc d₂ ∣ ≡⟨ cong ∣_∣ eq ⟩
- ∣ n₂ ℤ* + suc d₁ ∣ ≡⟨ ℤ.abs-*-commute n₂ (+ suc d₁) ⟩
- ∣ n₂ ∣ ℕ* suc d₁ ∎)
+ ∣ n₁ ℤ.* + suc d₂ ∣ ≡⟨ cong ∣_∣ eq ⟩
+ ∣ n₂ ℤ.* + suc d₁ ∣ ≡⟨ ℤ.abs-*-commute n₂ (+ suc d₁) ⟩
+ ∣ n₂ ∣ ℕ.* suc d₁ ∎)
1+d₂∣1+d₁ : suc d₂ ∣ suc d₁
1+d₂∣1+d₁ = ℤDiv.coprime-divisor (+ suc d₂) n₂ (+ suc d₁)
(C.sym $ toWitness c₂) $
ℕDiv.divides ∣ n₁ ∣ (begin
- ∣ n₂ ℤ* + suc d₁ ∣ ≡⟨ cong ∣_∣ (PropEq.sym eq) ⟩
- ∣ n₁ ℤ* + suc d₂ ∣ ≡⟨ ℤ.abs-*-commute n₁ (+ suc d₂) ⟩
- ∣ n₁ ∣ ℕ* suc d₂ ∎)
+ ∣ n₂ ℤ.* + suc d₁ ∣ ≡⟨ cong ∣_∣ (P.sym eq) ⟩
+ ∣ n₁ ℤ.* + suc d₂ ∣ ≡⟨ ℤ.abs-*-commute n₁ (+ suc d₂) ⟩
+ ∣ n₁ ∣ ℕ.* suc d₂ ∎)
helper n₁ d c₁ n₂ .d c₂ eq | refl with ℤ.cancel-*-right
n₁ n₂ (+ suc d) (λ ()) eq
helper n d c₁ .n .d c₂ eq | refl | refl with Bool.proof-irrelevance c₁ c₂
helper n d c .n .d .c eq | refl | refl | refl = refl
+
+------------------------------------------------------------------------
+-- Equality is decidable
+
+infix 4 _≟_
+
+_≟_ : Decidable {A = ℚ} _≡_
+p ≟ q with ℚ.numerator p ℤ.* ℚ.denominator q ℤ.≟
+ ℚ.numerator q ℤ.* ℚ.denominator p
+p ≟ q | yes pq≃qp = yes (≃⇒≡ pq≃qp)
+p ≟ q | no ¬pq≃qp = no (¬pq≃qp ∘ ≡⇒≃)
+
+------------------------------------------------------------------------
+-- Ordering
+
+infix 4 _≤_ _≤?_
+
+data _≤_ : ℚ → ℚ → Set where
+ *≤* : ∀ {p q} →
+ ℚ.numerator p ℤ.* ℚ.denominator q ℤ.≤
+ ℚ.numerator q ℤ.* ℚ.denominator p →
+ p ≤ q
+
+drop-*≤* : ∀ {p q} → p ≤ q →
+ ℚ.numerator p ℤ.* ℚ.denominator q ℤ.≤
+ ℚ.numerator q ℤ.* ℚ.denominator p
+drop-*≤* (*≤* pq≤qp) = pq≤qp
+
+_≤?_ : Decidable _≤_
+p ≤? q with ℚ.numerator p ℤ.* ℚ.denominator q ℤ.≤?
+ ℚ.numerator q ℤ.* ℚ.denominator p
+p ≤? q | yes pq≤qp = yes (*≤* pq≤qp)
+p ≤? q | no ¬pq≤qp = no (λ { (*≤* pq≤qp) → ¬pq≤qp pq≤qp })
+
+decTotalOrder : DecTotalOrder _ _ _
+decTotalOrder = record
+ { Carrier = ℚ
+ ; _≈_ = _≡_
+ ; _≤_ = _≤_
+ ; isDecTotalOrder = record
+ { isTotalOrder = record
+ { isPartialOrder = record
+ { isPreorder = record
+ { isEquivalence = P.isEquivalence
+ ; reflexive = refl′
+ ; trans = trans
+ }
+ ; antisym = antisym
+ }
+ ; total = total
+ }
+ ; _≟_ = _≟_
+ ; _≤?_ = _≤?_
+ }
+ }
+ where
+ module ℤO = DecTotalOrder ℤ.decTotalOrder
+
+ refl′ : _≡_ ⇒ _≤_
+ refl′ refl = *≤* ℤO.refl
+
+ trans : Transitive _≤_
+ trans {i = p} {j = q} {k = r} (*≤* le₁) (*≤* le₂)
+ = *≤* (ℤ.cancel-*-+-right-≤ _ _ _
+ (lemma
+ (ℚ.numerator p) (ℚ.denominator p)
+ (ℚ.numerator q) (ℚ.denominator q)
+ (ℚ.numerator r) (ℚ.denominator r)
+ (ℤ.*-+-right-mono (ℚ.denominator-1 r) le₁)
+ (ℤ.*-+-right-mono (ℚ.denominator-1 p) le₂)))
+ where
+ open Algebra.CommutativeRing ℤ.commutativeRing
+
+ lemma : ∀ n₁ d₁ n₂ d₂ n₃ d₃ →
+ n₁ ℤ.* d₂ ℤ.* d₃ ℤ.≤ n₂ ℤ.* d₁ ℤ.* d₃ →
+ n₂ ℤ.* d₃ ℤ.* d₁ ℤ.≤ n₃ ℤ.* d₂ ℤ.* d₁ →
+ n₁ ℤ.* d₃ ℤ.* d₂ ℤ.≤ n₃ ℤ.* d₁ ℤ.* d₂
+ lemma n₁ d₁ n₂ d₂ n₃ d₃
+ rewrite *-assoc n₁ d₂ d₃
+ | *-comm d₂ d₃
+ | sym (*-assoc n₁ d₃ d₂)
+ | *-assoc n₃ d₂ d₁
+ | *-comm d₂ d₁
+ | sym (*-assoc n₃ d₁ d₂)
+ | *-assoc n₂ d₁ d₃
+ | *-comm d₁ d₃
+ | sym (*-assoc n₂ d₃ d₁)
+ = ℤO.trans
+
+ antisym : Antisymmetric _≡_ _≤_
+ antisym (*≤* le₁) (*≤* le₂) = ≃⇒≡ (ℤO.antisym le₁ le₂)
+
+ total : Total _≤_
+ total p q =
+ [ inj₁ ∘′ *≤* , inj₂ ∘′ *≤* ]′
+ (ℤO.total (ℚ.numerator p ℤ.* ℚ.denominator q)
+ (ℚ.numerator q ℤ.* ℚ.denominator p))
diff --git a/src/Data/Star/BoundedVec.agda b/src/Data/Star/BoundedVec.agda
index 0d78b53..7d0cfdb 100644
--- a/src/Data/Star/BoundedVec.agda
+++ b/src/Data/Star/BoundedVec.agda
@@ -18,7 +18,6 @@ open import Function
import Data.Maybe as Maybe
open import Relation.Binary
open import Relation.Binary.Consequences
-open import Category.Monad
------------------------------------------------------------------------
-- The type
@@ -44,9 +43,7 @@ _∷_ = that
↑ : ∀ {a n} → BoundedVec a n → BoundedVec a (suc n)
↑ {a} = gmap inc lift
where
- open RawMonad Maybe.monad
-
- inc = _<$>_ (map-NonEmpty suc)
+ inc = Maybe.map (map-NonEmpty suc)
lift : Pointer (λ _ → a) (λ _ → ⊤) =[ inc ]⇒
Pointer (λ _ → a) (λ _ → ⊤)
diff --git a/src/Data/Stream.agda b/src/Data/Stream.agda
index bf76626..bc8dc41 100644
--- a/src/Data/Stream.agda
+++ b/src/Data/Stream.agda
@@ -11,6 +11,7 @@ open import Data.Colist using (Colist; []; _∷_)
open import Data.Vec using (Vec; []; _∷_)
open import Data.Nat using (ℕ; zero; suc)
open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
------------------------------------------------------------------------
-- The type
@@ -95,7 +96,8 @@ _++_ : ∀ {A} → Colist A → Stream A → Stream A
infix 4 _≈_
data _≈_ {A} : Stream A → Stream A → Set where
- _∷_ : ∀ x {xs ys} (xs≈ : ∞ (♭ xs ≈ ♭ ys)) → x ∷ xs ≈ x ∷ ys
+ _∷_ : ∀ {x y xs ys}
+ (x≡ : x ≡ y) (xs≈ : ∞ (♭ xs ≈ ♭ ys)) → x ∷ xs ≈ y ∷ ys
-- x ∈ xs means that x is a member of xs.
@@ -128,23 +130,29 @@ setoid A = record
}
where
refl : Reflexive _≈_
- refl {x ∷ _} = x ∷ ♯ refl
+ refl {_ ∷ _} = P.refl ∷ ♯ refl
sym : Symmetric _≈_
- sym (x ∷ xs≈) = x ∷ ♯ sym (♭ xs≈)
+ sym (x≡ ∷ xs≈) = P.sym x≡ ∷ ♯ sym (♭ xs≈)
trans : Transitive _≈_
- trans (x ∷ xs≈) (.x ∷ ys≈) = x ∷ ♯ trans (♭ xs≈) (♭ ys≈)
+ trans (x≡ ∷ xs≈) (y≡ ∷ ys≈) = P.trans x≡ y≡ ∷ ♯ trans (♭ xs≈) (♭ ys≈)
+
+head-cong : ∀ {A} {xs ys : Stream A} → xs ≈ ys → head xs ≡ head ys
+head-cong (x≡ ∷ _) = x≡
+
+tail-cong : ∀ {A} {xs ys : Stream A} → xs ≈ ys → tail xs ≈ tail ys
+tail-cong (_ ∷ xs≈) = ♭ xs≈
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≈)
+map-cong f (x≡ ∷ xs≈) = P.cong f x≡ ∷ ♯ map-cong f (♭ xs≈)
zipWith-cong : ∀ {A B C} (_∙_ : A → B → C) {xs xs′ ys ys′} →
xs ≈ xs′ → ys ≈ ys′ →
zipWith _∙_ xs ys ≈ zipWith _∙_ xs′ ys′
-zipWith-cong _∙_ (x ∷ xs≈) (y ∷ ys≈) =
- (x ∙ y) ∷ ♯ zipWith-cong _∙_ (♭ xs≈) (♭ ys≈)
+zipWith-cong _∙_ (x≡ ∷ xs≈) (y≡ ∷ ys≈) =
+ P.cong₂ _∙_ x≡ y≡ ∷ ♯ zipWith-cong _∙_ (♭ xs≈) (♭ ys≈)
infixr 5 _⋎-cong_
diff --git a/src/Data/String.agda b/src/Data/String.agda
index f98df01..0d99778 100644
--- a/src/Data/String.agda
+++ b/src/Data/String.agda
@@ -13,6 +13,7 @@ open import Data.Char as Char using (Char)
open import Data.Bool using (Bool; true; false)
open import Function
open import Relation.Nullary
+open import Relation.Nullary.Decidable
open import Relation.Binary
open import Relation.Binary.List.StrictLex as StrictLex
import Relation.Binary.On as On
@@ -56,6 +57,12 @@ toList = primStringToList
fromList : List Char → String
fromList = primStringFromList
+toList∘fromList : ∀ s → toList (fromList s) ≡ s
+toList∘fromList s = trustMe
+
+fromList∘toList : ∀ s → fromList (toList s) ≡ s
+fromList∘toList s = trustMe
+
toVec : (s : String) → Vec Char (List.length (toList s))
toVec s = Vec.fromList (toList s)
@@ -66,17 +73,37 @@ unlines : List String → String
unlines [] = ""
unlines (x ∷ xs) = x ++ "\n" ++ unlines xs
-infix 4 _==_
-
-_==_ : String → String → Bool
-_==_ = primStringEquality
+-- Informative equality test.
_≟_ : Decidable {A = String} _≡_
-s₁ ≟ s₂ with s₁ == s₂
+s₁ ≟ s₂ with primStringEquality s₁ s₂
... | true = yes trustMe
... | false = no whatever
where postulate whatever : _
+-- Boolean equality test.
+--
+-- Why is the definition _==_ = primStringEquality not used? One
+-- reason is that the present definition can sometimes improve type
+-- inference, at least with the version of Agda that is current at the
+-- time of writing: see unit-test below.
+
+infix 4 _==_
+
+_==_ : String → String → Bool
+s₁ == s₂ = ⌊ s₁ ≟ s₂ ⌋
+
+private
+
+ -- The following unit test does not type-check (at the time of
+ -- writing) if _==_ is replaced by primStringEquality.
+
+ data P : (String → Bool) → Set where
+ p : (c : String) → P (_==_ c)
+
+ unit-test : P (_==_ "")
+ unit-test = p _
+
setoid : Setoid _ _
setoid = PropEq.setoid String
diff --git a/src/Data/Sum.agda b/src/Data/Sum.agda
index 0313256..b1fe6fa 100644
--- a/src/Data/Sum.agda
+++ b/src/Data/Sum.agda
@@ -45,10 +45,10 @@ _-⊎-_ : ∀ {a b c d} {A : Set a} {B : Set b} →
(A → B → Set c) → (A → B → Set d) → (A → B → Set (c ⊔ d))
f -⊎- g = f -[ _⊎_ ]- g
-isInj₁ : {A B : Set} → A ⊎ B → Maybe A
+isInj₁ : ∀ {a b} {A : Set a} {B : Set b} → A ⊎ B → Maybe A
isInj₁ (inj₁ x) = just x
isInj₁ (inj₂ y) = nothing
-isInj₂ : {A B : Set} → A ⊎ B → Maybe B
+isInj₂ : ∀ {a b} {A : Set a} {B : Set b} → A ⊎ B → Maybe B
isInj₂ (inj₁ x) = nothing
isInj₂ (inj₂ y) = just y
diff --git a/src/Data/Vec.agda b/src/Data/Vec.agda
index 14d45d6..923ed2c 100644
--- a/src/Data/Vec.agda
+++ b/src/Data/Vec.agda
@@ -59,7 +59,7 @@ infixl 4 _⊛_
_⊛_ : ∀ {a b n} {A : Set a} {B : Set b} →
Vec (A → B) n → Vec A n → Vec B n
-[] ⊛ [] = []
+[] ⊛ _ = []
(f ∷ fs) ⊛ (x ∷ xs) = f x ∷ (fs ⊛ xs)
replicate : ∀ {a n} {A : Set a} → A → Vec A n
diff --git a/src/Data/Vec/Properties.agda b/src/Data/Vec/Properties.agda
index b3ae337..c327b4c 100644
--- a/src/Data/Vec/Properties.agda
+++ b/src/Data/Vec/Properties.agda
@@ -11,12 +11,15 @@ open import Category.Applicative.Indexed
open import Category.Monad
open import Category.Monad.Identity
open import Data.Vec
+open import Data.List.Any
+ using (here; there) renaming (module Membership-≡ to List)
open import Data.Nat
open import Data.Empty using (⊥-elim)
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 Data.Fin.Properties using (_+′_)
open import Data.Empty using (⊥-elim)
+open import Data.Product using (_×_ ; _,_)
open import Function
open import Function.Inverse using (_↔_)
open import Relation.Binary
@@ -48,6 +51,10 @@ open import Relation.Binary.PropositionalEquality as P
using (_≡_; _≢_; refl; _≗_)
open import Relation.Binary.HeterogeneousEquality using (_≅_; refl)
+∷-injective : ∀ {a n} {A : Set a} {x y : A} {xs ys : Vec A n} →
+ (x ∷ xs) ≡ (y ∷ ys) → x ≡ y × xs ≡ ys
+∷-injective refl = refl , refl
+
-- lookup is an applicative functor morphism.
lookup-morphism :
@@ -184,6 +191,17 @@ map-lookup-allFin {n = n} xs = begin
xs ∎
where open P.≡-Reasoning
+-- tabulate f contains f i.
+
+∈-tabulate : ∀ {n a} {A : Set a} (f : Fin n → A) i → f i ∈ tabulate f
+∈-tabulate f zero = here
+∈-tabulate f (suc i) = there (∈-tabulate (f ∘ suc) i)
+
+-- allFin n contains all elements in Fin n.
+
+∈-allFin : ∀ {n} (i : Fin n) → i ∈ allFin n
+∈-allFin = ∈-tabulate id
+
-- sum commutes with _++_.
sum-++-commute : ∀ {m n} (xs : Vec ℕ m) {ys : Vec ℕ n} →
@@ -202,10 +220,10 @@ sum-++-commute (x ∷ xs) {ys} = begin
-- foldr is a congruence.
-foldr-cong : ∀ {a} {A : Set a}
- {b₁} {B₁ : ℕ → Set b₁}
+foldr-cong : ∀ {a b} {A : Set a}
+ {B₁ : ℕ → Set b}
{f₁ : ∀ {n} → A → B₁ n → B₁ (suc n)} {e₁}
- {b₂} {B₂ : ℕ → Set b₂}
+ {B₂ : ℕ → Set b}
{f₂ : ∀ {n} → A → B₂ n → B₂ (suc n)} {e₂} →
(∀ {n x} {y₁ : B₁ n} {y₂ : B₂ n} →
y₁ ≅ y₂ → f₁ x y₁ ≅ f₂ x y₂) →
@@ -218,10 +236,10 @@ foldr-cong {B₁ = B₁} f₁=f₂ e₁=e₂ (x ∷ xs) =
-- foldl is a congruence.
-foldl-cong : ∀ {a} {A : Set a}
- {b₁} {B₁ : ℕ → Set b₁}
+foldl-cong : ∀ {a b} {A : Set a}
+ {B₁ : ℕ → Set b}
{f₁ : ∀ {n} → B₁ n → A → B₁ (suc n)} {e₁}
- {b₂} {B₂ : ℕ → Set b₂}
+ {B₂ : ℕ → Set b}
{f₂ : ∀ {n} → B₂ n → A → B₂ (suc n)} {e₂} →
(∀ {n x} {y₁ : B₁ n} {y₂ : B₂ n} →
y₁ ≅ y₂ → f₁ y₁ x ≅ f₂ y₂ x) →
@@ -266,6 +284,19 @@ foldr-fusion {B = B} {f} e {C} h fuse =
idIsFold : ∀ {a n} {A : Set a} → id ≗ foldr (Vec A) {n} _∷_ []
idIsFold = foldr-universal _ _ id refl (λ _ _ → refl)
+-- The _∈_ predicate is equivalent (in the following sense) to the
+-- corresponding predicate for lists.
+
+∈⇒List-∈ : ∀ {a} {A : Set a} {n x} {xs : Vec A n} →
+ x ∈ xs → x List.∈ toList xs
+∈⇒List-∈ here = here P.refl
+∈⇒List-∈ (there x∈) = there (∈⇒List-∈ x∈)
+
+List-∈⇒∈ : ∀ {a} {A : Set a} {x : A} {xs} →
+ x List.∈ xs → x ∈ fromList xs
+List-∈⇒∈ (here P.refl) = here
+List-∈⇒∈ (there x∈) = there (List-∈⇒∈ x∈)
+
-- Proof irrelevance for _[_]=_.
proof-irrelevance-[]= : ∀ {a} {A : Set a} {n} {xs : Vec A n} {i x} →
@@ -343,3 +374,11 @@ map-[]≔ f (x ∷ xs) (suc i) = P.cong (_∷_ _) $ map-[]≔ f xs i
[]≔-lookup [] ()
[]≔-lookup (x ∷ xs) zero = refl
[]≔-lookup (x ∷ xs) (suc i) = P.cong (_∷_ x) $ []≔-lookup xs i
+
+[]≔-++-inject+ : ∀ {a} {A : Set a} {m n x}
+ (xs : Vec A m) (ys : Vec A n) i →
+ (xs ++ ys) [ Fin.inject+ n i ]≔ x ≡ xs [ i ]≔ x ++ ys
+[]≔-++-inject+ [] ys ()
+[]≔-++-inject+ (x ∷ xs) ys zero = refl
+[]≔-++-inject+ (x ∷ xs) ys (suc i) =
+ P.cong (_∷_ x) $ []≔-++-inject+ xs ys i
diff --git a/src/Data/W/Indexed.agda b/src/Data/W/Indexed.agda
new file mode 100644
index 0000000..1b615ec
--- /dev/null
+++ b/src/Data/W/Indexed.agda
@@ -0,0 +1,42 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Indexed W-types aka Petersson-Synek trees
+------------------------------------------------------------------------
+
+module Data.W.Indexed where
+
+open import Level
+open import Data.Container.Indexed.Core
+open import Data.Product
+open import Relation.Unary
+
+-- The family of indexed W-types.
+
+module _ {ℓ c r} {O : Set ℓ} (C : Container O O c r) where
+
+ open Container C
+
+ data W (o : O) : Set (ℓ ⊔ c ⊔ r) where
+ sup : ⟦ C ⟧ W o → W o
+
+ -- Projections.
+
+ head : W ⊆ Command
+ head (sup (c , _)) = c
+
+ tail : ∀ {o} (w : W o) (r : Response (head w)) → W (next (head w) r)
+ tail (sup (_ , k)) r = k r
+
+ -- Induction, (primitive) recursion and iteration.
+
+ ind : ∀ {ℓ} (P : Pred (Σ O W) ℓ) →
+ (∀ {o} (cs : ⟦ C ⟧ W o) → □ C P (o , cs) → P (o , sup cs)) →
+ ∀ {o} (w : W o) → P (o , w)
+ ind P φ (sup (c , k)) = φ (c , k) (λ r → ind P φ (k r))
+
+ rec : ∀ {ℓ} {X : Pred O ℓ} → (⟦ C ⟧ (W ∩ X) ⊆ X) → W ⊆ X
+ rec φ (sup (c , k))= φ (c , λ r → (k r , rec φ (k r)))
+
+ iter : ∀ {ℓ} {X : Pred O ℓ} → (⟦ C ⟧ X ⊆ X) → W ⊆ X
+ iter φ (sup (c , k))= φ (c , λ r → iter φ (k r))
diff --git a/src/Function.agda b/src/Function.agda
index 4fd3bc0..52bcc32 100644
--- a/src/Function.agda
+++ b/src/Function.agda
@@ -33,12 +33,12 @@ Fun₂ A = A → A → A
_∘_ : ∀ {a b c}
{A : Set a} {B : A → Set b} {C : {x : A} → B x → Set c} →
- (∀ {x} (y : B x) → C y) → (g : (x : A) → B x) →
- ((x : A) → C (g x))
+ (∀ {x} (y : B x) → C y) → (g : (x : A) → B x) →
+ ((x : A) → C (g x))
f ∘ g = λ x → f (g x)
_∘′_ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} →
- (B → C) → (A → B) → (A → C)
+ (B → C) → (A → B) → (A → C)
f ∘′ g = _∘_ f g
id : ∀ {a} {A : Set a} → A → A
diff --git a/src/Function/Equality.agda b/src/Function/Equality.agda
index 85351f8..510171b 100644
--- a/src/Function/Equality.agda
+++ b/src/Function/Equality.agda
@@ -99,3 +99,16 @@ From ⇨ To = setoid From (B.Setoid.indexedSetoid To)
; trans = λ f∼g g∼h x → trans (f∼g x) (g∼h x)
}
} where open I.Setoid To
+
+-- Parameter swapping function.
+
+flip : ∀ {a₁ a₂} {A : B.Setoid a₁ a₂}
+ {b₁ b₂} {B : B.Setoid b₁ b₂}
+ {c₁ c₂} {C : B.Setoid c₁ c₂} →
+ A ⟶ B ⇨ C → B ⟶ A ⇨ C
+flip {B = B} f = record
+ { _⟨$⟩_ = λ b → record
+ { _⟨$⟩_ = λ a → f ⟨$⟩ a ⟨$⟩ b
+ ; cong = λ a₁≈a₂ → cong f a₁≈a₂ (B.Setoid.refl B) }
+ ; cong = λ b₁≈b₂ a₁≈a₂ → cong f a₁≈a₂ b₁≈b₂
+ }
diff --git a/src/Function/Inverse.agda b/src/Function/Inverse.agda
index fd685ac..de84be7 100644
--- a/src/Function/Inverse.agda
+++ b/src/Function/Inverse.agda
@@ -13,7 +13,8 @@ open import Function.Equality as F
using (_⟶_) renaming (_∘_ to _⟪∘⟫_)
open import Function.LeftInverse as Left hiding (id; _∘_)
open import Relation.Binary
-import Relation.Binary.PropositionalEquality as P
+open import Relation.Binary.PropositionalEquality as P using (_≗_)
+open import Relation.Unary using (Pred)
-- Inverses.
@@ -65,11 +66,14 @@ record Inverse {f₁ f₂ t₁ t₂}
-- 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)
+_↔̇_ : ∀ {i f t} {I : Set i} → Pred I f → Pred I t → Set _
+From ↔̇ To = ∀ {i} → From i ↔ To i
+
-- If two setoids are in bijective correspondence, then there is an
-- inverse between them.
diff --git a/src/Function/Related.agda b/src/Function/Related.agda
index 6eb2430..2a08835 100644
--- a/src/Function/Related.agda
+++ b/src/Function/Related.agda
@@ -270,7 +270,7 @@ module EquationalReasoning where
sym {bijection} = Inv.sym
infix 2 _∎
- infixr 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
@@ -282,6 +282,10 @@ module EquationalReasoning where
X ↔ Y → Y ∼[ k ] Z → X ∼[ k ] Z
X ↔⟨ X↔Y ⟩ Y⇔Z = X ∼⟨ ↔⇒ X↔Y ⟩ Y⇔Z
+ _↔⟨⟩_ : ∀ {k x y} (X : Set x) {Y : Set y} →
+ X ∼[ k ] Y → X ∼[ k ] Y
+ X ↔⟨⟩ X⇔Y = X⇔Y
+
_≡⟨_⟩_ : ∀ {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
diff --git a/src/Function/Related/TypeIsomorphisms.agda b/src/Function/Related/TypeIsomorphisms.agda
index 80b9b52..c4e3b65 100644
--- a/src/Function/Related/TypeIsomorphisms.agda
+++ b/src/Function/Related/TypeIsomorphisms.agda
@@ -28,7 +28,7 @@ open import Relation.Binary.Product.Pointwise
open import Relation.Binary.PropositionalEquality as P using (_≡_; _≗_)
open import Relation.Binary.Sum
open import Relation.Nullary
-import Relation.Nullary.Decidable as Dec
+open import Relation.Nullary.Decidable as Dec using (True)
------------------------------------------------------------------------
-- Σ is "associative"
@@ -428,3 +428,90 @@ Related-cong {A = A} {B} {C} {D} A≈B C≈D =
D ∼⟨ sym C≈D ⟩
C ∎)
where open Related.EquationalReasoning
+
+------------------------------------------------------------------------
+-- A lemma relating True dec and P, where dec : Dec P
+
+True↔ : ∀ {p} {P : Set p}
+ (dec : Dec P) → ((p₁ p₂ : P) → p₁ ≡ p₂) → True dec ↔ P
+True↔ (yes p) irr = record
+ { to = P.→-to-⟶ (λ _ → p)
+ ; from = P.→-to-⟶ (λ _ → _)
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.refl
+ ; right-inverse-of = irr p
+ }
+ }
+True↔ (no ¬p) _ = record
+ { to = P.→-to-⟶ (λ ())
+ ; from = P.→-to-⟶ (λ p → ¬p p)
+ ; inverse-of = record
+ { left-inverse-of = λ ()
+ ; right-inverse-of = λ p → ⊥-elim (¬p p)
+ }
+ }
+
+------------------------------------------------------------------------
+-- Equality between pairs can be expressed as a pair of equalities
+
+Σ-≡,≡↔≡ : ∀ {a b} {A : Set a} {B : A → Set b} {p₁ p₂ : Σ A B} →
+ (∃ λ (p : proj₁ p₁ ≡ proj₁ p₂) →
+ P.subst B p (proj₂ p₁) ≡ proj₂ p₂) ↔
+ (p₁ ≡ p₂)
+Σ-≡,≡↔≡ {A = A} {B} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = left-inverse-of
+ ; right-inverse-of = right-inverse-of
+ }
+ }
+ where
+ to : {p₁ p₂ : Σ A B} →
+ Σ (proj₁ p₁ ≡ proj₁ p₂)
+ (λ p → P.subst B p (proj₂ p₁) ≡ proj₂ p₂) →
+ p₁ ≡ p₂
+ to (P.refl , P.refl) = P.refl
+
+ from : {p₁ p₂ : Σ A B} →
+ p₁ ≡ p₂ →
+ Σ (proj₁ p₁ ≡ proj₁ p₂)
+ (λ p → P.subst B p (proj₂ p₁) ≡ proj₂ p₂)
+ from P.refl = P.refl , P.refl
+
+ left-inverse-of : {p₁ p₂ : Σ A B}
+ (p : Σ (proj₁ p₁ ≡ proj₁ p₂)
+ (λ x → P.subst B x (proj₂ p₁) ≡ proj₂ p₂)) →
+ from (to p) ≡ p
+ left-inverse-of (P.refl , P.refl) = P.refl
+
+ right-inverse-of : {p₁ p₂ : Σ A B} (p : p₁ ≡ p₂) → to (from p) ≡ p
+ right-inverse-of P.refl = P.refl
+
+×-≡,≡↔≡ : ∀ {a b} {A : Set a} {B : Set b} {p₁ p₂ : A × B} →
+ (proj₁ p₁ ≡ proj₁ p₂ × proj₂ p₁ ≡ proj₂ p₂) ↔
+ p₁ ≡ p₂
+×-≡,≡↔≡ {A = A} {B} = record
+ { to = P.→-to-⟶ to
+ ; from = P.→-to-⟶ from
+ ; inverse-of = record
+ { left-inverse-of = left-inverse-of
+ ; right-inverse-of = right-inverse-of
+ }
+ }
+ where
+ to : {p₁ p₂ : A × B} →
+ (proj₁ p₁ ≡ proj₁ p₂) × (proj₂ p₁ ≡ proj₂ p₂) → p₁ ≡ p₂
+ to (P.refl , P.refl) = P.refl
+
+ from : {p₁ p₂ : A × B} → p₁ ≡ p₂ →
+ (proj₁ p₁ ≡ proj₁ p₂) × (proj₂ p₁ ≡ proj₂ p₂)
+ from P.refl = P.refl , P.refl
+
+ left-inverse-of : {p₁ p₂ : A × B} →
+ (p : (proj₁ p₁ ≡ proj₁ p₂) × (proj₂ p₁ ≡ proj₂ p₂)) →
+ from (to p) ≡ p
+ left-inverse-of (P.refl , P.refl) = P.refl
+
+ right-inverse-of : {p₁ p₂ : A × B} (p : p₁ ≡ p₂) → to (from p) ≡ p
+ right-inverse-of P.refl = P.refl
diff --git a/src/Induction.agda b/src/Induction.agda
index f145c34..0effa83 100644
--- a/src/Induction.agda
+++ b/src/Induction.agda
@@ -19,23 +19,23 @@ open import Relation.Unary
-- A RecStruct describes the allowed structure of recursion. The
-- examples in Induction.Nat should explain what this is all about.
-RecStruct : ∀ {a} → Set a → Set (suc a)
-RecStruct {a} A = Pred A a → Pred A a
+RecStruct : ∀ {a} → Set a → (ℓ₁ ℓ₂ : Level) → Set _
+RecStruct A ℓ₁ ℓ₂ = Pred A ℓ₁ → Pred A ℓ₂
-- A recursor builder constructs an instance of a recursion structure
-- for a given input.
-RecursorBuilder : ∀ {a} {A : Set a} → RecStruct A → Set _
+RecursorBuilder : ∀ {a ℓ₁ ℓ₂} {A : Set a} → RecStruct A ℓ₁ ℓ₂ → Set _
RecursorBuilder Rec = ∀ P → Rec P ⊆′ P → Universal (Rec P)
-- A recursor can be used to actually compute/prove something useful.
-Recursor : ∀ {a} {A : Set a} → RecStruct A → Set _
+Recursor : ∀ {a ℓ₁ ℓ₂} {A : Set a} → RecStruct A ℓ₁ ℓ₂ → Set _
Recursor Rec = ∀ P → Rec P ⊆′ P → Universal P
-- And recursors can be constructed from recursor builders.
-build : ∀ {a} {A : Set a} {Rec : RecStruct A} →
+build : ∀ {a ℓ₁ ℓ₂} {A : Set a} {Rec : RecStruct A ℓ₁ ℓ₂} →
RecursorBuilder Rec →
Recursor Rec
build builder P f x = f x (builder P f x)
@@ -43,15 +43,16 @@ build builder P f x = f x (builder P f x)
-- We can repeat the exercise above for subsets of the type we are
-- recursing over.
-SubsetRecursorBuilder : ∀ {a ℓ} {A : Set a} →
- Pred A ℓ → RecStruct A → Set _
+SubsetRecursorBuilder : ∀ {a ℓ₁ ℓ₂ ℓ₃} {A : Set a} →
+ Pred A ℓ₁ → RecStruct A ℓ₂ ℓ₃ → Set _
SubsetRecursorBuilder Q Rec = ∀ P → Rec P ⊆′ P → Q ⊆′ Rec P
-SubsetRecursor : ∀ {a ℓ} {A : Set a} →
- Pred A ℓ → RecStruct A → Set _
+SubsetRecursor : ∀ {a ℓ₁ ℓ₂ ℓ₃} {A : Set a} →
+ Pred A ℓ₁ → RecStruct A ℓ₂ ℓ₃ → Set _
SubsetRecursor Q Rec = ∀ P → Rec P ⊆′ P → Q ⊆′ P
-subsetBuild : ∀ {a ℓ} {A : Set a} {Q : Pred A ℓ} {Rec : RecStruct A} →
+subsetBuild : ∀ {a ℓ₁ ℓ₂ ℓ₃}
+ {A : Set a} {Q : Pred A ℓ₁} {Rec : RecStruct A ℓ₂ ℓ₃} →
SubsetRecursorBuilder Q Rec →
SubsetRecursor Q Rec
subsetBuild builder P f x q = f x (builder P f x q)
diff --git a/src/Induction/Lexicographic.agda b/src/Induction/Lexicographic.agda
index a64d79f..18c620b 100644
--- a/src/Induction/Lexicographic.agda
+++ b/src/Induction/Lexicographic.agda
@@ -6,13 +6,15 @@
module Induction.Lexicographic where
-open import Induction
open import Data.Product
+open import Induction
+open import Level
-- The structure of lexicographic induction.
-Σ-Rec : ∀ {ℓ} {A : Set ℓ} {B : A → Set ℓ} →
- RecStruct A → (∀ x → RecStruct (B x)) → RecStruct (Σ A B)
+Σ-Rec : ∀ {a b ℓ₁ ℓ₂ ℓ₃} {A : Set a} {B : A → Set b} →
+ RecStruct A (ℓ₁ ⊔ b) ℓ₂ → (∀ x → RecStruct (B x) ℓ₁ ℓ₃) →
+ RecStruct (Σ A B) _ _
Σ-Rec RecA RecB P (x , y) =
-- Either x is constant and y is "smaller", ...
RecB x (λ y' → P (x , y')) y
@@ -20,16 +22,17 @@ open import Data.Product
-- ...or x is "smaller" and y is arbitrary.
RecA (λ x' → ∀ y' → P (x' , y')) x
-_⊗_ : ∀ {ℓ} {A B : Set ℓ} →
- RecStruct A → RecStruct B → RecStruct (A × B)
+_⊗_ : ∀ {a b ℓ₁ ℓ₂ ℓ₃} {A : Set a} {B : Set b} →
+ RecStruct A (ℓ₁ ⊔ b) ℓ₂ → RecStruct B ℓ₁ ℓ₃ →
+ RecStruct (A × B) _ _
RecA ⊗ RecB = Σ-Rec RecA (λ _ → RecB)
-- Constructs a recursor builder for lexicographic induction.
Σ-rec-builder :
- ∀ {ℓ} {A : Set ℓ} {B : A → Set ℓ}
- {RecA : RecStruct A}
- {RecB : ∀ x → RecStruct (B x)} →
+ ∀ {a b ℓ₁ ℓ₂ ℓ₃} {A : Set a} {B : A → Set b}
+ {RecA : RecStruct A (ℓ₁ ⊔ b) ℓ₂}
+ {RecB : ∀ x → RecStruct (B x) ℓ₁ ℓ₃} →
RecursorBuilder RecA → (∀ x → RecursorBuilder (RecB x)) →
RecursorBuilder (Σ-Rec RecA RecB)
Σ-rec-builder {RecA = RecA} {RecB = RecB} recA recB P f (x , y) =
@@ -49,7 +52,8 @@ RecA ⊗ RecB = Σ-Rec RecA (λ _ → RecB)
p₂x = p₂ x
-[_⊗_] : ∀ {ℓ} {A B : Set ℓ} {RecA : RecStruct A} {RecB : RecStruct B} →
+[_⊗_] : ∀ {a b ℓ₁ ℓ₂ ℓ₃} {A : Set a} {B : Set b}
+ {RecA : RecStruct A (ℓ₁ ⊔ b) ℓ₂} {RecB : RecStruct B ℓ₁ ℓ₃} →
RecursorBuilder RecA → RecursorBuilder RecB →
RecursorBuilder (RecA ⊗ RecB)
[ recA ⊗ recB ] = Σ-rec-builder recA (λ _ → recB)
diff --git a/src/Induction/Nat.agda b/src/Induction/Nat.agda
index e903a1c..23a61b4 100644
--- a/src/Induction/Nat.agda
+++ b/src/Induction/Nat.agda
@@ -8,48 +8,49 @@ module Induction.Nat where
open import Function
open import Data.Nat
-open import Data.Fin
-open import Data.Fin.Props
+open import Data.Fin using (_≺_)
+open import Data.Fin.Properties
open import Data.Product
open import Data.Unit
open import Induction
open import Induction.WellFounded as WF
+open import Level using (Lift)
open import Relation.Binary.PropositionalEquality
open import Relation.Unary
------------------------------------------------------------------------
-- Ordinary induction
-Rec : RecStruct ℕ
-Rec P zero = ⊤
-Rec P (suc n) = P n
+Rec : ∀ ℓ → RecStruct ℕ ℓ ℓ
+Rec _ P zero = Lift ⊤
+Rec _ P (suc n) = P n
-rec-builder : RecursorBuilder Rec
-rec-builder P f zero = tt
+rec-builder : ∀ {ℓ} → RecursorBuilder (Rec ℓ)
+rec-builder P f zero = _
rec-builder P f (suc n) = f n (rec-builder P f n)
-rec : Recursor Rec
+rec : ∀ {ℓ} → Recursor (Rec ℓ)
rec = build rec-builder
------------------------------------------------------------------------
-- Complete induction
-CRec : RecStruct ℕ
-CRec P zero = ⊤
-CRec P (suc n) = P n × CRec P n
+CRec : ∀ ℓ → RecStruct ℕ ℓ ℓ
+CRec _ P zero = Lift ⊤
+CRec _ P (suc n) = P n × CRec _ P n
-cRec-builder : RecursorBuilder CRec
-cRec-builder P f zero = tt
+cRec-builder : ∀ {ℓ} → RecursorBuilder (CRec ℓ)
+cRec-builder P f zero = _
cRec-builder P f (suc n) = f n ih , ih
where ih = cRec-builder P f n
-cRec : Recursor CRec
+cRec : ∀ {ℓ} → Recursor (CRec ℓ)
cRec = build cRec-builder
------------------------------------------------------------------------
-- Complete induction based on _<′_
-<-Rec : RecStruct ℕ
+<-Rec : ∀ {ℓ} → RecStruct ℕ ℓ ℓ
<-Rec = WfRec _<′_
mutual
@@ -62,24 +63,26 @@ mutual
<-well-founded′ (suc n) .n ≤′-refl = <-well-founded n
<-well-founded′ (suc n) m (≤′-step m<n) = <-well-founded′ n m m<n
-open WF.All <-well-founded public
- renaming ( wfRec-builder to <-rec-builder
- ; wfRec to <-rec
- )
+module _ {ℓ} where
+ open WF.All <-well-founded ℓ public
+ renaming ( wfRec-builder to <-rec-builder
+ ; wfRec to <-rec
+ )
------------------------------------------------------------------------
-- Complete induction based on _≺_
-≺-Rec : RecStruct ℕ
+≺-Rec : ∀ {ℓ} → RecStruct ℕ ℓ ℓ
≺-Rec = WfRec _≺_
≺-well-founded : Well-founded _≺_
≺-well-founded = Subrelation.well-founded ≺⇒<′ <-well-founded
-open WF.All ≺-well-founded public
- renaming ( wfRec-builder to ≺-rec-builder
- ; wfRec to ≺-rec
- )
+module _ {ℓ} where
+ open WF.All ≺-well-founded ℓ public
+ renaming ( wfRec-builder to ≺-rec-builder
+ ; wfRec to ≺-rec
+ )
------------------------------------------------------------------------
-- Examples
@@ -88,18 +91,145 @@ private
module Examples where
- -- The half function.
+ -- Doubles its input.
- half₁ : ℕ → ℕ
- half₁ = cRec _ λ
- { zero _ → zero
- ; (suc zero) _ → zero
- ; (suc (suc n)) (_ , half₁n , _) → suc half₁n
+ twice : ℕ → ℕ
+ twice = rec _ λ
+ { zero _ → zero
+ ; (suc n) twice-n → suc (suc twice-n)
}
- half₂ : ℕ → ℕ
- half₂ = <-rec _ λ
- { zero _ → zero
- ; (suc zero) _ → zero
- ; (suc (suc n)) rec → suc (rec n (≤′-step ≤′-refl))
+ -- Halves its input (rounding downwards).
+ --
+ -- The step function is mentioned in a proof below, so it has been
+ -- given a name. (The mutual keyword is used to avoid having to give
+ -- a type signature for the step function.)
+
+ mutual
+
+ half₁-step = λ
+ { zero _ → zero
+ ; (suc zero) _ → zero
+ ; (suc (suc n)) (_ , half₁n , _) → suc half₁n
+ }
+
+ half₁ : ℕ → ℕ
+ half₁ = cRec _ half₁-step
+
+ -- An alternative implementation of half₁.
+
+ mutual
+
+ half₂-step = λ
+ { zero _ → zero
+ ; (suc zero) _ → zero
+ ; (suc (suc n)) rec → suc (rec n (≤′-step ≤′-refl))
+ }
+
+ half₂ : ℕ → ℕ
+ half₂ = <-rec _ half₂-step
+
+ -- The application half₁ (2 + n) is definitionally equal to
+ -- 1 + half₁ n. Perhaps it is instructive to see why.
+
+ half₁-2+ : ∀ n → half₁ (2 + n) ≡ 1 + half₁ n
+ half₁-2+ n = begin
+
+ half₁ (2 + n) ≡⟨⟩
+
+ cRec (λ _ → ℕ) half₁-step (2 + n) ≡⟨⟩
+
+ half₁-step (2 + n) (cRec-builder (λ _ → ℕ) half₁-step (2 + n)) ≡⟨⟩
+
+ half₁-step (2 + n)
+ (let ih = cRec-builder (λ _ → ℕ) half₁-step (1 + n) in
+ half₁-step (1 + n) ih , ih) ≡⟨⟩
+
+ half₁-step (2 + n)
+ (let ih = cRec-builder (λ _ → ℕ) half₁-step n in
+ half₁-step (1 + n) (half₁-step n ih , ih) , half₁-step n ih , ih) ≡⟨⟩
+
+ 1 + half₁-step n (cRec-builder (λ _ → ℕ) half₁-step n) ≡⟨⟩
+
+ 1 + cRec (λ _ → ℕ) half₁-step n ≡⟨⟩
+
+ 1 + half₁ n ∎
+
+ where open ≡-Reasoning
+
+ -- The application half₂ (2 + n) is definitionally equal to
+ -- 1 + half₂ n. Perhaps it is instructive to see why.
+
+ half₂-2+ : ∀ n → half₂ (2 + n) ≡ 1 + half₂ n
+ half₂-2+ n = begin
+
+ half₂ (2 + n) ≡⟨⟩
+
+ <-rec (λ _ → ℕ) half₂-step (2 + n) ≡⟨⟩
+
+ half₂-step (2 + n) (<-rec-builder (λ _ → ℕ) half₂-step (2 + n)) ≡⟨⟩
+
+ 1 + <-rec-builder (λ _ → ℕ) half₂-step (2 + n) n (≤′-step ≤′-refl) ≡⟨⟩
+
+ 1 + Some.wfRec-builder (λ _ → ℕ) half₂-step (2 + n)
+ (<-well-founded (2 + n)) n (≤′-step ≤′-refl) ≡⟨⟩
+
+ 1 + Some.wfRec-builder (λ _ → ℕ) half₂-step (2 + n)
+ (acc (<-well-founded′ (2 + n))) n (≤′-step ≤′-refl) ≡⟨⟩
+
+ 1 + half₂-step n
+ (Some.wfRec-builder (λ _ → ℕ) half₂-step n
+ (<-well-founded′ (2 + n) n (≤′-step ≤′-refl))) ≡⟨⟩
+
+ 1 + half₂-step n
+ (Some.wfRec-builder (λ _ → ℕ) half₂-step n
+ (<-well-founded′ (1 + n) n ≤′-refl)) ≡⟨⟩
+
+ 1 + half₂-step n
+ (Some.wfRec-builder (λ _ → ℕ) half₂-step n (<-well-founded n)) ≡⟨⟩
+
+ 1 + half₂-step n (<-rec-builder (λ _ → ℕ) half₂-step n) ≡⟨⟩
+
+ 1 + <-rec (λ _ → ℕ) half₂-step n ≡⟨⟩
+
+ 1 + half₂ n ∎
+
+ where open ≡-Reasoning
+
+ -- Some properties that the functions above satisfy, proved using
+ -- cRec.
+
+ half₁-+₁ : ∀ n → half₁ (twice n) ≡ n
+ half₁-+₁ = cRec _ λ
+ { zero _ → refl
+ ; (suc zero) _ → refl
+ ; (suc (suc n)) (_ , half₁twice-n≡n , _) →
+ cong (suc ∘ suc) half₁twice-n≡n
+ }
+
+ half₂-+₁ : ∀ n → half₂ (twice n) ≡ n
+ half₂-+₁ = cRec _ λ
+ { zero _ → refl
+ ; (suc zero) _ → refl
+ ; (suc (suc n)) (_ , half₁twice-n≡n , _) →
+ cong (suc ∘ suc) half₁twice-n≡n
+ }
+
+ -- Some properties that the functions above satisfy, proved using
+ -- <-rec.
+
+ half₁-+₂ : ∀ n → half₁ (twice n) ≡ n
+ half₁-+₂ = <-rec _ λ
+ { zero _ → refl
+ ; (suc zero) _ → refl
+ ; (suc (suc n)) rec →
+ cong (suc ∘ suc) (rec n (≤′-step ≤′-refl))
+ }
+
+ half₂-+₂ : ∀ n → half₂ (twice n) ≡ n
+ half₂-+₂ = <-rec _ λ
+ { zero _ → refl
+ ; (suc zero) _ → refl
+ ; (suc (suc n)) rec →
+ cong (suc ∘ suc) (rec n (≤′-step ≤′-refl))
}
diff --git a/src/Induction/WellFounded.agda b/src/Induction/WellFounded.agda
index c455aba..c771b6d 100644
--- a/src/Induction/WellFounded.agda
+++ b/src/Induction/WellFounded.agda
@@ -11,33 +11,34 @@ module Induction.WellFounded where
open import Data.Product
open import Function
open import Induction
+open import Level
open import Relation.Unary
-- When using well-founded recursion you can recurse arbitrarily, as
-- long as the arguments become smaller, and "smaller" is
-- well-founded.
-WfRec : ∀ {a} {A : Set a} → Rel A a → RecStruct A
+WfRec : ∀ {a r} {A : Set a} → Rel A r → ∀ {ℓ} → RecStruct A ℓ _
WfRec _<_ P x = ∀ y → y < x → P y
-- The accessibility predicate: x is accessible if everything which is
-- smaller than x is also accessible (inductively).
-data Acc {a} {A : Set a} (_<_ : Rel A a) (x : A) : Set a where
+data Acc {a ℓ} {A : Set a} (_<_ : Rel A ℓ) (x : A) : Set (a ⊔ ℓ) where
acc : (rs : WfRec _<_ (Acc _<_) x) → Acc _<_ x
-- The accessibility predicate encodes what it means to be
-- well-founded; if all elements are accessible, then _<_ is
-- well-founded.
-Well-founded : ∀ {a} {A : Set a} → Rel A a → Set a
+Well-founded : ∀ {a ℓ} {A : Set a} → Rel A ℓ → Set _
Well-founded _<_ = ∀ x → Acc _<_ x
-- Well-founded induction for the subset of accessible elements:
-module Some {a} {A : Set a} {_<_ : Rel A a} where
+module Some {a lt} {A : Set a} {_<_ : Rel A lt} {ℓ} where
- wfRec-builder : SubsetRecursorBuilder (Acc _<_) (WfRec _<_)
+ wfRec-builder : SubsetRecursorBuilder (Acc _<_) (WfRec _<_ {ℓ = ℓ})
wfRec-builder P f x (acc rs) = λ y y<x →
f y (wfRec-builder P f y (rs y y<x))
@@ -47,10 +48,10 @@ module Some {a} {A : Set a} {_<_ : Rel A a} where
-- Well-founded induction for all elements, assuming they are all
-- accessible:
-module All {a} {A : Set a} {_<_ : Rel A a}
- (wf : Well-founded _<_) where
+module All {a lt} {A : Set a} {_<_ : Rel A lt}
+ (wf : Well-founded _<_) ℓ where
- wfRec-builder : RecursorBuilder (WfRec _<_)
+ wfRec-builder : RecursorBuilder (WfRec _<_ {ℓ = ℓ})
wfRec-builder P f x = Some.wfRec-builder P f x (wf x)
wfRec : Recursor (WfRec _<_)
@@ -61,7 +62,8 @@ module All {a} {A : Set a} {_<_ : Rel A a}
-- "Constructing Recursion Operators in Intuitionistic Type Theory" by
-- Lawrence C Paulson).
-module Subrelation {a} {A : Set a} {_<₁_ _<₂_ : Rel A a}
+module Subrelation {a ℓ₁ ℓ₂} {A : Set a}
+ {_<₁_ : Rel A ℓ₁} {_<₂_ : Rel A ℓ₂}
(<₁⇒<₂ : ∀ {x y} → x <₁ y → x <₂ y) where
accessible : Acc _<₂_ ⊆ Acc _<₁_
@@ -70,7 +72,7 @@ module Subrelation {a} {A : Set a} {_<₁_ _<₂_ : Rel A a}
well-founded : Well-founded _<₂_ → Well-founded _<₁_
well-founded wf = λ x → accessible (wf x)
-module Inverse-image {ℓ} {A B : Set ℓ} {_<_ : Rel B ℓ}
+module Inverse-image {a b ℓ} {A : Set a} {B : Set b} {_<_ : Rel B ℓ}
(f : A → B) where
accessible : ∀ {x} → Acc _<_ (f x) → Acc (_<_ on f) x
@@ -79,11 +81,11 @@ module Inverse-image {ℓ} {A B : Set ℓ} {_<_ : Rel B ℓ}
well-founded : Well-founded _<_ → Well-founded (_<_ on f)
well-founded wf = λ x → accessible (wf (f x))
-module Transitive-closure {a} {A : Set a} (_<_ : Rel A a) where
+module Transitive-closure {a ℓ} {A : Set a} (_<_ : Rel A ℓ) where
infix 4 _<⁺_
- data _<⁺_ : Rel A a where
+ data _<⁺_ : Rel A (a ⊔ ℓ) where
[_] : ∀ {x y} (x<y : x < y) → x <⁺ y
trans : ∀ {x y z} (x<y : x <⁺ y) (y<z : y <⁺ z) → x <⁺ z
@@ -103,11 +105,11 @@ module Transitive-closure {a} {A : Set a} (_<_ : Rel A a) where
well-founded : Well-founded _<_ → Well-founded _<⁺_
well-founded wf = λ x → accessible (wf x)
-module Lexicographic {ℓ} {A : Set ℓ} {B : A → Set ℓ}
- (RelA : Rel A ℓ)
- (RelB : ∀ x → Rel (B x) ℓ) where
+module Lexicographic {a b ℓ₁ ℓ₂} {A : Set a} {B : A → Set b}
+ (RelA : Rel A ℓ₁)
+ (RelB : ∀ x → Rel (B x) ℓ₂) where
- data _<_ : Rel (Σ A B) ℓ where
+ data _<_ : Rel (Σ A B) (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) where
left : ∀ {x₁ y₁ x₂ y₂} (x₁<x₂ : RelA x₁ x₂) → (x₁ , y₁) < (x₂ , y₂)
right : ∀ {x y₁ y₂} (y₁<y₂ : RelB x y₁ y₂) → (x , y₁) < (x , y₂)
diff --git a/src/Level.agda b/src/Level.agda
index 4c2bd5e..648bf49 100644
--- a/src/Level.agda
+++ b/src/Level.agda
@@ -8,26 +8,9 @@ module Level where
-- Levels.
-infixl 6 _⊔_
-
-postulate
- Level : Set
- zero : Level
- suc : Level → Level
- _⊔_ : Level → Level → Level
-
--- MAlonzo compiles Level to (). This should be safe, because it is
--- not possible to pattern match on levels.
-
-{-# COMPILED_TYPE Level () #-}
-{-# COMPILED zero () #-}
-{-# COMPILED suc (\_ -> ()) #-}
-{-# COMPILED _⊔_ (\_ _ -> ()) #-}
-
-{-# BUILTIN LEVEL Level #-}
-{-# BUILTIN LEVELZERO zero #-}
-{-# BUILTIN LEVELSUC suc #-}
-{-# BUILTIN LEVELMAX _⊔_ #-}
+open import Agda.Primitive public
+ using (Level; _⊔_)
+ renaming (lzero to zero; lsuc to suc)
-- Lifting.
diff --git a/src/Record.agda b/src/Record.agda
index 8444fb1..5e8ad4f 100644
--- a/src/Record.agda
+++ b/src/Record.agda
@@ -194,7 +194,7 @@ _·_ : ∀ {s} {Sig : Signature s} (r : Record Sig)
_·_ {Sig = ∅} r ℓ {}
_·_ {Sig = Sig , ℓ′ ∶ A} (rec r) ℓ {ℓ∈} with ℓ ≟ ℓ′
... | yes _ = Σ.proj₂ r
-... | no _ = _·_ (Σ.proj₁ r) ℓ {ℓ∈}
+... | no _ = _·_ (Σ.proj₁ r) ℓ {ℓ∈}
_·_ {Sig = Sig , ℓ′ ≔ a} (rec r) ℓ {ℓ∈} with ℓ ≟ ℓ′
... | yes _ = Manifest-Σ.proj₂ r
... | no _ = _·_ (Manifest-Σ.proj₁ r) ℓ {ℓ∈}
@@ -224,7 +224,7 @@ mutual
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)
+ ... | 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 62d7f54..f06c210 100644
--- a/src/Reflection.agda
+++ b/src/Reflection.agda
@@ -71,11 +71,16 @@ data Relevance : Set where
-- Arguments.
-data Arg A : Set where
- arg : (v : Visibility) (r : Relevance) (x : A) → Arg A
+data Arg-info : Set where
+ arg-info : (v : Visibility) (r : Relevance) → Arg-info
-{-# BUILTIN ARG Arg #-}
-{-# BUILTIN ARGARG arg #-}
+data Arg (A : Set) : Set where
+ arg : (i : Arg-info) (x : A) → Arg A
+
+{-# BUILTIN ARGINFO Arg-info #-}
+{-# BUILTIN ARGARGINFO arg-info #-}
+{-# BUILTIN ARG Arg #-}
+{-# BUILTIN ARGARG arg #-}
-- Terms.
@@ -191,14 +196,17 @@ private
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₁ : ∀ {A i i′} {x x′ : A} → arg i x ≡ arg i′ x′ → i ≡ i′
arg₁ refl = refl
- arg₂ : ∀ {A v v′ r r′} {x x′ : A} → arg v r x ≡ arg v′ r′ x′ → r ≡ r′
+ arg₂ : ∀ {A i i′} {x x′ : A} → arg i x ≡ arg i′ x′ → x ≡ x′
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
+ arg-info₁ : ∀ {v v′ r r′} → arg-info v r ≡ arg-info v′ r′ → v ≡ v′
+ arg-info₁ refl = refl
+
+ arg-info₂ : ∀ {v v′ r r′} → arg-info v r ≡ arg-info v′ r′ → r ≡ r′
+ arg-info₂ refl = refl
cons₁ : ∀ {A : Set} {x y} {xs ys : List A} → x ∷ xs ≡ y ∷ ys → x ≡ y
cons₁ refl = refl
@@ -268,22 +276,26 @@ irrelevant ≟-Relevance irrelevant = yes refl
relevant ≟-Relevance irrelevant = no λ()
irrelevant ≟-Relevance relevant = no λ()
+_≟-Arg-info_ : Decidable (_≡_ {A = Arg-info})
+arg-info v r ≟-Arg-info arg-info v′ r′ =
+ Dec.map′ (cong₂′ arg-info)
+ < arg-info₁ , arg-info₂ >
+ (v ≟-Visibility v′ ×-dec r ≟-Relevance r′)
+
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′)
+ arg i a ≟-ArgTerm arg i′ a′ =
+ Dec.map′ (cong₂′ arg)
+ < arg₁ , arg₂ >
+ (i ≟-Arg-info i′ ×-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′)
+ arg i a ≟-ArgType arg i′ a′ =
+ Dec.map′ (cong₂′ arg)
+ < arg₁ , arg₂ >
+ (i ≟-Arg-info i′ ×-dec a ≟-Type a′)
_≟-Args_ : Decidable (_≡_ {A = List (Arg Term)})
[] ≟-Args [] = yes refl
diff --git a/src/Relation/Binary.agda b/src/Relation/Binary.agda
index c49ea00..447b284 100644
--- a/src/Relation/Binary.agda
+++ b/src/Relation/Binary.agda
@@ -215,6 +215,58 @@ record StrictPartialOrder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂))
open IsStrictPartialOrder isStrictPartialOrder public
+ asymmetric : Asymmetric _<_
+ asymmetric {x} {y} =
+ trans∧irr⟶asym Eq.refl trans irrefl {x = x} {y = y}
+
+------------------------------------------------------------------------
+-- Decidable strict partial orders
+
+record IsDecStrictPartialOrder {a ℓ₁ ℓ₂} {A : Set a}
+ (_≈_ : Rel A ℓ₁) (_<_ : Rel A ℓ₂) :
+ Set (a ⊔ ℓ₁ ⊔ ℓ₂) where
+ infix 4 _≟_ _<?_
+ field
+ isStrictPartialOrder : IsStrictPartialOrder _≈_ _<_
+ _≟_ : Decidable _≈_
+ _<?_ : Decidable _<_
+
+ private
+ module SPO = IsStrictPartialOrder isStrictPartialOrder
+ open SPO public hiding (module Eq)
+
+ module Eq where
+
+ isDecEquivalence : IsDecEquivalence _≈_
+ isDecEquivalence = record
+ { isEquivalence = SPO.isEquivalence
+ ; _≟_ = _≟_
+ }
+
+ open IsDecEquivalence isDecEquivalence public
+
+record DecStrictPartialOrder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where
+ infix 4 _≈_ _<_
+ field
+ Carrier : Set c
+ _≈_ : Rel Carrier ℓ₁
+ _<_ : Rel Carrier ℓ₂
+ isDecStrictPartialOrder : IsDecStrictPartialOrder _≈_ _<_
+
+ private
+ module DSPO = IsDecStrictPartialOrder isDecStrictPartialOrder
+ open DSPO public hiding (module Eq)
+
+ strictPartialOrder : StrictPartialOrder c ℓ₁ ℓ₂
+ strictPartialOrder = record { isStrictPartialOrder = isStrictPartialOrder }
+
+ module Eq where
+
+ decSetoid : DecSetoid c ℓ₁
+ decSetoid = record { isDecEquivalence = DSPO.Eq.isDecEquivalence }
+
+ open DecSetoid decSetoid public
+
------------------------------------------------------------------------
-- Total orders
diff --git a/src/Relation/Binary/HeterogeneousEquality.agda b/src/Relation/Binary/HeterogeneousEquality.agda
index 83f588e..d6e6896 100644
--- a/src/Relation/Binary/HeterogeneousEquality.agda
+++ b/src/Relation/Binary/HeterogeneousEquality.agda
@@ -9,6 +9,7 @@ module Relation.Binary.HeterogeneousEquality where
open import Data.Product
open import Function
open import Function.Inverse using (Inverse)
+open import Data.Unit.Core
open import Level
open import Relation.Nullary
open import Relation.Binary
@@ -27,7 +28,7 @@ open Core public using (_≅_; refl)
-- Nonequality.
-_≇_ : ∀ {a} {A : Set a} → A → ∀ {b} {B : Set b} → B → Set
+_≇_ : ∀ {ℓ} {A : Set ℓ} → A → {B : Set ℓ} → B → Set ℓ
x ≇ y = ¬ x ≅ y
------------------------------------------------------------------------
@@ -44,11 +45,10 @@ open Core public using (≅-to-≡)
reflexive : ∀ {a} {A : Set a} → _⇒_ {A = A} _≡_ (λ x y → x ≅ y)
reflexive refl = refl
-sym : ∀ {a b} {A : Set a} {B : Set b} {x : A} {y : B} → x ≅ y → y ≅ x
+sym : ∀ {ℓ} {A B : Set ℓ} {x : A} {y : B} → x ≅ y → y ≅ x
sym refl = refl
-trans : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
- {x : A} {y : B} {z : C} →
+trans : ∀ {ℓ} {A B C : Set ℓ} {x : A} {y : B} {z : C} →
x ≅ y → y ≅ z → x ≅ z
trans refl eq = eq
@@ -73,6 +73,10 @@ cong : ∀ {a b} {A : Set a} {B : A → Set b} {x y}
(f : (x : A) → B x) → x ≅ y → f x ≅ f y
cong f refl = refl
+cong-app : ∀ {a b} {A : Set a} {B : A → Set b} {f g : (x : A) → B x} →
+ f ≅ g → (x : A) → f x ≅ g x
+cong-app refl x = refl
+
cong₂ : ∀ {a b c} {A : Set a} {B : A → Set b} {C : ∀ x → B x → Set c}
{x y u v}
(f : (x : A) (y : B x) → C x y) → x ≅ y → u ≅ v → f x u ≅ f y v
@@ -81,7 +85,7 @@ cong₂ f refl refl = refl
resp₂ : ∀ {a ℓ} {A : Set a} (∼ : Rel A ℓ) → ∼ Respects₂ (λ x y → x ≅ y)
resp₂ _∼_ = subst⟶resp₂ _∼_ subst
-proof-irrelevance : ∀ {a b} {A : Set a} {B : Set b} {x : A} {y : B}
+proof-irrelevance : ∀ {ℓ} {A B : Set ℓ} {x : A} {y : B}
(p q : x ≅ y) → p ≡ q
proof-irrelevance refl refl = refl
@@ -170,24 +174,23 @@ module ≅-Reasoning where
infixr 2 _≅⟨_⟩_ _≡⟨_⟩_ _≡⟨⟩_
infix 1 begin_
- data _IsRelatedTo_ {a} {A : Set a} (x : A) {b} {B : Set b} (y : B) :
- Set where
+ data _IsRelatedTo_ {ℓ} {A : Set ℓ} (x : A) {B : Set ℓ} (y : B) :
+ Set ℓ where
relTo : (x≅y : x ≅ y) → x IsRelatedTo y
- begin_ : ∀ {a} {A : Set a} {x : A} {b} {B : Set b} {y : B} →
+ begin_ : ∀ {ℓ} {A : Set ℓ} {x : A} {B} {y : B} →
x IsRelatedTo y → x ≅ y
begin relTo x≅y = x≅y
- _≅⟨_⟩_ : ∀ {a} {A : Set a} (x : A) {b} {B : Set b} {y : B}
- {c} {C : Set c} {z : C} →
+ _≅⟨_⟩_ : ∀ {ℓ} {A : Set ℓ} (x : A) {B} {y : B} {C} {z : C} →
x ≅ y → y IsRelatedTo z → x IsRelatedTo z
_ ≅⟨ x≅y ⟩ relTo y≅z = relTo (trans x≅y y≅z)
- _≡⟨_⟩_ : ∀ {a} {A : Set a} (x : A) {y c} {C : Set c} {z : C} →
+ _≡⟨_⟩_ : ∀ {ℓ} {A : Set ℓ} (x : A) {y C} {z : C} →
x ≡ y → y IsRelatedTo z → x IsRelatedTo z
_ ≡⟨ x≡y ⟩ relTo y≅z = relTo (trans (reflexive x≡y) y≅z)
- _≡⟨⟩_ : ∀ {a} {A : Set a} (x : A) {b} {B : Set b} {y : B} →
+ _≡⟨⟩_ : ∀ {ℓ} {A : Set ℓ} (x : A) {B} {y : B} →
x IsRelatedTo y → x IsRelatedTo y
_ ≡⟨⟩ x≅y = x≅y
@@ -215,3 +218,18 @@ Extensionality a b =
where
ext′ : P.Extensionality ℓ₁ ℓ₂
ext′ = P.extensionality-for-lower-levels ℓ₁ (suc ℓ₂) ext
+
+
+------------------------------------------------------------------------
+-- Inspect on steroids
+
+-- 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 Reveal_is_ {a} {A : Set a} (x : Hidden A) (y : A) : Set a where
+ [_] : (eq : reveal x ≅ y) → Reveal x is y
+
+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 ]
diff --git a/src/Relation/Binary/HeterogeneousEquality/Core.agda b/src/Relation/Binary/HeterogeneousEquality/Core.agda
index 6a2f24b..15c104e 100644
--- a/src/Relation/Binary/HeterogeneousEquality/Core.agda
+++ b/src/Relation/Binary/HeterogeneousEquality/Core.agda
@@ -16,8 +16,8 @@ open import Relation.Binary.Core using (_≡_; refl)
infix 4 _≅_
-data _≅_ {a} {A : Set a} (x : A) : ∀ {b} {B : Set b} → B → Set where
- refl : x ≅ x
+data _≅_ {ℓ} {A : Set ℓ} (x : A) : {B : Set ℓ} → B → Set ℓ where
+ refl : x ≅ x
------------------------------------------------------------------------
-- Conversion
diff --git a/src/Relation/Binary/List/Pointwise.agda b/src/Relation/Binary/List/Pointwise.agda
index 9e8cc84..d6f0fc1 100644
--- a/src/Relation/Binary/List/Pointwise.agda
+++ b/src/Relation/Binary/List/Pointwise.agda
@@ -7,12 +7,14 @@
module Relation.Binary.List.Pointwise where
open import Function
+open import Function.Inverse using (Inverse)
open import Data.Product hiding (map)
open import Data.List hiding (map)
open import Level
open import Relation.Nullary
+import Relation.Nullary.Decidable as Dec using (map′)
open import Relation.Binary renaming (Rel to Rel₂)
-open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
infixr 5 _∷_
@@ -127,9 +129,9 @@ isDecEquivalence : ∀ {a ℓ} {A : Set a}
{_≈_ : Rel₂ A ℓ} → IsDecEquivalence _≈_ →
IsDecEquivalence (Rel _≈_)
isDecEquivalence eq = record
- { isEquivalence = isEquivalence Dec.isEquivalence
- ; _≟_ = decidable Dec._≟_
- } where module Dec = IsDecEquivalence eq
+ { isEquivalence = isEquivalence DE.isEquivalence
+ ; _≟_ = decidable DE._≟_
+ } where module DE = IsDecEquivalence eq
isPartialOrder : ∀ {a ℓ₁ ℓ₂} {A : Set a}
{_≈_ : Rel₂ A ℓ₁} {_≤_ : Rel₂ A ℓ₂} →
@@ -140,16 +142,6 @@ isPartialOrder po = record
; antisym = antisymmetric PO.antisym
} where module PO = IsPartialOrder po
--- Rel _≡_ coincides with _≡_.
-
-Rel≡⇒≡ : ∀ {a} {A : Set a} → Rel {A = A} _≡_ ⇒ _≡_
-Rel≡⇒≡ [] = PropEq.refl
-Rel≡⇒≡ (PropEq.refl ∷ xs∼ys) with Rel≡⇒≡ xs∼ys
-Rel≡⇒≡ (PropEq.refl ∷ xs∼ys) | PropEq.refl = PropEq.refl
-
-≡⇒Rel≡ : ∀ {a} {A : Set a} → _≡_ ⇒ Rel {A = A} _≡_
-≡⇒Rel≡ PropEq.refl = refl PropEq.refl
-
preorder : ∀ {p₁ p₂ p₃} → Preorder p₁ p₂ p₃ → Preorder _ _ _
preorder p = record
{ isPreorder = isPreorder (Preorder.isPreorder p)
@@ -169,3 +161,30 @@ poset : ∀ {c ℓ₁ ℓ₂} → Poset c ℓ₁ ℓ₂ → Poset _ _ _
poset p = record
{ isPartialOrder = isPartialOrder (Poset.isPartialOrder p)
}
+
+-- Rel _≡_ coincides with _≡_.
+
+Rel≡⇒≡ : ∀ {a} {A : Set a} → Rel {A = A} _≡_ ⇒ _≡_
+Rel≡⇒≡ [] = P.refl
+Rel≡⇒≡ (P.refl ∷ xs∼ys) with Rel≡⇒≡ xs∼ys
+Rel≡⇒≡ (P.refl ∷ xs∼ys) | P.refl = P.refl
+
+≡⇒Rel≡ : ∀ {a} {A : Set a} → _≡_ ⇒ Rel {A = A} _≡_
+≡⇒Rel≡ P.refl = refl P.refl
+
+Rel↔≡ : ∀ {a} {A : Set a} →
+ Inverse (setoid (P.setoid A)) (P.setoid (List A))
+Rel↔≡ = record
+ { to = record { _⟨$⟩_ = id; cong = Rel≡⇒≡ }
+ ; from = record { _⟨$⟩_ = id; cong = ≡⇒Rel≡ }
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl P.refl
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+
+decidable-≡ : ∀ {a} {A : Set a} →
+ Decidable {A = A} _≡_ → Decidable {A = List A} _≡_
+decidable-≡ dec xs ys = Dec.map′ Rel≡⇒≡ ≡⇒Rel≡ (xs ≟ ys)
+ where
+ open DecSetoid (decSetoid (P.decSetoid dec))
diff --git a/src/Relation/Binary/Product/Pointwise.agda b/src/Relation/Binary/Product/Pointwise.agda
index 4d78604..8b238b1 100644
--- a/src/Relation/Binary/Product/Pointwise.agda
+++ b/src/Relation/Binary/Product/Pointwise.agda
@@ -23,9 +23,10 @@ open import Function.Related
open import Function.Surjection as Surj
using (Surjection; _↠_; module Surjection)
open import Level
+import Relation.Nullary.Decidable as Dec
open import Relation.Nullary.Product
open import Relation.Binary
-import Relation.Binary.PropositionalEquality as P
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
module _ {a₁ a₂ ℓ₁ ℓ₂} {A₁ : Set a₁} {A₂ : Set a₂} where
@@ -250,6 +251,16 @@ s₁ ×-strictPartialOrder s₂ = record
------------------------------------------------------------------------
-- Some properties related to "relatedness"
+private
+
+ to-cong : ∀ {a b} {A : Set a} {B : Set b} →
+ (_≡_ ×-Rel _≡_) ⇒ _≡_ {A = A × B}
+ to-cong {i = (x , y)} {j = (.x , .y)} (P.refl , P.refl) = P.refl
+
+ from-cong : ∀ {a b} {A : Set a} {B : Set b} →
+ _≡_ {A = A × B} ⇒ (_≡_ ×-Rel _≡_)
+ from-cong P.refl = (P.refl , P.refl)
+
×-Rel↔≡ : ∀ {a b} {A : Set a} {B : Set b} →
Inverse (P.setoid A ×-setoid P.setoid B) (P.setoid (A × B))
×-Rel↔≡ = record
@@ -260,12 +271,13 @@ s₁ ×-strictPartialOrder s₂ = record
; right-inverse-of = λ _ → P.refl
}
}
- where
- to-cong : (P._≡_ ×-Rel P._≡_) ⇒ P._≡_
- 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)
+_×-≟_ : ∀ {a b} {A : Set a} {B : Set b} →
+ Decidable {A = A} _≡_ → Decidable {A = B} _≡_ →
+ Decidable {A = A × B} _≡_
+(dec₁ ×-≟ dec₂) p₁ p₂ = Dec.map′ to-cong from-cong (p₁ ≟ p₂)
+ where
+ open DecSetoid (P.decSetoid dec₁ ×-decSetoid P.decSetoid dec₂)
_×-⟶_ :
∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈}
diff --git a/src/Relation/Binary/Properties/DecTotalOrder.agda b/src/Relation/Binary/Properties/DecTotalOrder.agda
new file mode 100644
index 0000000..13673c0
--- /dev/null
+++ b/src/Relation/Binary/Properties/DecTotalOrder.agda
@@ -0,0 +1,26 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties satisfied by decidable total orders
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Relation.Binary.Properties.DecTotalOrder
+ {d₁ d₂ d₃} (DT : DecTotalOrder d₁ d₂ d₃) where
+
+open Relation.Binary.DecTotalOrder DT hiding (trans)
+import Relation.Binary.NonStrictToStrict as Conv
+open Conv _≈_ _≤_
+
+strictTotalOrder : StrictTotalOrder _ _ _
+strictTotalOrder = record
+ { isStrictTotalOrder = record
+ { isEquivalence = isEquivalence
+ ; trans = trans isPartialOrder
+ ; compare = trichotomous Eq.sym _≟_ antisym total
+ ; <-resp-≈ = <-resp-≈ isEquivalence ≤-resp-≈
+ }
+ }
+
+open StrictTotalOrder strictTotalOrder public
diff --git a/src/Relation/Binary/Properties/Poset.agda b/src/Relation/Binary/Properties/Poset.agda
new file mode 100644
index 0000000..1171bc5
--- /dev/null
+++ b/src/Relation/Binary/Properties/Poset.agda
@@ -0,0 +1,29 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties satisfied by posets
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Relation.Binary.Properties.Poset
+ {p₁ p₂ p₃} (P : Poset p₁ p₂ p₃) where
+
+open Relation.Binary.Poset P hiding (trans)
+import Relation.Binary.NonStrictToStrict as Conv
+open Conv _≈_ _≤_
+
+------------------------------------------------------------------------
+-- Posets can be turned into strict partial orders
+
+strictPartialOrder : StrictPartialOrder _ _ _
+strictPartialOrder = record
+ { isStrictPartialOrder = record
+ { isEquivalence = isEquivalence
+ ; irrefl = irrefl
+ ; trans = trans isPartialOrder
+ ; <-resp-≈ = <-resp-≈ isEquivalence ≤-resp-≈
+ }
+ }
+
+open StrictPartialOrder strictPartialOrder
diff --git a/src/Relation/Binary/Properties/Preorder.agda b/src/Relation/Binary/Properties/Preorder.agda
new file mode 100644
index 0000000..716300e
--- /dev/null
+++ b/src/Relation/Binary/Properties/Preorder.agda
@@ -0,0 +1,28 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties satisfied by preorders
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Relation.Binary.Properties.Preorder
+ {p₁ p₂ p₃} (P : Preorder p₁ p₂ p₃) where
+
+open import Function
+open import Data.Product as Prod
+
+open Relation.Binary.Preorder P
+
+------------------------------------------------------------------------
+-- For every preorder there is an induced equivalence
+
+InducedEquivalence : Setoid _ _
+InducedEquivalence = record
+ { _≈_ = λ x y → x ∼ y × y ∼ x
+ ; isEquivalence = record
+ { refl = (refl , refl)
+ ; sym = swap
+ ; trans = Prod.zip trans (flip trans)
+ }
+ }
diff --git a/src/Relation/Binary/Properties/StrictPartialOrder.agda b/src/Relation/Binary/Properties/StrictPartialOrder.agda
new file mode 100644
index 0000000..d62d245
--- /dev/null
+++ b/src/Relation/Binary/Properties/StrictPartialOrder.agda
@@ -0,0 +1,32 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties satisfied by strict partial orders
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Relation.Binary.Properties.StrictPartialOrder
+ {s₁ s₂ s₃} (SPO : StrictPartialOrder s₁ s₂ s₃) where
+
+open Relation.Binary.StrictPartialOrder SPO
+ renaming (trans to <-trans)
+import Relation.Binary.StrictToNonStrict as Conv
+open Conv _≈_ _<_
+
+------------------------------------------------------------------------
+-- Strict partial orders can be converted to posets
+
+poset : Poset _ _ _
+poset = record
+ { isPartialOrder = record
+ { isPreorder = record
+ { isEquivalence = isEquivalence
+ ; reflexive = reflexive
+ ; trans = trans isEquivalence <-resp-≈ <-trans
+ }
+ ; antisym = antisym isEquivalence <-trans irrefl
+ }
+ }
+
+open Poset poset public
diff --git a/src/Relation/Binary/Properties/StrictTotalOrder.agda b/src/Relation/Binary/Properties/StrictTotalOrder.agda
new file mode 100644
index 0000000..7d7a836
--- /dev/null
+++ b/src/Relation/Binary/Properties/StrictTotalOrder.agda
@@ -0,0 +1,34 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties satisfied by strict partial orders
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Relation.Binary.Properties.StrictTotalOrder
+ {s₁ s₂ s₃} (STO : StrictTotalOrder s₁ s₂ s₃)
+ where
+
+open Relation.Binary.StrictTotalOrder STO
+import Relation.Binary.StrictToNonStrict as Conv
+open Conv _≈_ _<_
+import Relation.Binary.Properties.StrictPartialOrder as SPO
+open import Relation.Binary.Consequences
+
+------------------------------------------------------------------------
+-- Strict total orders can be converted to decidable total orders
+
+decTotalOrder : DecTotalOrder _ _ _
+decTotalOrder = record
+ { isDecTotalOrder = record
+ { isTotalOrder = record
+ { isPartialOrder = SPO.isPartialOrder strictPartialOrder
+ ; total = total compare
+ }
+ ; _≟_ = _≟_
+ ; _≤?_ = decidable' compare
+ }
+ }
+
+open DecTotalOrder decTotalOrder public
diff --git a/src/Relation/Binary/Properties/TotalOrder.agda b/src/Relation/Binary/Properties/TotalOrder.agda
new file mode 100644
index 0000000..5af5f11
--- /dev/null
+++ b/src/Relation/Binary/Properties/TotalOrder.agda
@@ -0,0 +1,22 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties satisfied by total orders
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Relation.Binary.Properties.TotalOrder
+ {t₁ t₂ t₃} (T : TotalOrder t₁ t₂ t₃) where
+
+open Relation.Binary.TotalOrder T
+open import Relation.Binary.Consequences
+
+decTotalOrder : Decidable _≈_ → DecTotalOrder _ _ _
+decTotalOrder ≟ = record
+ { isDecTotalOrder = record
+ { isTotalOrder = isTotalOrder
+ ; _≟_ = ≟
+ ; _≤?_ = total+dec⟶dec reflexive antisym total ≟
+ }
+ }
diff --git a/src/Relation/Binary/PropositionalEquality.agda b/src/Relation/Binary/PropositionalEquality.agda
index e408b62..9db5fee 100644
--- a/src/Relation/Binary/PropositionalEquality.agda
+++ b/src/Relation/Binary/PropositionalEquality.agda
@@ -32,6 +32,10 @@ cong : ∀ {a b} {A : Set a} {B : Set b}
(f : A → B) {x y} → x ≡ y → f x ≡ f y
cong f refl = refl
+cong-app : ∀ {a b} {A : Set a} {B : A → Set b} {f g : (x : A) → B x} →
+ f ≡ g → (x : A) → f x ≡ g x
+cong-app refl x = refl
+
cong₂ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
(f : A → B → C) {x y u v} → x ≡ y → u ≡ v → f x u ≡ f y v
cong₂ f refl refl = refl
diff --git a/src/Relation/Binary/Props/DecTotalOrder.agda b/src/Relation/Binary/Props/DecTotalOrder.agda
index 2623d62..c604222 100644
--- a/src/Relation/Binary/Props/DecTotalOrder.agda
+++ b/src/Relation/Binary/Props/DecTotalOrder.agda
@@ -1,7 +1,8 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Properties satisfied by decidable total orders
+-- Compatibility module. Pending for removal. Use
+-- Relation.Binary.Properties.DecTotalOrder instead.
------------------------------------------------------------------------
open import Relation.Binary
@@ -9,18 +10,4 @@ open import Relation.Binary
module Relation.Binary.Props.DecTotalOrder
{d₁ d₂ d₃} (DT : DecTotalOrder d₁ d₂ d₃) where
-open Relation.Binary.DecTotalOrder DT hiding (trans)
-import Relation.Binary.NonStrictToStrict as Conv
-open Conv _≈_ _≤_
-
-strictTotalOrder : StrictTotalOrder _ _ _
-strictTotalOrder = record
- { isStrictTotalOrder = record
- { isEquivalence = isEquivalence
- ; trans = trans isPartialOrder
- ; compare = trichotomous Eq.sym _≟_ antisym total
- ; <-resp-≈ = <-resp-≈ isEquivalence ≤-resp-≈
- }
- }
-
-open StrictTotalOrder strictTotalOrder public
+open import Relation.Binary.Properties.DecTotalOrder DT public
diff --git a/src/Relation/Binary/Props/Poset.agda b/src/Relation/Binary/Props/Poset.agda
index d40b234..6ae3495 100644
--- a/src/Relation/Binary/Props/Poset.agda
+++ b/src/Relation/Binary/Props/Poset.agda
@@ -1,7 +1,8 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Properties satisfied by posets
+-- Compatibility module. Pending for removal. Use
+-- Relation.Binary.Properties.Poset instead.
------------------------------------------------------------------------
open import Relation.Binary
@@ -9,21 +10,4 @@ open import Relation.Binary
module Relation.Binary.Props.Poset
{p₁ p₂ p₃} (P : Poset p₁ p₂ p₃) where
-open Relation.Binary.Poset P hiding (trans)
-import Relation.Binary.NonStrictToStrict as Conv
-open Conv _≈_ _≤_
-
-------------------------------------------------------------------------
--- Posets can be turned into strict partial orders
-
-strictPartialOrder : StrictPartialOrder _ _ _
-strictPartialOrder = record
- { isStrictPartialOrder = record
- { isEquivalence = isEquivalence
- ; irrefl = irrefl
- ; trans = trans isPartialOrder
- ; <-resp-≈ = <-resp-≈ isEquivalence ≤-resp-≈
- }
- }
-
-open StrictPartialOrder strictPartialOrder
+open import Relation.Binary.Properties.Poset P public
diff --git a/src/Relation/Binary/Props/Preorder.agda b/src/Relation/Binary/Props/Preorder.agda
index 41e8e68..6a0c767 100644
--- a/src/Relation/Binary/Props/Preorder.agda
+++ b/src/Relation/Binary/Props/Preorder.agda
@@ -1,7 +1,8 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Properties satisfied by preorders
+-- Compatibility module. Pending for removal. Use
+-- Relation.Binary.Properties.Preorder instead.
------------------------------------------------------------------------
open import Relation.Binary
@@ -9,20 +10,4 @@ open import Relation.Binary
module Relation.Binary.Props.Preorder
{p₁ p₂ p₃} (P : Preorder p₁ p₂ p₃) where
-open import Function
-open import Data.Product as Prod
-
-open Relation.Binary.Preorder P
-
-------------------------------------------------------------------------
--- For every preorder there is an induced equivalence
-
-InducedEquivalence : Setoid _ _
-InducedEquivalence = record
- { _≈_ = λ x y → x ∼ y × y ∼ x
- ; isEquivalence = record
- { refl = (refl , refl)
- ; sym = swap
- ; trans = Prod.zip trans (flip trans)
- }
- }
+open import Relation.Binary.Properties.Preorder P public
diff --git a/src/Relation/Binary/Props/StrictPartialOrder.agda b/src/Relation/Binary/Props/StrictPartialOrder.agda
index e9faf98..66ca71d 100644
--- a/src/Relation/Binary/Props/StrictPartialOrder.agda
+++ b/src/Relation/Binary/Props/StrictPartialOrder.agda
@@ -1,7 +1,8 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Properties satisfied by strict partial orders
+-- Compatibility module. Pending for removal. Use
+-- Relation.Binary.Properties.StrictPartialOrder instead.
------------------------------------------------------------------------
open import Relation.Binary
@@ -9,24 +10,4 @@ open import Relation.Binary
module Relation.Binary.Props.StrictPartialOrder
{s₁ s₂ s₃} (SPO : StrictPartialOrder s₁ s₂ s₃) where
-open Relation.Binary.StrictPartialOrder SPO
- renaming (trans to <-trans)
-import Relation.Binary.StrictToNonStrict as Conv
-open Conv _≈_ _<_
-
-------------------------------------------------------------------------
--- Strict partial orders can be converted to posets
-
-poset : Poset _ _ _
-poset = record
- { isPartialOrder = record
- { isPreorder = record
- { isEquivalence = isEquivalence
- ; reflexive = reflexive
- ; trans = trans isEquivalence <-resp-≈ <-trans
- }
- ; antisym = antisym isEquivalence <-trans irrefl
- }
- }
-
-open Poset poset public
+open import Relation.Binary.Properties.StrictPartialOrder SPO public
diff --git a/src/Relation/Binary/Props/StrictTotalOrder.agda b/src/Relation/Binary/Props/StrictTotalOrder.agda
index 2e42554..e9539d0 100644
--- a/src/Relation/Binary/Props/StrictTotalOrder.agda
+++ b/src/Relation/Binary/Props/StrictTotalOrder.agda
@@ -1,7 +1,8 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Properties satisfied by strict partial orders
+-- Compatibility module. Pending for removal. Use
+-- Relation.Binary.Properties.StrictTotalOrder instead.
------------------------------------------------------------------------
open import Relation.Binary
@@ -10,25 +11,4 @@ module Relation.Binary.Props.StrictTotalOrder
{s₁ s₂ s₃} (STO : StrictTotalOrder s₁ s₂ s₃)
where
-open Relation.Binary.StrictTotalOrder STO
-import Relation.Binary.StrictToNonStrict as Conv
-open Conv _≈_ _<_
-import Relation.Binary.Props.StrictPartialOrder as SPO
-open import Relation.Binary.Consequences
-
-------------------------------------------------------------------------
--- Strict total orders can be converted to decidable total orders
-
-decTotalOrder : DecTotalOrder _ _ _
-decTotalOrder = record
- { isDecTotalOrder = record
- { isTotalOrder = record
- { isPartialOrder = SPO.isPartialOrder strictPartialOrder
- ; total = total compare
- }
- ; _≟_ = _≟_
- ; _≤?_ = decidable' compare
- }
- }
-
-open DecTotalOrder decTotalOrder public
+open import Relation.Binary.Properties.StrictTotalOrder STO public
diff --git a/src/Relation/Binary/Props/TotalOrder.agda b/src/Relation/Binary/Props/TotalOrder.agda
index 4ae8b96..dd492cc 100644
--- a/src/Relation/Binary/Props/TotalOrder.agda
+++ b/src/Relation/Binary/Props/TotalOrder.agda
@@ -1,7 +1,8 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Properties satisfied by total orders
+-- Compatibility module. Pending for removal. Use
+-- Relation.Binary.Properties.TotalOrder instead.
------------------------------------------------------------------------
open import Relation.Binary
@@ -9,14 +10,4 @@ open import Relation.Binary
module Relation.Binary.Props.TotalOrder
{t₁ t₂ t₃} (T : TotalOrder t₁ t₂ t₃) where
-open Relation.Binary.TotalOrder T
-open import Relation.Binary.Consequences
-
-decTotalOrder : Decidable _≈_ → DecTotalOrder _ _ _
-decTotalOrder ≟ = record
- { isDecTotalOrder = record
- { isTotalOrder = isTotalOrder
- ; _≟_ = ≟
- ; _≤?_ = total+dec⟶dec reflexive antisym total ≟
- }
- }
+open import Relation.Binary.Properties.TotalOrder T public
diff --git a/src/Relation/Binary/SetoidReasoning.agda b/src/Relation/Binary/SetoidReasoning.agda
new file mode 100644
index 0000000..e3a3c98
--- /dev/null
+++ b/src/Relation/Binary/SetoidReasoning.agda
@@ -0,0 +1,46 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Convenient syntax for "equational reasoning" in multiple Setoids
+------------------------------------------------------------------------
+
+-- Example use:
+--
+-- open import Data.Maybe
+-- import Relation.Binary.SetoidReasoning as SetR
+--
+-- begin⟨ S ⟩
+-- x
+-- ≈⟨ drop-just (begin⟨ setoid S ⟩
+-- just x
+-- ≈⟨ justx≈mz ⟩
+-- mz
+-- ≈⟨ mz≈justy ⟩
+-- just y ∎) ⟩
+-- y
+-- ≈⟨ y≈z ⟩
+-- z ∎
+
+open import Relation.Binary.EqReasoning as EqR using (_IsRelatedTo_)
+open import Relation.Binary
+open import Relation.Binary.Core
+
+open Setoid
+
+module Relation.Binary.SetoidReasoning where
+
+infix 1 begin⟨_⟩_
+infixr 2 _≈⟨_⟩_ _≡⟨_⟩_
+infix 2 _∎
+
+begin⟨_⟩_ : ∀ {c l} (S : Setoid c l) → {x y : Carrier S} → _IsRelatedTo_ S x y → _≈_ S x y
+begin⟨_⟩_ S p = EqR.begin_ S p
+
+_∎ : ∀ {c l} {S : Setoid c l} → (x : Carrier S) → _IsRelatedTo_ S x x
+_∎ {S = S} = EqR._∎ S
+
+_≈⟨_⟩_ : ∀ {c l} {S : Setoid c l} → (x : Carrier S) → {y z : Carrier S} → _≈_ S x y → _IsRelatedTo_ S y z → _IsRelatedTo_ S x z
+_≈⟨_⟩_ {S = S} = EqR._≈⟨_⟩_ S
+
+_≡⟨_⟩_ : ∀ {c l} {S : Setoid c l} → (x : Carrier S) → {y z : Carrier S} → x ≡ y → _IsRelatedTo_ S y z → _IsRelatedTo_ S x z
+_≡⟨_⟩_ {S = S} = EqR._≡⟨_⟩_ S
diff --git a/src/Relation/Binary/StrictPartialOrderReasoning.agda b/src/Relation/Binary/StrictPartialOrderReasoning.agda
index facb252..62a13eb 100644
--- a/src/Relation/Binary/StrictPartialOrderReasoning.agda
+++ b/src/Relation/Binary/StrictPartialOrderReasoning.agda
@@ -11,5 +11,5 @@ module Relation.Binary.StrictPartialOrderReasoning
{p₁ p₂ p₃} (S : StrictPartialOrder p₁ p₂ p₃) where
import Relation.Binary.PreorderReasoning as PreR
-import Relation.Binary.Props.StrictPartialOrder as SPO
+import Relation.Binary.Properties.StrictPartialOrder as SPO
open PreR (SPO.preorder S) public renaming (_∼⟨_⟩_ to _<⟨_⟩_)
diff --git a/src/Relation/Binary/Sum.agda b/src/Relation/Binary/Sum.agda
index 90c56f3..8a84e63 100644
--- a/src/Relation/Binary/Sum.agda
+++ b/src/Relation/Binary/Sum.agda
@@ -25,8 +25,9 @@ open import Function.Surjection as Surj
using (Surjection; _↠_; module Surjection)
open import Level
open import Relation.Nullary
+import Relation.Nullary.Decidable as Dec
open import Relation.Binary
-import Relation.Binary.PropositionalEquality as P
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
module _ {a₁ a₂} {A₁ : Set a₁} {A₂ : Set a₂} where
@@ -403,6 +404,18 @@ sto₁ ⊎-<-strictTotalOrder sto₂ = record
------------------------------------------------------------------------
-- Some properties related to "relatedness"
+private
+
+ to-cong : ∀ {a b} {A : Set a} {B : Set b} →
+ (_≡_ ⊎-Rel _≡_) ⇒ _≡_ {A = A ⊎ B}
+ to-cong (₁∼₂ ())
+ to-cong (₁∼₁ P.refl) = P.refl
+ to-cong (₂∼₂ P.refl) = P.refl
+
+ from-cong : ∀ {a b} {A : Set a} {B : Set b} →
+ _≡_ {A = A ⊎ B} ⇒ (_≡_ ⊎-Rel _≡_)
+ from-cong P.refl = P.refl ⊎-refl P.refl
+
⊎-Rel↔≡ : ∀ {a b} (A : Set a) (B : Set b) →
Inverse (P.setoid A ⊎-setoid P.setoid B) (P.setoid (A ⊎ B))
⊎-Rel↔≡ _ _ = record
@@ -413,14 +426,13 @@ sto₁ ⊎-<-strictTotalOrder sto₂ = record
; right-inverse-of = λ _ → P.refl
}
}
- where
- to-cong : (P._≡_ ⊎-Rel P._≡_) ⇒ P._≡_
- to-cong (₁∼₂ ())
- to-cong (₁∼₁ P.refl) = P.refl
- to-cong (₂∼₂ P.refl) = P.refl
- from-cong : P._≡_ ⇒ (P._≡_ ⊎-Rel P._≡_)
- from-cong P.refl = P.refl ⊎-refl P.refl
+_⊎-≟_ : ∀ {a b} {A : Set a} {B : Set b} →
+ Decidable {A = A} _≡_ → Decidable {A = B} _≡_ →
+ Decidable {A = A ⊎ B} _≡_
+(dec₁ ⊎-≟ dec₂) s₁ s₂ = Dec.map′ to-cong from-cong (s₁ ≟ s₂)
+ where
+ open DecSetoid (P.decSetoid dec₁ ⊎-decSetoid P.decSetoid dec₂)
_⊎-⟶_ :
∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈}
diff --git a/src/Relation/Nullary/Decidable.agda b/src/Relation/Nullary/Decidable.agda
index 7c20aff..a7bd74f 100644
--- a/src/Relation/Nullary/Decidable.agda
+++ b/src/Relation/Nullary/Decidable.agda
@@ -11,10 +11,12 @@ 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.Equality using (_⟨$⟩_; module Π)
open import Function.Equivalence
using (_⇔_; equivalence; module Equivalence)
+open import Function.Injection using (Injection; module Injection)
open import Level using (Lift)
+open import Relation.Binary using (Setoid; module Setoid; Decidable)
open import Relation.Binary.PropositionalEquality
open import Relation.Nullary
@@ -58,6 +60,21 @@ map′ : ∀ {p q} {P : Set p} {Q : Set q} →
(P → Q) → (Q → P) → Dec P → Dec Q
map′ P→Q Q→P = map (equivalence P→Q Q→P)
+module _ {a₁ a₂ b₁ b₂} {A : Setoid a₁ a₂} {B : Setoid b₁ b₂} where
+
+ open Injection
+ open Setoid A using () renaming (_≈_ to _≈A_)
+ open Setoid B using () renaming (_≈_ to _≈B_)
+
+ -- If there is an injection from one setoid to another, and the
+ -- latter's equivalence relation is decidable, then the former's
+ -- equivalence relation is also decidable.
+
+ via-injection : Injection A B → Decidable _≈B_ → Decidable _≈A_
+ via-injection inj dec x y with dec (to inj ⟨$⟩ x) (to inj ⟨$⟩ y)
+ ... | yes injx≈injy = yes (Injection.injective inj injx≈injy)
+ ... | no injx≉injy = no (λ x≈y → injx≉injy (Π.cong (to inj) x≈y))
+
-- If a decision procedure returns "yes", then we can extract the
-- proof using from-yes.
diff --git a/src/Relation/Nullary/Sum.agda b/src/Relation/Nullary/Sum.agda
index c3663ab..7e65a50 100644
--- a/src/Relation/Nullary/Sum.agda
+++ b/src/Relation/Nullary/Sum.agda
@@ -13,6 +13,8 @@ open import Relation.Nullary
-- Some properties which are preserved by _⊎_.
+infixr 1 _¬-⊎_ _⊎-dec_
+
_¬-⊎_ : ∀ {p q} {P : Set p} {Q : Set q} →
¬ P → ¬ Q → ¬ (P ⊎ Q)
(¬p ¬-⊎ ¬q) (inj₁ p) = ¬p p
diff --git a/src/Relation/Unary.agda b/src/Relation/Unary.agda
index 004b20a..03aa8e3 100644
--- a/src/Relation/Unary.agda
+++ b/src/Relation/Unary.agda
@@ -8,11 +8,12 @@ module Relation.Unary where
open import Data.Empty
open import Function
-open import Data.Unit
+open import Data.Unit hiding (setoid)
open import Data.Product
open import Data.Sum
open import Level
open import Relation.Nullary
+open import Relation.Binary using (Setoid; IsEquivalence)
------------------------------------------------------------------------
-- Unary relations
@@ -97,30 +98,66 @@ module _ {a} {A : Set a} -- The universe of discourse.
⊆-U : ∀ {ℓ} → (P : Pred A ℓ) → P ⊆ U
⊆-U P _ = _
+ -- Positive version of non-disjointness, dual to inclusion.
+
+ infix 4 _≬_
+
+ _≬_ : ∀ {ℓ₁ ℓ₂} → Pred A ℓ₁ → Pred A ℓ₂ → Set _
+ P ≬ Q = ∃ λ x → x ∈ P × x ∈ Q
+
-- Set union.
- infixl 6 _∪_
+ infixr 6 _∪_
_∪_ : ∀ {ℓ₁ ℓ₂} → Pred A ℓ₁ → Pred A ℓ₂ → Pred A _
P ∪ Q = λ x → x ∈ P ⊎ x ∈ Q
-- Set intersection.
- infixl 7 _∩_
+ infixr 7 _∩_
_∩_ : ∀ {ℓ₁ ℓ₂} → Pred A ℓ₁ → Pred A ℓ₂ → Pred A _
P ∩ Q = λ x → x ∈ P × x ∈ Q
+ -- Implication.
+
+ infixl 8 _⇒_
+
+ _⇒_ : ∀ {ℓ₁ ℓ₂} → Pred A ℓ₁ → Pred A ℓ₂ → Pred A _
+ P ⇒ Q = λ x → x ∈ P → x ∈ Q
+
+ -- Infinitary union and intersection.
+
+ infix 9 ⋃ ⋂
+
+ ⋃ : ∀ {ℓ i} (I : Set i) → (I → Pred A ℓ) → Pred A _
+ ⋃ I P = λ x → Σ[ i ∈ I ] P i x
+
+ syntax ⋃ I (λ i → P) = ⋃[ i ∶ I ] P
+
+ ⋂ : ∀ {ℓ i} (I : Set i) → (I → Pred A ℓ) → Pred A _
+ ⋂ I P = λ x → (i : I) → P i x
+
+ syntax ⋂ I (λ i → P) = ⋂[ i ∶ I ] P
+
------------------------------------------------------------------------
-- Unary relation combinators
infixr 2 _⟨×⟩_
+infixr 2 _⟨⊙⟩_
infixr 1 _⟨⊎⟩_
infixr 0 _⟨→⟩_
+infixl 9 _⟨·⟩_
+infixr 9 _⟨∘⟩_
+infixr 2 _//_ _\\_
_⟨×⟩_ : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b} →
Pred A ℓ₁ → Pred B ℓ₂ → Pred (A × B) _
-P ⟨×⟩ Q = uncurry (λ p q → P p × Q q)
+(P ⟨×⟩ Q) (x , y) = x ∈ P × y ∈ Q
+
+_⟨⊙⟩_ : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b} →
+ Pred A ℓ₁ → Pred B ℓ₂ → Pred (A × B) _
+(P ⟨⊙⟩ Q) (x , y) = x ∈ P ⊎ y ∈ Q
_⟨⊎⟩_ : ∀ {a b ℓ} {A : Set a} {B : Set b} →
Pred A ℓ → Pred B ℓ → Pred (A ⊎ B) _
@@ -130,8 +167,35 @@ _⟨→⟩_ : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b} →
Pred A ℓ₁ → Pred B ℓ₂ → Pred (A → B) _
(P ⟨→⟩ Q) f = P ⊆ Q ∘ f
+_⟨·⟩_ : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b}
+ (P : Pred A ℓ₁) (Q : Pred B ℓ₂) →
+ (P ⟨×⟩ (P ⟨→⟩ Q)) ⊆ Q ∘ uncurry (flip _$_)
+(P ⟨·⟩ Q) (p , f) = f p
+
+-- Converse.
+
+_~ : ∀ {a b ℓ} {A : Set a} {B : Set b} →
+ Pred (A × B) ℓ → Pred (B × A) ℓ
+P ~ = P ∘ swap
+
+-- Composition.
+
+_⟨∘⟩_ : ∀ {a b c ℓ₁ ℓ₂} {A : Set a} {B : Set b} {C : Set c} →
+ Pred (A × B) ℓ₁ → Pred (B × C) ℓ₂ → Pred (A × C) _
+(P ⟨∘⟩ Q) (x , z) = ∃ λ y → (x , y) ∈ P × (y , z) ∈ Q
+
+-- Post and pre-division.
+
+_//_ : ∀ {a b c ℓ₁ ℓ₂} {A : Set a} {B : Set b} {C : Set c} →
+ Pred (A × C) ℓ₁ → Pred (B × C) ℓ₂ → Pred (A × B) _
+(P // Q) (x , y) = Q ∘ _,_ y ⊆ P ∘ _,_ x
+
+_\\_ : ∀ {a b c ℓ₁ ℓ₂} {A : Set a} {B : Set b} {C : Set c} →
+ Pred (A × C) ℓ₁ → Pred (A × B) ℓ₂ → Pred (B × C) _
+P \\ Q = (P ~ // Q ~) ~
+
------------------------------------------------------------------------
-- Properties of unary relations
-Decidable : ∀ {a p} {A : Set a} (P : A → Set p) → Set _
+Decidable : ∀ {a ℓ} {A : Set a} (P : Pred A ℓ) → Set _
Decidable P = ∀ x → Dec (P x)
diff --git a/src/Relation/Unary/PredicateTransformer.agda b/src/Relation/Unary/PredicateTransformer.agda
new file mode 100644
index 0000000..05b684d
--- /dev/null
+++ b/src/Relation/Unary/PredicateTransformer.agda
@@ -0,0 +1,117 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Predicate transformers
+------------------------------------------------------------------------
+
+module Relation.Unary.PredicateTransformer where
+
+open import Level hiding (_⊔_)
+open import Function
+open import Data.Product
+open import Relation.Nullary
+open import Relation.Unary
+open import Relation.Binary using (REL)
+
+------------------------------------------------------------------------
+
+-- Heterogeneous and homogeneous predicate transformers
+
+PT : ∀ {a b} → Set a → Set b → (ℓ₁ ℓ₂ : Level) → Set _
+PT A B ℓ₁ ℓ₂ = Pred A ℓ₁ → Pred B ℓ₂
+
+Pt : ∀ {a} → Set a → (ℓ : Level) → Set _
+Pt A ℓ = PT A A ℓ ℓ
+
+-- Composition and identity
+
+_⍮_ : ∀ {a b c ℓ₁ ℓ₂ ℓ₃} {A : Set a} {B : Set b} {C : Set c} →
+ PT B C ℓ₂ ℓ₃ → PT A B ℓ₁ ℓ₂ → PT A C ℓ₁ _
+S ⍮ T = S ∘ T
+
+skip : ∀ {a ℓ} {A : Set a} → PT A A ℓ ℓ
+skip P = P
+
+------------------------------------------------------------------------
+-- Operations on predicates extend pointwise to predicate transformers
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ -- The bottom and the top of the predicate transformer lattice.
+
+ abort : PT A B zero zero
+ abort = λ _ → ∅
+
+ magic : PT A B zero zero
+ magic = λ _ → U
+
+ -- Negation.
+
+ ∼_ : ∀ {ℓ₁ ℓ₂} → PT A B ℓ₁ ℓ₂ → PT A B ℓ₁ ℓ₂
+ ∼ T = ∁ ∘ T
+
+ -- Refinement.
+
+ infix 4 _⊑_ _⊒_ _⊑′_ _⊒′_
+
+ _⊑_ : ∀ {ℓ₁ ℓ₂} → PT A B ℓ₁ ℓ₂ → PT A B ℓ₁ ℓ₂ → Set _
+ S ⊑ T = ∀ {X} → S X ⊆ T X
+
+ _⊑′_ : ∀ {ℓ₁ ℓ₂} → PT A B ℓ₁ ℓ₂ → PT A B ℓ₁ ℓ₂ → Set _
+ S ⊑′ T = ∀ X → S X ⊆ T X
+
+ _⊒_ : ∀ {ℓ₁ ℓ₂} → PT A B ℓ₁ ℓ₂ → PT A B ℓ₁ ℓ₂ → Set _
+ T ⊒ S = T ⊑ S
+
+ _⊒′_ : ∀ {ℓ₁ ℓ₂} → PT A B ℓ₁ ℓ₂ → PT A B ℓ₁ ℓ₂ → Set _
+ T ⊒′ S = S ⊑′ T
+
+ -- The dual of refinement.
+
+ infix 4 _⋢_
+
+ _⋢_ : ∀ {ℓ₁ ℓ₂} → PT A B ℓ₁ ℓ₂ → PT A B ℓ₁ ℓ₂ → Set _
+ S ⋢ T = ∃ λ X → S X ≬ T X
+
+ -- Union.
+
+ infixl 6 _⊓_
+
+ _⊓_ : ∀ {ℓ₁ ℓ₂} → PT A B ℓ₁ ℓ₂ → PT A B ℓ₁ ℓ₂ → PT A B ℓ₁ ℓ₂
+ S ⊓ T = λ X → S X ∪ T X
+
+ -- Intersection.
+
+ infixl 7 _⊔_
+
+ _⊔_ : ∀ {ℓ₁ ℓ₂} → PT A B ℓ₁ ℓ₂ → PT A B ℓ₁ ℓ₂ → PT A B ℓ₁ ℓ₂
+ S ⊔ T = λ X → S X ∩ T X
+
+ -- Implication.
+
+ infixl 8 _⇛_
+
+ _⇛_ : ∀ {ℓ₁ ℓ₂} → PT A B ℓ₁ ℓ₂ → PT A B ℓ₁ ℓ₂ → PT A B ℓ₁ ℓ₂
+ S ⇛ T = λ X → S X ⇒ T X
+
+ -- Infinitary union and intersection.
+
+ infix 9 ⨆ ⨅
+
+ ⨆ : ∀ {ℓ₁ ℓ₂ i} (I : Set i) → (I → PT A B ℓ₁ ℓ₂) → PT A B ℓ₁ _
+ ⨆ I T = λ X → ⋃[ i ∶ I ] T i X
+
+ syntax ⨆ I (λ i → T) = ⨆[ i ∶ I ] T
+
+ ⨅ : ∀ {ℓ₁ ℓ₂ i} (I : Set i) → (I → PT A B ℓ₁ ℓ₂) → PT A B ℓ₁ _
+ ⨅ I T = λ X → ⋂[ i ∶ I ] T i X
+
+ syntax ⨅ I (λ i → T) = ⨅[ i ∶ I ] T
+
+ -- Angelic and demonic update.
+
+ ⟨_⟩ : ∀ {ℓ} → REL A B ℓ → PT B A ℓ _
+ ⟨ R ⟩ P = λ x → R x ≬ P
+
+ [_] : ∀ {ℓ} → REL A B ℓ → PT B A ℓ _
+ [ R ] P = λ x → R x ⊆ P
diff --git a/src/Size.agda b/src/Size.agda
index 25614ac..992d44c 100644
--- a/src/Size.agda
+++ b/src/Size.agda
@@ -7,10 +7,12 @@
module Size where
postulate
- Size : Set
- ↑_ : Size → Size
- ∞ : Size
+ Size : Set
+ Size<_ : Size → Set
+ ↑_ : Size → Size
+ ∞ : Size
-{-# BUILTIN SIZE Size #-}
-{-# BUILTIN SIZESUC ↑_ #-}
-{-# BUILTIN SIZEINF ∞ #-}
+{-# BUILTIN SIZE Size #-}
+{-# BUILTIN SIZELT Size<_ #-}
+{-# BUILTIN SIZESUC ↑_ #-}
+{-# BUILTIN SIZEINF ∞ #-}