lib-0.7/0000755000000000000000000000000012101774707010322 5ustar0000000000000000lib-0.7/README.agda0000644000000000000000000002435612101774706012106 0ustar0000000000000000module README where ------------------------------------------------------------------------ -- The Agda standard library, version 0.7 -- -- 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 ------------------------------------------------------------------------ -- This version of the library has been tested using Agda 2.3.2. -- Note that no guarantees are currently made about forwards or -- backwards compatibility, the library is still at an experimental -- stage. -- To make use of the library, add the path to the library’s root -- directory (src) to the Agda search path, either using the -- --include-path flag or by customising the Emacs mode variable -- agda2-include-dirs (M-x customize-group RET agda2 RET). -- To compile the library using the MAlonzo compiler you first need to -- install some supporting Haskell code, for instance as follows: -- -- cd ffi -- cabal install -- -- Currently the library does not support the Epic or JavaScript -- compiler backends. -- Contributions to this library are welcome (but to avoid wasted work -- it is suggested that you discuss large changes before implementing -- them). Please send contributions in the form of darcs patches (run -- darcs send --output  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. ------------------------------------------------------------------------ -- Module hierarchy ------------------------------------------------------------------------ -- The top-level module names of the library are currently allocated -- as follows: -- -- • Algebra -- Abstract algebra (monoids, groups, rings etc.), along with -- properties needed to specify these structures (associativity, -- commutativity, etc.), and operations on and proofs about the -- structures. -- • Category -- Category theory-inspired idioms used to structure functional -- programs (functors and monads, for instance). -- • Coinduction -- Support for coinduction. -- • Data -- Data types and properties about data types. -- • Function -- Combinators and properties related to functions. -- • Foreign -- Related to the foreign function interface. -- • Induction -- A general framework for induction (includes lexicographic and -- well-founded induction). -- • IO -- Input/output-related functions. -- • Irrelevance -- Definitions related to (proscriptive) irrelevance. -- • Level -- Universe levels. -- • Record -- An encoding of record types with manifest fields and "with". -- • Reflection -- Support for reflection. -- • Relation -- Properties of and proofs about relations (mostly homogeneous -- binary relations). -- • Size -- Sizes used by the sized types mechanism. -- • Universe -- A definition of universes. ------------------------------------------------------------------------ -- A selection of useful library modules ------------------------------------------------------------------------ -- Note that module names in source code are often hyperlinked to the -- corresponding module. In the Emacs mode you can follow these -- hyperlinks by typing M-. or clicking with the middle mouse button. -- • Some data types import Data.Bool -- Booleans. import Data.Char -- Characters. import Data.Empty -- The empty type. import Data.Fin -- Finite sets. import Data.List -- Lists. import Data.Maybe -- The maybe type. import Data.Nat -- Natural numbers. import Data.Product -- Products. import Data.Stream -- Streams. import Data.String -- Strings. import Data.Sum -- Disjoint sums. import Data.Unit -- The unit type. import Data.Vec -- Fixed-length vectors. -- • Some types used to structure computations import Category.Functor -- Functors. import Category.Applicative -- Applicative functors. import Category.Monad -- Monads. -- • Equality -- Propositional equality: import Relation.Binary.PropositionalEquality -- Convenient syntax for "equational reasoning" using a preorder: import Relation.Binary.PreorderReasoning -- Solver for commutative ring or semiring equalities: import Algebra.RingSolver -- • Properties of functions, sets and relations -- Monoids, rings and similar algebraic structures: import Algebra -- Negation, decidability, and similar operations on sets: import Relation.Nullary -- Properties of homogeneous binary relations: import Relation.Binary -- • Induction -- An abstraction of various forms of recursion/induction: import Induction -- Well-founded induction: import Induction.WellFounded -- Various forms of induction for natural numbers: import Induction.Nat -- • Support for coinduction import Coinduction -- • IO import IO ------------------------------------------------------------------------ -- Record hierarchies ------------------------------------------------------------------------ -- When an abstract hierarchy of some sort (for instance semigroup → -- monoid → group) is included in the library the basic approach is to -- specify the properties of every concept in terms of a record -- containing just properties, parameterised on the underlying -- operations, sets etc.: -- -- record IsSemigroup {A} (≈ : Rel A) (∙ : Op₂ A) : Set where -- open FunctionProperties ≈ -- field -- isEquivalence : IsEquivalence ≈ -- assoc : Associative ∙ -- ∙-cong : ∙ Preserves₂ ≈ ⟶ ≈ ⟶ ≈ -- -- More specific concepts are then specified in terms of the simpler -- ones: -- -- record IsMonoid {A} (≈ : Rel A) (∙ : Op₂ A) (ε : A) : Set where -- open FunctionProperties ≈ -- field -- isSemigroup : IsSemigroup ≈ ∙ -- identity : Identity ε ∙ -- -- open IsSemigroup isSemigroup public -- -- Note here that open IsSemigroup isSemigroup public ensures that the -- fields of the isSemigroup record can be accessed directly; this -- technique enables the user of an IsMonoid record to use underlying -- records without having to manually open an entire record hierarchy. -- This is not always possible, though. Consider the following definition -- of preorders: -- -- record IsPreorder {A : Set} -- (_≈_ : Rel A) -- The underlying equality. -- (_∼_ : Rel A) -- The relation. -- : Set where -- field -- isEquivalence : IsEquivalence _≈_ -- -- Reflexivity is expressed in terms of an underlying equality: -- reflexive : _≈_ ⇒ _∼_ -- trans : Transitive _∼_ -- -- module Eq = IsEquivalence isEquivalence -- -- ... -- -- The Eq module in IsPreorder is not opened publicly, because it -- contains some fields which clash with fields or other definitions -- in IsPreorder. -- Records packing up properties with the corresponding operations, -- sets, etc. are sometimes also defined: -- -- record Semigroup : Set₁ where -- infixl 7 _∙_ -- infix 4 _≈_ -- field -- Carrier : Set -- _≈_ : Rel Carrier -- _∙_ : Op₂ Carrier -- isSemigroup : IsSemigroup _≈_ _∙_ -- -- open IsSemigroup isSemigroup public -- -- setoid : Setoid -- setoid = record { isEquivalence = isEquivalence } -- -- record Monoid : Set₁ where -- infixl 7 _∙_ -- infix 4 _≈_ -- field -- Carrier : Set -- _≈_ : Rel Carrier -- _∙_ : Op₂ Carrier -- ε : Carrier -- isMonoid : IsMonoid _≈_ _∙_ ε -- -- open IsMonoid isMonoid public -- -- semigroup : Semigroup -- semigroup = record { isSemigroup = isSemigroup } -- -- open Semigroup semigroup public using (setoid) -- -- Note that the Monoid record does not include a Semigroup field. -- Instead the Monoid /module/ includes a "repackaging function" -- semigroup which converts a Monoid to a Semigroup. -- The above setup may seem a bit complicated, but we think it makes the -- library quite easy to work with, while also providing enough -- flexibility. ------------------------------------------------------------------------ -- More documentation ------------------------------------------------------------------------ -- Some examples showing where the natural numbers/integers and some -- related operations and properties are defined, and how they can be -- used: import README.Nat import README.Integer -- Some examples showing how the AVL tree module can be used. import README.AVL -- An example showing how the Record module can be used. import README.Record -- An example showing how the case expression can be used. import README.Case ------------------------------------------------------------------------ -- Core modules ------------------------------------------------------------------------ -- Some modules have names ending in ".Core". These modules are -- internal, and have (mostly) been created to avoid mutual recursion -- between modules. They should not be imported directly; their -- contents are reexported by other modules. ------------------------------------------------------------------------ -- All library modules ------------------------------------------------------------------------ -- For short descriptions of every library module, see Everything: import Everything -- Note that the Everything module is generated automatically. If you -- have downloaded the library from its darcs repository and want to -- type check README then you can (try to) construct Everything by -- running "cabal install && GenerateEverything". -- Note that all library sources are located under src or ffi. The -- modules README, README.* and Everything are not really part of the -- library, so these modules are located in the top-level directory -- instead. lib-0.7/LICENCE0000644000000000000000000000264112101774706011311 0ustar0000000000000000Copyright (c) 2007-2013 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 Permission is hereby granted, free of charge, to any person obtaining a copy of this software and associated documentation files (the "Software"), to deal in the Software without restriction, including without limitation the rights to use, copy, modify, merge, publish, distribute, sublicense, and/or sell copies of the Software, and to permit persons to whom the Software is furnished to do so, subject to the following conditions: The above copyright notice and this permission notice shall be included in all copies or substantial portions of the Software. THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE SOFTWARE. lib-0.7/Header0000644000000000000000000000045612101774706011441 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- All library modules, along with short descriptions ------------------------------------------------------------------------ -- Note that core modules are not included. module Everything where lib-0.7/Setup.hs0000644000000000000000000000005712101774706011757 0ustar0000000000000000import Distribution.Simple main = defaultMain lib-0.7/Everything.agda0000644000000000000000000003070512101774707013271 0ustar0000000000000000------------------------------------------------------------------------ -- 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 lib-0.7/.boring0000644000000000000000000000012212101774706011575 0ustar0000000000000000\.l?agda\.el$ \.agdai$ (^|/)MAlonzo($|/) ^dist($|/) ^html($|/) ^Everything\.agda$ lib-0.7/AllNonAsciiChars.hs0000644000000000000000000000155712101774706014002 0ustar0000000000000000-- | This module extracts all the non-ASCII characters used by the -- library code (along with how many times they are used). module Main where import qualified Data.List as L import Data.Char import Data.Function import Control.Applicative import System.FilePath.Find import System.IO readUTF8File :: FilePath -> IO String readUTF8File f = do h <- openFile f ReadMode hSetEncoding h utf8 hGetContents h main :: IO () main = do agdaFiles <- find always (extension ==? ".agda" ||? extension ==? ".lagda") "src" nonAsciiChars <- filter (not . isAscii) . concat <$> mapM readUTF8File agdaFiles let table = reverse $ L.sortBy (compare `on` snd) $ map (\cs -> (head cs, length cs)) $ L.group $ L.sort $ nonAsciiChars mapM_ (\(c, count) -> putStrLn (c : ": " ++ show count)) table lib-0.7/GNUmakefile0000644000000000000000000000035712101774706012400 0ustar0000000000000000AGDA=agda test: Everything.agda $(AGDA) -i. -isrc README.agda setup: Everything.agda agda-lib-ffi .PHONY: Everything.agda Everything.agda: cabal install GenerateEverything .PHONY: agda-lib-ffi agda-lib-ffi: cd ffi && cabal install lib-0.7/lib.cabal0000644000000000000000000000101212101774706012045 0ustar0000000000000000name: lib version: 0.1.0.1 cabal-version: >= 1.8 build-type: Simple description: Helper programs. executable GenerateEverything hs-source-dirs: . main-is: GenerateEverything.hs build-depends: base >= 4.2 && < 4.7, filemanip == 0.3.*, filepath >= 1.1 && < 1.4 executable AllNonAsciiChars hs-source-dirs: . main-is: AllNonAsciiChars.hs build-depends: base >= 4.2 && < 4.7, filemanip == 0.3.* lib-0.7/release-notes0000644000000000000000000000761312101774706013021 0ustar0000000000000000------------------------------------------------------------------------ Version 0.7 ------------------------------------------------------------------------ Version 0.7 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.3.2. 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.6 ------------------------------------------------------------------------ Version 0.6 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.3.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.5 ------------------------------------------------------------------------ Version 0.5 of the standard library has now been released, see http://wiki.portal.chalmers.se/agda/pmwiki.php?n=Libraries.StandardLibrary. The library has been tested using Agda version 2.2.10. Note that no guarantees are made about backwards or forwards compatibility, the library is still at an experimental stage. ------------------------------------------------------------------------ Version 0.4 ------------------------------------------------------------------------ Version 0.4 of the standard library has now been released, see http://wiki.portal.chalmers.se/agda/pmwiki.php?n=Libraries.StandardLibrary. The library has been tested using Agda version 2.2.8. Note that no guarantees are made about backwards or forwards compatibility, the library is still at an experimental stage. ------------------------------------------------------------------------ Version 0.3 ------------------------------------------------------------------------ Version 0.3 of the standard library has now been released, see http://wiki.portal.chalmers.se/agda/pmwiki.php?n=Libraries.StandardLibrary. The library has been tested using Agda version 2.2.6. Note that no guarantees are made about backwards or forwards compatibility, the library is still at an experimental stage. ------------------------------------------------------------------------ Version 0.2 ------------------------------------------------------------------------ Version 0.2 of the "standard" library has now been released, see http://wiki.portal.chalmers.se/agda/pmwiki.php?n=Libraries.StandardLibrary. The library has been tested using Agda version 2.2.4. Note that no guarantees are made about backwards or forwards compatibility, the library is still at an experimental stage. Note also that the library sources are now located in the sub-directory lib-/src of the installation tarball. ------------------------------------------------------------------------ Version 0.1 ------------------------------------------------------------------------ Version 0.1 of the "standard" library has now been released, see http://wiki.portal.chalmers.se/agda/pmwiki.php?n=Libraries.StandardLibrary. The library has been tested using Agda version 2.2.2. Note that no guarantees are made about backwards or forwards compatibility, the library is still at an experimental stage. lib-0.7/GenerateEverything.hs0000644000000000000000000000550612101774706014462 0ustar0000000000000000{-# LANGUAGE PatternGuards #-} import qualified Data.List as List import Control.Applicative import System.Environment import System.IO import System.Exit import System.FilePath import System.FilePath.Find headerFile = "Header" outputFile = "Everything.agda" srcDir = "src" main = do args <- getArgs case args of [] -> return () _ -> hPutStr stderr usage >> exitFailure header <- readFileUTF8 headerFile modules <- filter isLibraryModule . List.sort <$> find always (extension ==? ".agda" ||? extension ==? ".lagda") srcDir headers <- mapM extractHeader modules writeFileUTF8 outputFile $ header ++ format (zip modules headers) -- | Usage info. usage :: String usage = unlines [ "GenerateEverything: A utility program for Agda's standard library." , "" , "Usage: GenerateEverything" , "" , "This program should be run in the base directory of a clean checkout of" , "the library." , "" , "The program generates documentation for the library by extracting" , "headers from library modules. The output is written to " ++ outputFile , "with the file " ++ headerFile ++ " inserted verbatim at the beginning." ] -- | Returns 'True' for all Agda files except for core modules. isLibraryModule :: FilePath -> Bool isLibraryModule f = takeExtension f `elem` [".agda", ".lagda"] && dropExtension (takeFileName f) /= "Core" -- | Reads a module and extracts the header. extractHeader :: FilePath -> IO [String] extractHeader mod = fmap (extract . lines) $ readFileUTF8 mod where delimiter = all (== '-') extract (d1 : "-- The Agda standard library" : "--" : ss) | delimiter d1 , (info, d2 : rest) <- span ("-- " `List.isPrefixOf`) ss , delimiter d2 = info extract _ = error $ mod ++ " is malformed." -- | Formats the extracted module information. format :: [(FilePath, [String])] -- ^ Pairs of module names and headers. All lines in the -- headers are already prefixed with \"-- \". -> String format = unlines . concat . map fmt where fmt (mod, header) = "" : header ++ ["import " ++ fileToMod mod] -- | Translates a file name to the corresponding module name. It is -- assumed that the file name corresponds to an Agda module under -- 'srcDir'. fileToMod :: FilePath -> String fileToMod = map slashToDot . dropExtension . makeRelative srcDir where slashToDot c | isPathSeparator c = '.' | otherwise = c -- | A variant of 'readFile' which uses the 'utf8' encoding. readFileUTF8 :: FilePath -> IO String readFileUTF8 f = do h <- openFile f ReadMode hSetEncoding h utf8 hGetContents h -- | A variant of 'writeFile' which uses the 'utf8' encoding. writeFileUTF8 :: FilePath -> String -> IO () writeFileUTF8 f s = withFile f WriteMode $ \h -> do hSetEncoding h utf8 hPutStr h s lib-0.7/src/0000755000000000000000000000000012101774706011110 5ustar0000000000000000lib-0.7/src/Coinduction.agda0000644000000000000000000000302712101774706014206 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Basic types related to coinduction ------------------------------------------------------------------------ module Coinduction where import Level ------------------------------------------------------------------------ -- A type used to make recursive arguments coinductive infix 1000 ♯_ postulate ∞ : ∀ {a} (A : Set a) → Set a ♯_ : ∀ {a} {A : Set a} → A → ∞ A ♭ : ∀ {a} {A : Set a} → ∞ A → A {-# BUILTIN INFINITY ∞ #-} {-# BUILTIN SHARP ♯_ #-} {-# BUILTIN FLAT ♭ #-} ------------------------------------------------------------------------ -- Rec, a type which is analogous to the Rec type constructor used in -- (the current version of) ΠΣ data Rec {a} (A : ∞ (Set a)) : Set a where fold : (x : ♭ A) → Rec A unfold : ∀ {a} {A : ∞ (Set a)} → Rec A → ♭ A unfold (fold x) = x {- -- If --guardedness-preserving-type-constructors is enabled one can -- define types like ℕ by recursion: open import Data.Sum open import Data.Unit ℕ : Set ℕ = ⊤ ⊎ Rec (♯ ℕ) zero : ℕ zero = inj₁ _ suc : ℕ → ℕ suc n = inj₂ (fold n) ℕ-rec : (P : ℕ → Set) → P zero → (∀ n → P n → P (suc n)) → ∀ n → P n ℕ-rec P z s (inj₁ _) = z ℕ-rec P z s (inj₂ (fold n)) = s n (ℕ-rec P z s n) -- This feature is very experimental, though: it may lead to -- inconsistencies. -} lib-0.7/src/Irrelevance.agda0000644000000000000000000000102512101774706014163 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- The irrelevance axiom ------------------------------------------------------------------------ module Irrelevance where import Level ------------------------------------------------------------------------ -- The irrelevance axiom -- There is no guarantee that this axiom is sound. Use it at your own -- risk. postulate .irrelevance-axiom : ∀ {a} {A : Set a} → .A → A {-# BUILTIN IRRAXIOM irrelevance-axiom #-} lib-0.7/src/IO.agda0000644000000000000000000000772012101774706012243 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- IO ------------------------------------------------------------------------ module IO where open import Coinduction open import Data.Unit open import Data.String open import Data.Colist open import Function import IO.Primitive as Prim open import Level ------------------------------------------------------------------------ -- The IO monad -- One cannot write "infinitely large" computations with the -- postulated IO monad in IO.Primitive without turning off the -- termination checker (or going via the FFI, or perhaps abusing -- something else). The following coinductive deep embedding is -- introduced to avoid this problem. Possible non-termination is -- isolated to the run function below. infixl 1 _>>=_ _>>_ data IO {a} (A : Set a) : Set (suc a) where lift : (m : Prim.IO A) → IO A return : (x : A) → IO A _>>=_ : {B : Set a} (m : ∞ (IO B)) (f : (x : B) → ∞ (IO A)) → IO A _>>_ : {B : Set a} (m₁ : ∞ (IO B)) (m₂ : ∞ (IO A)) → IO A -- The use of abstract ensures that the run function will not be -- unfolded infinitely by the type checker. abstract {-# NO_TERMINATION_CHECK #-} run : ∀ {a} {A : Set a} → IO A → Prim.IO A run (lift m) = m run (return x) = Prim.return x run (m >>= f) = Prim._>>=_ (run (♭ m )) λ x → run (♭ (f x)) run (m₁ >> m₂) = Prim._>>=_ (run (♭ m₁)) λ _ → run (♭ m₂) ------------------------------------------------------------------------ -- Utilities sequence : ∀ {a} {A : Set a} → Colist (IO A) → IO (Colist A) sequence [] = return [] sequence (c ∷ cs) = ♯ c >>= λ x → ♯ (♯ sequence (♭ cs) >>= λ xs → ♯ return (x ∷ ♯ xs)) -- The reason for not defining sequence′ in terms of sequence is -- efficiency (the unused results could cause unnecessary memory use). sequence′ : ∀ {a} {A : Set a} → Colist (IO A) → IO (Lift ⊤) sequence′ [] = return _ sequence′ (c ∷ cs) = ♯ c >> ♯ (♯ sequence′ (♭ cs) >> ♯ return _) mapM : ∀ {a b} {A : Set a} {B : Set b} → (A → IO B) → Colist A → IO (Colist B) mapM f = sequence ∘ map f mapM′ : {A B : Set} → (A → IO B) → Colist A → IO (Lift ⊤) mapM′ f = sequence′ ∘ map f ------------------------------------------------------------------------ -- Simple lazy IO -- Note that the functions below produce commands which, when -- executed, may raise exceptions. -- Note also that the semantics of these functions depends on the -- version of the Haskell base library. If the version is 4.2.0.0 (or -- later?), then the functions use the character encoding specified by -- the locale. For older versions of the library (going back to at -- least version 3) the functions use ISO-8859-1. getContents : IO Costring getContents = lift Prim.getContents readFile : String → IO Costring readFile f = lift (Prim.readFile f) -- Reads a finite file. Raises an exception if the file path refers to -- a non-physical file (like "/dev/zero"). readFiniteFile : String → IO String readFiniteFile f = lift (Prim.readFiniteFile f) writeFile∞ : String → Costring → IO ⊤ writeFile∞ f s = ♯ lift (Prim.writeFile f s) >> ♯ return _ writeFile : String → String → IO ⊤ writeFile f s = writeFile∞ f (toCostring s) appendFile∞ : String → Costring → IO ⊤ appendFile∞ f s = ♯ lift (Prim.appendFile f s) >> ♯ return _ appendFile : String → String → IO ⊤ appendFile f s = appendFile∞ f (toCostring s) putStr∞ : Costring → IO ⊤ putStr∞ s = ♯ lift (Prim.putStr s) >> ♯ return _ putStr : String → IO ⊤ putStr s = putStr∞ (toCostring s) putStrLn∞ : Costring → IO ⊤ putStrLn∞ s = ♯ lift (Prim.putStrLn s) >> ♯ return _ putStrLn : String → IO ⊤ putStrLn s = putStrLn∞ (toCostring s) lib-0.7/src/Algebra.agda0000644000000000000000000003254612101774706013275 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Definitions of algebraic structures like monoids and rings -- (packed in records together with sets, operations, etc.) ------------------------------------------------------------------------ module Algebra where open import Relation.Binary open import Algebra.FunctionProperties open import Algebra.Structures open import Function open import Level ------------------------------------------------------------------------ -- Semigroups, (commutative) monoids and (abelian) groups record Semigroup c ℓ : Set (suc (c ⊔ ℓ)) where infixl 7 _∙_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _∙_ : Op₂ Carrier isSemigroup : IsSemigroup _≈_ _∙_ open IsSemigroup isSemigroup public setoid : Setoid _ _ setoid = record { isEquivalence = isEquivalence } -- A raw monoid is a monoid without any laws. record RawMonoid c ℓ : Set (suc (c ⊔ ℓ)) where infixl 7 _∙_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _∙_ : Op₂ Carrier ε : Carrier record Monoid c ℓ : Set (suc (c ⊔ ℓ)) where infixl 7 _∙_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _∙_ : Op₂ Carrier ε : Carrier isMonoid : IsMonoid _≈_ _∙_ ε open IsMonoid isMonoid public semigroup : Semigroup _ _ semigroup = record { isSemigroup = isSemigroup } open Semigroup semigroup public using (setoid) rawMonoid : RawMonoid _ _ rawMonoid = record { _≈_ = _≈_ ; _∙_ = _∙_ ; ε = ε } record CommutativeMonoid c ℓ : Set (suc (c ⊔ ℓ)) where infixl 7 _∙_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _∙_ : Op₂ Carrier ε : Carrier isCommutativeMonoid : IsCommutativeMonoid _≈_ _∙_ ε open IsCommutativeMonoid isCommutativeMonoid public monoid : Monoid _ _ monoid = record { isMonoid = isMonoid } open Monoid monoid public using (setoid; semigroup; rawMonoid) record Group c ℓ : Set (suc (c ⊔ ℓ)) where infix 8 _⁻¹ infixl 7 _∙_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _∙_ : Op₂ Carrier ε : Carrier _⁻¹ : Op₁ Carrier isGroup : IsGroup _≈_ _∙_ ε _⁻¹ open IsGroup isGroup public monoid : Monoid _ _ monoid = record { isMonoid = isMonoid } open Monoid monoid public using (setoid; semigroup; rawMonoid) record AbelianGroup c ℓ : Set (suc (c ⊔ ℓ)) where infix 8 _⁻¹ infixl 7 _∙_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _∙_ : Op₂ Carrier ε : Carrier _⁻¹ : Op₁ Carrier isAbelianGroup : IsAbelianGroup _≈_ _∙_ ε _⁻¹ open IsAbelianGroup isAbelianGroup public group : Group _ _ group = record { isGroup = isGroup } open Group group public using (setoid; semigroup; monoid; rawMonoid) commutativeMonoid : CommutativeMonoid _ _ commutativeMonoid = record { isCommutativeMonoid = isCommutativeMonoid } ------------------------------------------------------------------------ -- Various kinds of semirings record NearSemiring c ℓ : Set (suc (c ⊔ ℓ)) where infixl 7 _*_ infixl 6 _+_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _+_ : Op₂ Carrier _*_ : Op₂ Carrier 0# : Carrier isNearSemiring : IsNearSemiring _≈_ _+_ _*_ 0# open IsNearSemiring isNearSemiring public +-monoid : Monoid _ _ +-monoid = record { isMonoid = +-isMonoid } open Monoid +-monoid public using (setoid) renaming ( semigroup to +-semigroup ; rawMonoid to +-rawMonoid) *-semigroup : Semigroup _ _ *-semigroup = record { isSemigroup = *-isSemigroup } record SemiringWithoutOne c ℓ : Set (suc (c ⊔ ℓ)) where infixl 7 _*_ infixl 6 _+_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _+_ : Op₂ Carrier _*_ : Op₂ Carrier 0# : Carrier isSemiringWithoutOne : IsSemiringWithoutOne _≈_ _+_ _*_ 0# open IsSemiringWithoutOne isSemiringWithoutOne public nearSemiring : NearSemiring _ _ nearSemiring = record { isNearSemiring = isNearSemiring } open NearSemiring nearSemiring public using ( setoid ; +-semigroup; +-rawMonoid; +-monoid ; *-semigroup ) +-commutativeMonoid : CommutativeMonoid _ _ +-commutativeMonoid = record { isCommutativeMonoid = +-isCommutativeMonoid } record SemiringWithoutAnnihilatingZero c ℓ : Set (suc (c ⊔ ℓ)) where infixl 7 _*_ infixl 6 _+_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _+_ : Op₂ Carrier _*_ : Op₂ Carrier 0# : Carrier 1# : Carrier isSemiringWithoutAnnihilatingZero : IsSemiringWithoutAnnihilatingZero _≈_ _+_ _*_ 0# 1# open IsSemiringWithoutAnnihilatingZero isSemiringWithoutAnnihilatingZero public +-commutativeMonoid : CommutativeMonoid _ _ +-commutativeMonoid = record { isCommutativeMonoid = +-isCommutativeMonoid } open CommutativeMonoid +-commutativeMonoid public using (setoid) renaming ( semigroup to +-semigroup ; rawMonoid to +-rawMonoid ; monoid to +-monoid ) *-monoid : Monoid _ _ *-monoid = record { isMonoid = *-isMonoid } open Monoid *-monoid public using () renaming ( semigroup to *-semigroup ; rawMonoid to *-rawMonoid ) record Semiring c ℓ : Set (suc (c ⊔ ℓ)) where infixl 7 _*_ infixl 6 _+_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _+_ : Op₂ Carrier _*_ : Op₂ Carrier 0# : Carrier 1# : Carrier isSemiring : IsSemiring _≈_ _+_ _*_ 0# 1# open IsSemiring isSemiring public semiringWithoutAnnihilatingZero : SemiringWithoutAnnihilatingZero _ _ semiringWithoutAnnihilatingZero = record { isSemiringWithoutAnnihilatingZero = isSemiringWithoutAnnihilatingZero } open SemiringWithoutAnnihilatingZero semiringWithoutAnnihilatingZero public using ( setoid ; +-semigroup; +-rawMonoid; +-monoid ; +-commutativeMonoid ; *-semigroup; *-rawMonoid; *-monoid ) semiringWithoutOne : SemiringWithoutOne _ _ semiringWithoutOne = record { isSemiringWithoutOne = isSemiringWithoutOne } open SemiringWithoutOne semiringWithoutOne public using (nearSemiring) record CommutativeSemiringWithoutOne c ℓ : Set (suc (c ⊔ ℓ)) where infixl 7 _*_ infixl 6 _+_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _+_ : Op₂ Carrier _*_ : Op₂ Carrier 0# : Carrier isCommutativeSemiringWithoutOne : IsCommutativeSemiringWithoutOne _≈_ _+_ _*_ 0# open IsCommutativeSemiringWithoutOne isCommutativeSemiringWithoutOne public semiringWithoutOne : SemiringWithoutOne _ _ semiringWithoutOne = record { isSemiringWithoutOne = isSemiringWithoutOne } open SemiringWithoutOne semiringWithoutOne public using ( setoid ; +-semigroup; +-rawMonoid; +-monoid ; +-commutativeMonoid ; *-semigroup ; nearSemiring ) record CommutativeSemiring c ℓ : Set (suc (c ⊔ ℓ)) where infixl 7 _*_ infixl 6 _+_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _+_ : Op₂ Carrier _*_ : Op₂ Carrier 0# : Carrier 1# : Carrier isCommutativeSemiring : IsCommutativeSemiring _≈_ _+_ _*_ 0# 1# open IsCommutativeSemiring isCommutativeSemiring public semiring : Semiring _ _ semiring = record { isSemiring = isSemiring } open Semiring semiring public using ( setoid ; +-semigroup; +-rawMonoid; +-monoid ; +-commutativeMonoid ; *-semigroup; *-rawMonoid; *-monoid ; nearSemiring; semiringWithoutOne ; semiringWithoutAnnihilatingZero ) *-commutativeMonoid : CommutativeMonoid _ _ *-commutativeMonoid = record { isCommutativeMonoid = *-isCommutativeMonoid } commutativeSemiringWithoutOne : CommutativeSemiringWithoutOne _ _ commutativeSemiringWithoutOne = record { isCommutativeSemiringWithoutOne = isCommutativeSemiringWithoutOne } ------------------------------------------------------------------------ -- (Commutative) rings -- A raw ring is a ring without any laws. record RawRing c : Set (suc c) where infix 8 -_ infixl 7 _*_ infixl 6 _+_ field Carrier : Set c _+_ : Op₂ Carrier _*_ : Op₂ Carrier -_ : Op₁ Carrier 0# : Carrier 1# : Carrier record Ring c ℓ : Set (suc (c ⊔ ℓ)) where infix 8 -_ infixl 7 _*_ infixl 6 _+_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _+_ : Op₂ Carrier _*_ : Op₂ Carrier -_ : Op₁ Carrier 0# : Carrier 1# : Carrier isRing : IsRing _≈_ _+_ _*_ -_ 0# 1# open IsRing isRing public +-abelianGroup : AbelianGroup _ _ +-abelianGroup = record { isAbelianGroup = +-isAbelianGroup } semiring : Semiring _ _ semiring = record { isSemiring = isSemiring } open Semiring semiring public using ( setoid ; +-semigroup; +-rawMonoid; +-monoid ; +-commutativeMonoid ; *-semigroup; *-rawMonoid; *-monoid ; nearSemiring; semiringWithoutOne ; semiringWithoutAnnihilatingZero ) open AbelianGroup +-abelianGroup public using () renaming (group to +-group) rawRing : RawRing _ rawRing = record { _+_ = _+_ ; _*_ = _*_ ; -_ = -_ ; 0# = 0# ; 1# = 1# } record CommutativeRing c ℓ : Set (suc (c ⊔ ℓ)) where infix 8 -_ infixl 7 _*_ infixl 6 _+_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _+_ : Op₂ Carrier _*_ : Op₂ Carrier -_ : Op₁ Carrier 0# : Carrier 1# : Carrier isCommutativeRing : IsCommutativeRing _≈_ _+_ _*_ -_ 0# 1# open IsCommutativeRing isCommutativeRing public ring : Ring _ _ ring = record { isRing = isRing } commutativeSemiring : CommutativeSemiring _ _ commutativeSemiring = record { isCommutativeSemiring = isCommutativeSemiring } open Ring ring public using (rawRing; +-group; +-abelianGroup) open CommutativeSemiring commutativeSemiring public using ( setoid ; +-semigroup; +-rawMonoid; +-monoid; +-commutativeMonoid ; *-semigroup; *-rawMonoid; *-monoid; *-commutativeMonoid ; nearSemiring; semiringWithoutOne ; semiringWithoutAnnihilatingZero; semiring ; commutativeSemiringWithoutOne ) ------------------------------------------------------------------------ -- (Distributive) lattices and boolean algebras record Lattice c ℓ : Set (suc (c ⊔ ℓ)) where infixr 7 _∧_ infixr 6 _∨_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _∨_ : Op₂ Carrier _∧_ : Op₂ Carrier isLattice : IsLattice _≈_ _∨_ _∧_ open IsLattice isLattice public setoid : Setoid _ _ setoid = record { isEquivalence = isEquivalence } record DistributiveLattice c ℓ : Set (suc (c ⊔ ℓ)) where infixr 7 _∧_ infixr 6 _∨_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _∨_ : Op₂ Carrier _∧_ : Op₂ Carrier isDistributiveLattice : IsDistributiveLattice _≈_ _∨_ _∧_ open IsDistributiveLattice isDistributiveLattice public lattice : Lattice _ _ lattice = record { isLattice = isLattice } open Lattice lattice public using (setoid) record BooleanAlgebra c ℓ : Set (suc (c ⊔ ℓ)) where infix 8 ¬_ infixr 7 _∧_ infixr 6 _∨_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _∨_ : Op₂ Carrier _∧_ : Op₂ Carrier ¬_ : Op₁ Carrier ⊤ : Carrier ⊥ : Carrier isBooleanAlgebra : IsBooleanAlgebra _≈_ _∨_ _∧_ ¬_ ⊤ ⊥ open IsBooleanAlgebra isBooleanAlgebra public distributiveLattice : DistributiveLattice _ _ distributiveLattice = record { isDistributiveLattice = isDistributiveLattice } open DistributiveLattice distributiveLattice public using (setoid; lattice) lib-0.7/src/Level.agda0000644000000000000000000000153612101774706013002 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Universe levels ------------------------------------------------------------------------ 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 _⊔_ #-} -- Lifting. record Lift {a ℓ} (A : Set a) : Set (a ⊔ ℓ) where constructor lift field lower : A open Lift public lib-0.7/src/Reflection.agda0000644000000000000000000003010112101774706014013 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Support for reflection ------------------------------------------------------------------------ module Reflection where open import Data.Bool as Bool using (Bool); open Bool.Bool open import Data.List using (List); open Data.List.List open import Data.Nat using (ℕ) renaming (_≟_ to _≟-ℕ_) open import Data.Product open import Function open import Relation.Binary open import Relation.Binary.PropositionalEquality open import Relation.Binary.PropositionalEquality.TrustMe open import Relation.Nullary open import Relation.Nullary.Decidable as Dec open import Relation.Nullary.Product ------------------------------------------------------------------------ -- Names -- Names. postulate Name : Set {-# BUILTIN QNAME Name #-} private primitive primQNameEquality : Name → Name → Bool -- Equality of names is decidable. infix 4 _==_ _≟-Name_ private _==_ : Name → Name → Bool _==_ = primQNameEquality _≟-Name_ : Decidable {A = Name} _≡_ s₁ ≟-Name s₂ with s₁ == s₂ ... | true = yes trustMe ... | false = no whatever where postulate whatever : _ ------------------------------------------------------------------------ -- Terms -- Is the argument visible (explicit), hidden (implicit), or an -- instance argument? data Visibility : Set where visible hidden instance : Visibility {-# BUILTIN HIDING Visibility #-} {-# BUILTIN VISIBLE visible #-} {-# BUILTIN HIDDEN hidden #-} {-# BUILTIN INSTANCE instance #-} -- Arguments can be relevant or irrelevant. data Relevance : Set where relevant irrelevant : Relevance {-# BUILTIN RELEVANCE Relevance #-} {-# BUILTIN RELEVANT relevant #-} {-# BUILTIN IRRELEVANT irrelevant #-} -- Arguments. data Arg A : Set where arg : (v : Visibility) (r : Relevance) (x : A) → Arg A {-# BUILTIN ARG Arg #-} {-# BUILTIN ARGARG arg #-} -- Terms. mutual data Term : Set where -- Variable applied to arguments. var : (x : ℕ) (args : List (Arg Term)) → Term -- Constructor applied to arguments. con : (c : Name) (args : List (Arg Term)) → Term -- Identifier applied to arguments. def : (f : Name) (args : List (Arg Term)) → Term -- Different kinds of λ-abstraction. lam : (v : Visibility) (t : Term) → Term -- Pi-type. pi : (t₁ : Arg Type) (t₂ : Type) → Term -- A sort. sort : Sort → Term -- Anything else. unknown : Term data Type : Set where el : (s : Sort) (t : Term) → Type data Sort : Set where -- A Set of a given (possibly neutral) level. set : (t : Term) → Sort -- A Set of a given concrete level. lit : (n : ℕ) → Sort -- Anything else. unknown : Sort {-# BUILTIN AGDASORT Sort #-} {-# BUILTIN AGDATYPE Type #-} {-# BUILTIN AGDATERM Term #-} {-# BUILTIN AGDATERMVAR var #-} {-# BUILTIN AGDATERMCON con #-} {-# BUILTIN AGDATERMDEF def #-} {-# BUILTIN AGDATERMLAM lam #-} {-# BUILTIN AGDATERMPI pi #-} {-# BUILTIN AGDATERMSORT sort #-} {-# BUILTIN AGDATERMUNSUPPORTED unknown #-} {-# BUILTIN AGDATYPEEL el #-} {-# BUILTIN AGDASORTSET set #-} {-# BUILTIN AGDASORTLIT lit #-} {-# BUILTIN AGDASORTUNSUPPORTED unknown #-} ------------------------------------------------------------------------ -- Definitions postulate -- Function definition. Function : Set -- Data type definition. Data-type : Set -- Record type definition. Record : Set {-# BUILTIN AGDAFUNDEF Function #-} {-# BUILTIN AGDADATADEF Data-type #-} {-# BUILTIN AGDARECORDDEF Record #-} -- Definitions. data Definition : Set where function : Function → Definition data-type : Data-type → Definition record′ : Record → Definition constructor′ : Definition axiom : Definition primitive′ : Definition {-# BUILTIN AGDADEFINITION Definition #-} {-# BUILTIN AGDADEFINITIONFUNDEF function #-} {-# BUILTIN AGDADEFINITIONDATADEF data-type #-} {-# BUILTIN AGDADEFINITIONRECORDDEF record′ #-} {-# BUILTIN AGDADEFINITIONDATACONSTRUCTOR constructor′ #-} {-# BUILTIN AGDADEFINITIONPOSTULATE axiom #-} {-# BUILTIN AGDADEFINITIONPRIMITIVE primitive′ #-} private primitive primQNameType : Name → Type primQNameDefinition : Name → Definition primDataConstructors : Data-type → List Name -- The type of the thing with the given name. type : Name → Type type = primQNameType -- The definition of the thing with the given name. definition : Name → Definition definition = primQNameDefinition -- The constructors of the given data type. constructors : Data-type → List Name constructors = primDataConstructors ------------------------------------------------------------------------ -- Term equality is decidable -- Boring helper functions. private cong₂′ : ∀ {A B C : Set} (f : A → B → C) {x y u v} → x ≡ y × u ≡ v → f x u ≡ f y v cong₂′ f = uncurry (cong₂ f) cong₃′ : ∀ {A B C D : Set} (f : A → B → C → D) {x y u v r s} → x ≡ y × u ≡ v × r ≡ s → f x u r ≡ f y v s cong₃′ f (refl , refl , refl) = refl arg₁ : ∀ {A v v′ r r′} {x x′ : A} → arg v r x ≡ arg v′ r′ x′ → v ≡ v′ arg₁ refl = refl arg₂ : ∀ {A v v′ r r′} {x x′ : A} → arg v r x ≡ arg v′ r′ x′ → r ≡ r′ arg₂ refl = refl arg₃ : ∀ {A v v′ r r′} {x x′ : A} → arg v r x ≡ arg v′ r′ x′ → x ≡ x′ arg₃ refl = refl cons₁ : ∀ {A : Set} {x y} {xs ys : List A} → x ∷ xs ≡ y ∷ ys → x ≡ y cons₁ refl = refl cons₂ : ∀ {A : Set} {x y} {xs ys : List A} → x ∷ xs ≡ y ∷ ys → xs ≡ ys cons₂ refl = refl var₁ : ∀ {x x′ args args′} → var x args ≡ var x′ args′ → x ≡ x′ var₁ refl = refl var₂ : ∀ {x x′ args args′} → var x args ≡ var x′ args′ → args ≡ args′ var₂ refl = refl con₁ : ∀ {c c′ args args′} → con c args ≡ con c′ args′ → c ≡ c′ con₁ refl = refl con₂ : ∀ {c c′ args args′} → con c args ≡ con c′ args′ → args ≡ args′ con₂ refl = refl def₁ : ∀ {f f′ args args′} → def f args ≡ def f′ args′ → f ≡ f′ def₁ refl = refl def₂ : ∀ {f f′ args args′} → def f args ≡ def f′ args′ → args ≡ args′ def₂ refl = refl lam₁ : ∀ {v v′ t t′} → lam v t ≡ lam v′ t′ → v ≡ v′ lam₁ refl = refl lam₂ : ∀ {v v′ t t′} → lam v t ≡ lam v′ t′ → t ≡ t′ lam₂ refl = refl pi₁ : ∀ {t₁ t₁′ t₂ t₂′} → pi t₁ t₂ ≡ pi t₁′ t₂′ → t₁ ≡ t₁′ pi₁ refl = refl pi₂ : ∀ {t₁ t₁′ t₂ t₂′} → pi t₁ t₂ ≡ pi t₁′ t₂′ → t₂ ≡ t₂′ pi₂ refl = refl sort₁ : ∀ {x y} → sort x ≡ sort y → x ≡ y sort₁ refl = refl set₁ : ∀ {x y} → set x ≡ set y → x ≡ y set₁ refl = refl lit₁ : ∀ {x y} → lit x ≡ lit y → x ≡ y lit₁ refl = refl el₁ : ∀ {s s′ t t′} → el s t ≡ el s′ t′ → s ≡ s′ el₁ refl = refl el₂ : ∀ {s s′ t t′} → el s t ≡ el s′ t′ → t ≡ t′ el₂ refl = refl _≟-Visibility_ : Decidable (_≡_ {A = Visibility}) visible ≟-Visibility visible = yes refl hidden ≟-Visibility hidden = yes refl instance ≟-Visibility instance = yes refl visible ≟-Visibility hidden = no λ() visible ≟-Visibility instance = no λ() hidden ≟-Visibility visible = no λ() hidden ≟-Visibility instance = no λ() instance ≟-Visibility visible = no λ() instance ≟-Visibility hidden = no λ() _≟-Relevance_ : Decidable (_≡_ {A = Relevance}) relevant ≟-Relevance relevant = yes refl irrelevant ≟-Relevance irrelevant = yes refl relevant ≟-Relevance irrelevant = no λ() irrelevant ≟-Relevance relevant = no λ() mutual infix 4 _≟_ _≟-Args_ _≟-ArgType_ _≟-ArgTerm_ : Decidable (_≡_ {A = Arg Term}) arg e r a ≟-ArgTerm arg e′ r′ a′ = Dec.map′ (cong₃′ arg) < arg₁ , < arg₂ , arg₃ > > (e ≟-Visibility e′ ×-dec r ≟-Relevance r′ ×-dec a ≟ a′) _≟-ArgType_ : Decidable (_≡_ {A = Arg Type}) arg e r a ≟-ArgType arg e′ r′ a′ = Dec.map′ (cong₃′ arg) < arg₁ , < arg₂ , arg₃ > > (e ≟-Visibility e′ ×-dec r ≟-Relevance r′ ×-dec a ≟-Type a′) _≟-Args_ : Decidable (_≡_ {A = List (Arg Term)}) [] ≟-Args [] = yes refl (x ∷ xs) ≟-Args (y ∷ ys) = Dec.map′ (cong₂′ _∷_) < cons₁ , cons₂ > (x ≟-ArgTerm y ×-dec xs ≟-Args ys) [] ≟-Args (_ ∷ _) = no λ() (_ ∷ _) ≟-Args [] = no λ() _≟_ : Decidable (_≡_ {A = Term}) var x args ≟ var x′ args′ = Dec.map′ (cong₂′ var) < var₁ , var₂ > (x ≟-ℕ x′ ×-dec args ≟-Args args′) con c args ≟ con c′ args′ = Dec.map′ (cong₂′ con) < con₁ , con₂ > (c ≟-Name c′ ×-dec args ≟-Args args′) def f args ≟ def f′ args′ = Dec.map′ (cong₂′ def) < def₁ , def₂ > (f ≟-Name f′ ×-dec args ≟-Args args′) lam v t ≟ lam v′ t′ = Dec.map′ (cong₂′ lam) < lam₁ , lam₂ > (v ≟-Visibility v′ ×-dec t ≟ t′) pi t₁ t₂ ≟ pi t₁′ t₂′ = Dec.map′ (cong₂′ pi) < pi₁ , pi₂ > (t₁ ≟-ArgType t₁′ ×-dec t₂ ≟-Type t₂′) sort s ≟ sort s′ = Dec.map′ (cong sort) sort₁ (s ≟-Sort s′) unknown ≟ unknown = yes refl var x args ≟ con c args′ = no λ() var x args ≟ def f args′ = no λ() var x args ≟ lam v t = no λ() var x args ≟ pi t₁ t₂ = no λ() var x args ≟ sort _ = no λ() var x args ≟ unknown = no λ() con c args ≟ var x args′ = no λ() con c args ≟ def f args′ = no λ() con c args ≟ lam v t = no λ() con c args ≟ pi t₁ t₂ = no λ() con c args ≟ sort _ = no λ() con c args ≟ unknown = no λ() def f args ≟ var x args′ = no λ() def f args ≟ con c args′ = no λ() def f args ≟ lam v t = no λ() def f args ≟ pi t₁ t₂ = no λ() def f args ≟ sort _ = no λ() def f args ≟ unknown = no λ() lam v t ≟ var x args = no λ() lam v t ≟ con c args = no λ() lam v t ≟ def f args = no λ() lam v t ≟ pi t₁ t₂ = no λ() lam v t ≟ sort _ = no λ() lam v t ≟ unknown = no λ() pi t₁ t₂ ≟ var x args = no λ() pi t₁ t₂ ≟ con c args = no λ() pi t₁ t₂ ≟ def f args = no λ() pi t₁ t₂ ≟ lam v t = no λ() pi t₁ t₂ ≟ sort _ = no λ() pi t₁ t₂ ≟ unknown = no λ() sort _ ≟ var x args = no λ() sort _ ≟ con c args = no λ() sort _ ≟ def f args = no λ() sort _ ≟ lam v t = no λ() sort _ ≟ pi t₁ t₂ = no λ() sort _ ≟ unknown = no λ() unknown ≟ var x args = no λ() unknown ≟ con c args = no λ() unknown ≟ def f args = no λ() unknown ≟ lam v t = no λ() unknown ≟ pi t₁ t₂ = no λ() unknown ≟ sort _ = no λ() _≟-Type_ : Decidable (_≡_ {A = Type}) el s t ≟-Type el s′ t′ = Dec.map′ (cong₂′ el) < el₁ , el₂ > (s ≟-Sort s′ ×-dec t ≟ t′) _≟-Sort_ : Decidable (_≡_ {A = Sort}) set t ≟-Sort set t′ = Dec.map′ (cong set) set₁ (t ≟ t′) lit n ≟-Sort lit n′ = Dec.map′ (cong lit) lit₁ (n ≟-ℕ n′) unknown ≟-Sort unknown = yes refl set _ ≟-Sort lit _ = no λ() set _ ≟-Sort unknown = no λ() lit _ ≟-Sort set _ = no λ() lit _ ≟-Sort unknown = no λ() unknown ≟-Sort set _ = no λ() unknown ≟-Sort lit _ = no λ() lib-0.7/src/Function.agda0000644000000000000000000000605312101774706013517 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Simple combinators working solely on and with functions ------------------------------------------------------------------------ module Function where open import Level infixr 9 _∘_ _∘′_ infixl 8 _ˢ_ infixl 1 _on_ infixl 1 _⟨_⟩_ infixr 0 _-[_]-_ _$_ infixl 0 _∋_ ------------------------------------------------------------------------ -- Types -- Unary functions on a given set. Fun₁ : ∀ {i} → Set i → Set i Fun₁ A = A → A -- Binary functions on a given set. Fun₂ : ∀ {i} → Set i → Set i Fun₂ A = A → A → A ------------------------------------------------------------------------ -- Functions _∘_ : ∀ {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)) f ∘ g = λ x → f (g x) _∘′_ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → (B → C) → (A → B) → (A → C) f ∘′ g = _∘_ f g id : ∀ {a} {A : Set a} → A → A id x = x const : ∀ {a b} {A : Set a} {B : Set b} → A → B → A const x = λ _ → x -- The S combinator. (Written infix as in Conor McBride's paper -- "Outrageous but Meaningful Coincidences: Dependent type-safe syntax -- and evaluation".) _ˢ_ : ∀ {a b c} {A : Set a} {B : A → Set b} {C : (x : A) → B x → Set c} → ((x : A) (y : B x) → C x y) → (g : (x : A) → B x) → ((x : A) → C x (g x)) f ˢ g = λ x → f x (g x) flip : ∀ {a b c} {A : Set a} {B : Set b} {C : A → B → Set c} → ((x : A) (y : B) → C x y) → ((y : B) (x : A) → C x y) flip f = λ y x → f x y -- Note that _$_ is right associative, like in Haskell. If you want a -- left associative infix application operator, use -- Category.Functor._<$>_, available from -- Category.Monad.Identity.IdentityMonad. _$_ : ∀ {a b} {A : Set a} {B : A → Set b} → ((x : A) → B x) → ((x : A) → B x) f $ x = f x _⟨_⟩_ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → A → (A → B → C) → B → C x ⟨ f ⟩ y = f x y _on_ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → (B → B → C) → (A → B) → (A → A → C) _*_ on f = λ x y → f x * f y _-[_]-_ : ∀ {a b c d e} {A : Set a} {B : Set b} {C : Set c} {D : Set d} {E : Set e} → (A → B → C) → (C → D → E) → (A → B → D) → (A → B → E) f -[ _*_ ]- g = λ x y → f x y * g x y -- In Agda you cannot annotate every subexpression with a type -- signature. This function can be used instead. _∋_ : ∀ {a} (A : Set a) → A → A A ∋ x = x -- Case expressions (to be used with pattern-matching lambdas, see -- README.Case). infix 0 case_return_of_ case_of_ case_return_of_ : ∀ {a b} {A : Set a} (x : A) (B : A → Set b) → ((x : A) → B x) → B x case x return B of f = f x case_of_ : ∀ {a b} {A : Set a} {B : Set b} → A → (A → B) → B case x of f = case x return _ of f lib-0.7/src/Induction.agda0000644000000000000000000000404312101774706013663 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- An abstraction of various forms of recursion/induction ------------------------------------------------------------------------ -- The idea underlying Induction.* comes from Epigram 1, see Section 4 -- of "The view from the left" by McBride and McKinna. -- Note: The types in this module can perhaps be easier to understand -- if they are normalised. Note also that Agda can do the -- normalisation for you. module Induction where open import Level 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 -- A recursor builder constructs an instance of a recursion structure -- for a given input. 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 Rec = ∀ P → Rec P ⊆′ P → Universal P -- And recursors can be constructed from recursor builders. build : ∀ {a} {A : Set a} {Rec : RecStruct A} → RecursorBuilder Rec → Recursor Rec 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 Q Rec = ∀ P → Rec P ⊆′ P → Q ⊆′ Rec P 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} → SubsetRecursorBuilder Q Rec → SubsetRecursor Q Rec subsetBuild builder P f x q = f x (builder P f x q) lib-0.7/src/Size.agda0000644000000000000000000000060212101774706012636 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Sizes for Agda's sized types ------------------------------------------------------------------------ module Size where postulate Size : Set ↑_ : Size → Size ∞ : Size {-# BUILTIN SIZE Size #-} {-# BUILTIN SIZESUC ↑_ #-} {-# BUILTIN SIZEINF ∞ #-} lib-0.7/src/Record.agda0000644000000000000000000002035112101774706013145 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Record types with manifest fields and "with", based on Randy -- Pollack's "Dependently Typed Records in Type Theory" ------------------------------------------------------------------------ -- For an example of how this module can be used, see README.Record. -- The module is parametrised by the type of labels, which should come -- with decidable equality. open import Relation.Binary open import Relation.Binary.PropositionalEquality module Record (Label : Set) (_≟_ : Decidable (_≡_ {A = Label})) where open import Data.Bool open import Data.Empty open import Data.List open import Data.Product hiding (proj₁; proj₂) open import Data.Unit open import Function open import Level open import Relation.Nullary open import Relation.Nullary.Decidable ------------------------------------------------------------------------ -- A Σ-type with a manifest field -- A variant of Σ where the value of the second field is "manifest" -- (given by the first). infix 4 _, record Manifest-Σ {a b} (A : Set a) {B : A → Set b} (f : (x : A) → B x) : Set a where constructor _, field proj₁ : A proj₂ : B proj₁ proj₂ = f proj₁ ------------------------------------------------------------------------ -- Signatures and records mutual infixl 5 _,_∶_ _,_≔_ data Signature s : Set (suc s) where ∅ : Signature s _,_∶_ : (Sig : Signature s) (ℓ : Label) (A : Record Sig → Set s) → Signature s _,_≔_ : (Sig : Signature s) (ℓ : Label) {A : Record Sig → Set s} (a : (r : Record Sig) → A r) → Signature s -- Record is a record type to ensure that the signature can be -- inferred from a value of type Record Sig. record Record {s} (Sig : Signature s) : Set s where constructor rec field fun : Record-fun Sig Record-fun : ∀ {s} → Signature s → Set s Record-fun ∅ = Lift ⊤ Record-fun (Sig , ℓ ∶ A) = Σ (Record Sig) A Record-fun (Sig , ℓ ≔ a) = Manifest-Σ (Record Sig) a ------------------------------------------------------------------------ -- Variants of Signature and Record -- It may be easier to define values of type Record′ (Sig⇒Sig′ Sig) -- than of type Record Sig. mutual data Signature′ s : Set (suc s) where ∅ : Signature′ s _,_∶_ : (Sig : Signature′ s) (ℓ : Label) (A : Record′ Sig → Set s) → Signature′ s _,_≔_ : (Sig : Signature′ s) (ℓ : Label) {A : Record′ Sig → Set s} (a : (r : Record′ Sig) → A r) → Signature′ s Record′ : ∀ {s} → Signature′ s → Set s Record′ ∅ = Lift ⊤ Record′ (Sig , ℓ ∶ A) = Σ (Record′ Sig) A Record′ (Sig , ℓ ≔ a) = Manifest-Σ (Record′ Sig) a -- We can convert between the two kinds of signatures/records. mutual Sig′⇒Sig : ∀ {s} → Signature′ s → Signature s Sig′⇒Sig ∅ = ∅ Sig′⇒Sig (Sig , ℓ ∶ A) = Sig′⇒Sig Sig , ℓ ∶ (A ∘ Rec⇒Rec′ _) Sig′⇒Sig (Sig , ℓ ≔ a) = Sig′⇒Sig Sig , ℓ ≔ (a ∘ Rec⇒Rec′ _) Rec⇒Rec′ : ∀ {s} (Sig : Signature′ s) → Record (Sig′⇒Sig Sig) → Record′ Sig Rec⇒Rec′ ∅ (rec r) = r Rec⇒Rec′ (Sig , ℓ ∶ A) (rec r) = (Rec⇒Rec′ _ (Σ.proj₁ r) , Σ.proj₂ r) Rec⇒Rec′ (Sig , ℓ ≔ a) (rec r) = (Rec⇒Rec′ _ (Manifest-Σ.proj₁ r) ,) mutual Sig⇒Sig′ : ∀ {s} → Signature s → Signature′ s Sig⇒Sig′ ∅ = ∅ Sig⇒Sig′ (Sig , ℓ ∶ A) = Sig⇒Sig′ Sig , ℓ ∶ (A ∘ Rec′⇒Rec _) Sig⇒Sig′ (Sig , ℓ ≔ a) = Sig⇒Sig′ Sig , ℓ ≔ (a ∘ Rec′⇒Rec _) Rec′⇒Rec : ∀ {s} (Sig : Signature s) → Record′ (Sig⇒Sig′ Sig) → Record Sig Rec′⇒Rec ∅ r = rec r Rec′⇒Rec (Sig , ℓ ∶ A) r = rec (Rec′⇒Rec _ (Σ.proj₁ r) , Σ.proj₂ r) Rec′⇒Rec (Sig , ℓ ≔ a) r = rec (Rec′⇒Rec _ (Manifest-Σ.proj₁ r) ,) ------------------------------------------------------------------------ -- Labels -- A signature's labels, starting with the last one. labels : ∀ {s} → Signature s → List Label labels ∅ = [] labels (Sig , ℓ ∶ A) = ℓ ∷ labels Sig labels (Sig , ℓ ≔ a) = ℓ ∷ labels Sig -- Inhabited if the label is part of the signature. infix 4 _∈_ _∈_ : ∀ {s} → Label → Signature s → Set ℓ ∈ Sig = foldr (λ ℓ′ → if ⌊ ℓ ≟ ℓ′ ⌋ then (λ _ → ⊤) else id) ⊥ (labels Sig) ------------------------------------------------------------------------ -- Projections -- Signature restriction and projection. (Restriction means removal of -- a given field and all subsequent fields.) Restrict : ∀ {s} (Sig : Signature s) (ℓ : Label) → ℓ ∈ Sig → Signature s Restrict ∅ ℓ () Restrict (Sig , ℓ′ ∶ A) ℓ ℓ∈ with ℓ ≟ ℓ′ ... | yes _ = Sig ... | no _ = Restrict Sig ℓ ℓ∈ Restrict (Sig , ℓ′ ≔ a) ℓ ℓ∈ with ℓ ≟ ℓ′ ... | yes _ = Sig ... | no _ = Restrict Sig ℓ ℓ∈ Restricted : ∀ {s} (Sig : Signature s) (ℓ : Label) → ℓ ∈ Sig → Set s Restricted Sig ℓ ℓ∈ = Record (Restrict Sig ℓ ℓ∈) Proj : ∀ {s} (Sig : Signature s) (ℓ : Label) {ℓ∈ : ℓ ∈ Sig} → Restricted Sig ℓ ℓ∈ → Set s Proj ∅ ℓ {} Proj (Sig , ℓ′ ∶ A) ℓ {ℓ∈} with ℓ ≟ ℓ′ ... | yes _ = A ... | no _ = Proj Sig ℓ {ℓ∈} Proj (_,_≔_ Sig ℓ′ {A = A} a) ℓ {ℓ∈} with ℓ ≟ ℓ′ ... | yes _ = A ... | no _ = Proj Sig ℓ {ℓ∈} -- Record restriction and projection. infixl 5 _∣_ _∣_ : ∀ {s} {Sig : Signature s} → Record Sig → (ℓ : Label) {ℓ∈ : ℓ ∈ Sig} → Restricted Sig ℓ ℓ∈ _∣_ {Sig = ∅} r ℓ {} _∣_ {Sig = Sig , ℓ′ ∶ A} (rec r) ℓ {ℓ∈} with ℓ ≟ ℓ′ ... | yes _ = Σ.proj₁ r ... | no _ = _∣_ (Σ.proj₁ r) ℓ {ℓ∈} _∣_ {Sig = Sig , ℓ′ ≔ a} (rec r) ℓ {ℓ∈} with ℓ ≟ ℓ′ ... | yes _ = Manifest-Σ.proj₁ r ... | no _ = _∣_ (Manifest-Σ.proj₁ r) ℓ {ℓ∈} infixl 5 _·_ _·_ : ∀ {s} {Sig : Signature s} (r : Record Sig) (ℓ : Label) {ℓ∈ : ℓ ∈ Sig} → Proj Sig ℓ {ℓ∈} (r ∣ ℓ) _·_ {Sig = ∅} r ℓ {} _·_ {Sig = Sig , ℓ′ ∶ A} (rec r) ℓ {ℓ∈} with ℓ ≟ ℓ′ ... | yes _ = Σ.proj₂ r ... | no _ = _·_ (Σ.proj₁ r) ℓ {ℓ∈} _·_ {Sig = Sig , ℓ′ ≔ a} (rec r) ℓ {ℓ∈} with ℓ ≟ ℓ′ ... | yes _ = Manifest-Σ.proj₂ r ... | no _ = _·_ (Manifest-Σ.proj₁ r) ℓ {ℓ∈} ------------------------------------------------------------------------ -- With -- Sig With ℓ ≔ a is the signature Sig, but with the ℓ field set to a. mutual infixl 5 _With_≔_ _With_≔_ : ∀ {s} (Sig : Signature s) (ℓ : Label) {ℓ∈ : ℓ ∈ Sig} → ((r : Restricted Sig ℓ ℓ∈) → Proj Sig ℓ r) → Signature s _With_≔_ ∅ ℓ {} a _With_≔_ (Sig , ℓ′ ∶ A) ℓ {ℓ∈} a with ℓ ≟ ℓ′ ... | yes _ = Sig , ℓ′ ≔ a ... | no _ = _With_≔_ Sig ℓ {ℓ∈} a , ℓ′ ∶ A ∘ drop-With _With_≔_ (Sig , ℓ′ ≔ a′) ℓ {ℓ∈} a with ℓ ≟ ℓ′ ... | yes _ = Sig , ℓ′ ≔ a ... | no _ = _With_≔_ Sig ℓ {ℓ∈} a , ℓ′ ≔ a′ ∘ drop-With drop-With : ∀ {s} {Sig : Signature s} {ℓ : Label} {ℓ∈ : ℓ ∈ Sig} {a : (r : Restricted Sig ℓ ℓ∈) → Proj Sig ℓ r} → Record (_With_≔_ Sig ℓ {ℓ∈} a) → Record Sig drop-With {Sig = ∅} {ℓ∈ = ()} r drop-With {Sig = Sig , ℓ′ ∶ A} {ℓ} (rec r) with ℓ ≟ ℓ′ ... | yes _ = rec (Manifest-Σ.proj₁ r , Manifest-Σ.proj₂ r) ... | no _ = rec (drop-With (Σ.proj₁ r) , Σ.proj₂ r) drop-With {Sig = Sig , ℓ′ ≔ a} {ℓ} (rec r) with ℓ ≟ ℓ′ ... | yes _ = rec (Manifest-Σ.proj₁ r ,) ... | no _ = rec (drop-With (Manifest-Σ.proj₁ r) ,) lib-0.7/src/Universe.agda0000644000000000000000000000152212101774706013526 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Universes ------------------------------------------------------------------------ module Universe where open import Data.Product open import Function open import Level -- Universes. record Universe u e : Set (suc (u ⊔ e)) where field -- Codes. U : Set u -- Decoding function. El : U → Set e -- Indexed universes. record Indexed-universe i u e : Set (suc (i ⊔ u ⊔ e)) where field -- Index set. I : Set i -- Codes. U : I → Set u -- Decoding function. El : ∀ {i} → U i → Set e -- An indexed universe can be turned into an unindexed one. unindexed-universe : Universe (i ⊔ u) e unindexed-universe = record { U = ∃ λ i → U i ; El = El ∘ proj₂ } lib-0.7/src/Category/0000755000000000000000000000000012101774706012665 5ustar0000000000000000lib-0.7/src/Category/Functor.agda0000644000000000000000000000104412101774706015122 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Functors ------------------------------------------------------------------------ -- Note that currently the functor laws are not included here. module Category.Functor where open import Function open import Level record RawFunctor {ℓ} (F : Set ℓ → Set ℓ) : Set (suc ℓ) where infixl 4 _<$>_ _<$_ field _<$>_ : ∀ {A B} → (A → B) → F A → F B _<$_ : ∀ {A B} → A → F B → F A x <$ y = const x <$> y lib-0.7/src/Category/Monad.agda0000644000000000000000000000203212101774706014536 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Monads ------------------------------------------------------------------------ -- Note that currently the monad laws are not included here. module Category.Monad where open import Function open import Category.Monad.Indexed open import Data.Unit RawMonad : ∀ {f} → (Set f → Set f) → Set _ RawMonad M = RawIMonad {I = ⊤} (λ _ _ → M) RawMonadZero : ∀ {f} → (Set f → Set f) → Set _ RawMonadZero M = RawIMonadZero {I = ⊤} (λ _ _ → M) RawMonadPlus : ∀ {f} → (Set f → Set f) → Set _ RawMonadPlus M = RawIMonadPlus {I = ⊤} (λ _ _ → M) module RawMonad {f} {M : Set f → Set f} (Mon : RawMonad M) where open RawIMonad Mon public module RawMonadZero {f} {M : Set f → Set f} (Mon : RawMonadZero M) where open RawIMonadZero Mon public module RawMonadPlus {f} {M : Set f → Set f} (Mon : RawMonadPlus M) where open RawIMonadPlus Mon public lib-0.7/src/Category/Applicative.agda0000644000000000000000000000117212101774706015745 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Applicative functors ------------------------------------------------------------------------ -- Note that currently the applicative functor laws are not included -- here. module Category.Applicative where open import Data.Unit open import Category.Applicative.Indexed RawApplicative : ∀ {f} → (Set f → Set f) → Set _ RawApplicative F = RawIApplicative {I = ⊤} (λ _ _ → F) module RawApplicative {f} {F : Set f → Set f} (app : RawApplicative F) where open RawIApplicative app public lib-0.7/src/Category/Applicative/0000755000000000000000000000000012101774706015126 5ustar0000000000000000lib-0.7/src/Category/Applicative/Indexed.agda0000644000000000000000000000467112101774706017334 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Indexed applicative functors ------------------------------------------------------------------------ -- Note that currently the applicative functor laws are not included -- here. module Category.Applicative.Indexed where open import Category.Functor open import Data.Product open import Function open import Level open import Relation.Binary.PropositionalEquality as P using (_≡_) IFun : ∀ {i} → Set i → (ℓ : Level) → Set _ IFun I ℓ = I → I → Set ℓ → Set ℓ record RawIApplicative {i f} {I : Set i} (F : IFun I f) : Set (i ⊔ suc f) where infixl 4 _⊛_ _<⊛_ _⊛>_ infix 4 _⊗_ field pure : ∀ {i A} → A → F i i A _⊛_ : ∀ {i j k A B} → F i j (A → B) → F j k A → F i k B rawFunctor : ∀ {i j} → RawFunctor (F i j) rawFunctor = record { _<$>_ = λ g x → pure g ⊛ x } private open module RF {i j : I} = RawFunctor (rawFunctor {i = i} {j = j}) public _<⊛_ : ∀ {i j k A B} → F i j A → F j k B → F i k A x <⊛ y = const <$> x ⊛ y _⊛>_ : ∀ {i j k A B} → F i j A → F j k B → F i k B x ⊛> y = flip const <$> x ⊛ y _⊗_ : ∀ {i j k A B} → F i j A → F j k B → F i k (A × B) x ⊗ y = (_,_) <$> x ⊛ y zipWith : ∀ {i j k A B C} → (A → B → C) → F i j A → F j k B → F i k C zipWith f x y = f <$> x ⊛ y -- Applicative functor morphisms, specialised to propositional -- equality. record Morphism {i f} {I : Set i} {F₁ F₂ : IFun I f} (A₁ : RawIApplicative F₁) (A₂ : RawIApplicative F₂) : Set (i ⊔ suc f) where module A₁ = RawIApplicative A₁ module A₂ = RawIApplicative A₂ field op : ∀ {i j X} → F₁ i j X → F₂ i j X op-pure : ∀ {i X} (x : X) → op (A₁.pure {i = i} x) ≡ A₂.pure x op-⊛ : ∀ {i j k X Y} (f : F₁ i j (X → Y)) (x : F₁ j k X) → op (A₁._⊛_ f x) ≡ A₂._⊛_ (op f) (op x) op-<$> : ∀ {i j X Y} (f : X → Y) (x : F₁ i j X) → op (A₁._<$>_ f x) ≡ A₂._<$>_ f (op x) op-<$> f x = begin op (A₁._⊛_ (A₁.pure f) x) ≡⟨ op-⊛ _ _ ⟩ A₂._⊛_ (op (A₁.pure f)) (op x) ≡⟨ P.cong₂ A₂._⊛_ (op-pure _) P.refl ⟩ A₂._⊛_ (A₂.pure f) (op x) ∎ where open P.≡-Reasoning lib-0.7/src/Category/Monad/0000755000000000000000000000000012101774706013723 5ustar0000000000000000lib-0.7/src/Category/Monad/Indexed.agda0000644000000000000000000000305412101774706016123 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Indexed monads ------------------------------------------------------------------------ -- Note that currently the monad laws are not included here. module Category.Monad.Indexed where open import Category.Applicative.Indexed open import Function open import Level record RawIMonad {i f} {I : Set i} (M : IFun I f) : Set (i ⊔ suc f) where infixl 1 _>>=_ _>>_ infixr 1 _=<<_ field return : ∀ {i A} → A → M i i A _>>=_ : ∀ {i j k A B} → M i j A → (A → M j k B) → M i k B _>>_ : ∀ {i j k A B} → M i j A → M j k B → M i k B m₁ >> m₂ = m₁ >>= λ _ → m₂ _=<<_ : ∀ {i j k A B} → (A → M j k B) → M i j A → M i k B f =<< c = c >>= f join : ∀ {i j k A} → M i j (M j k A) → M i k A join m = m >>= id rawIApplicative : RawIApplicative M rawIApplicative = record { pure = return ; _⊛_ = λ f x → f >>= λ f' → x >>= λ x' → return (f' x') } open RawIApplicative rawIApplicative public record RawIMonadZero {i f} {I : Set i} (M : IFun I f) : Set (i ⊔ suc f) where field monad : RawIMonad M ∅ : ∀ {i j A} → M i j A open RawIMonad monad public record RawIMonadPlus {i f} {I : Set i} (M : IFun I f) : Set (i ⊔ suc f) where infixr 3 _∣_ field monadZero : RawIMonadZero M _∣_ : ∀ {i j A} → M i j A → M i j A → M i j A open RawIMonadZero monadZero public lib-0.7/src/Category/Monad/Identity.agda0000644000000000000000000000070512101774706016334 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- The identity monad ------------------------------------------------------------------------ module Category.Monad.Identity where open import Category.Monad Identity : ∀ {f} → Set f → Set f Identity A = A IdentityMonad : ∀ {f} → RawMonad (Identity {f}) IdentityMonad = record { return = λ x → x ; _>>=_ = λ x f → f x } lib-0.7/src/Category/Monad/Partiality.agda0000644000000000000000000010564712101774706016700 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- The partiality monad ------------------------------------------------------------------------ module Category.Monad.Partiality where open import Coinduction open import Category.Monad open import Data.Bool open import Data.Nat using (ℕ; zero; suc; _+_) open import Data.Product as Prod hiding (map) open import Data.Sum hiding (map) open import Function open import Function.Equivalence using (_⇔_; equivalence) open import Level using (_⊔_) open import Relation.Binary as B hiding (Rel) open import Relation.Binary.PropositionalEquality as P using (_≡_) open import Relation.Nullary open import Relation.Nullary.Decidable hiding (map) open import Relation.Nullary.Negation ------------------------------------------------------------------------ -- The partiality monad data _⊥ {a} (A : Set a) : Set a where now : (x : A) → A ⊥ later : (x : ∞ (A ⊥)) → A ⊥ monad : ∀ {f} → RawMonad {f = f} _⊥ monad = record { return = now ; _>>=_ = _>>=_ } where _>>=_ : ∀ {A B} → A ⊥ → (A → B ⊥) → B ⊥ now x >>= f = f x later x >>= f = later (♯ (♭ x >>= f)) private module M {f} = RawMonad (monad {f}) -- Non-termination. never : ∀ {a} {A : Set a} → A ⊥ never = later (♯ never) -- run x for n steps peels off at most n "later" constructors from x. run_for_steps : ∀ {a} {A : Set a} → A ⊥ → ℕ → A ⊥ run now x for n steps = now x run later x for zero steps = later x run later x for suc n steps = run ♭ x for n steps -- Is the computation done? isNow : ∀ {a} {A : Set a} → A ⊥ → Bool isNow (now x) = true isNow (later x) = false ------------------------------------------------------------------------ -- Kinds -- The partiality monad comes with two forms of equality (weak and -- strong) and one ordering. Strong equality is stronger than the -- ordering, which is stronger than weak equality. -- The three relations are defined using a single data type, indexed -- by a "kind". data OtherKind : Set where geq weak : OtherKind data Kind : Set where strong : Kind other : (k : OtherKind) → Kind -- Kind equality is decidable. _≟-Kind_ : Decidable (_≡_ {A = Kind}) _≟-Kind_ strong strong = yes P.refl _≟-Kind_ strong (other k) = no λ() _≟-Kind_ (other k) strong = no λ() _≟-Kind_ (other geq) (other geq) = yes P.refl _≟-Kind_ (other geq) (other weak) = no λ() _≟-Kind_ (other weak) (other geq) = no λ() _≟-Kind_ (other weak) (other weak) = yes P.refl -- A predicate which is satisfied only for equalities. Note that, for -- concrete inputs, this predicate evaluates to ⊤ or ⊥. Equality : Kind → Set Equality k = False (k ≟-Kind other geq) ------------------------------------------------------------------------ -- Equality/ordering module Equality {a ℓ} {A : Set a} -- The "return type". (_∼_ : A → A → Set ℓ) where -- The three relations. data Rel : Kind → A ⊥ → A ⊥ → Set (a ⊔ ℓ) where now : ∀ {k x y} (x∼y : x ∼ y) → Rel k (now x) (now y) later : ∀ {k x y} (x∼y : ∞ (Rel k (♭ x) (♭ y))) → Rel k (later x) (later y) laterʳ : ∀ {x y} (x≈y : Rel (other weak) x (♭ y) ) → Rel (other weak) x (later y) laterˡ : ∀ {k x y} (x∼y : Rel (other k) (♭ x) y ) → Rel (other k) (later x) y infix 4 _≅_ _≳_ _≲_ _≈_ _≅_ : A ⊥ → A ⊥ → Set _ _≅_ = Rel strong _≳_ : A ⊥ → A ⊥ → Set _ _≳_ = Rel (other geq) _≲_ : A ⊥ → A ⊥ → Set _ _≲_ = flip _≳_ _≈_ : A ⊥ → A ⊥ → Set _ _≈_ = Rel (other weak) -- x ⇓ y means that x terminates with y. infix 4 _⇓[_]_ _⇓_ _⇓[_]_ : A ⊥ → Kind → A → Set _ x ⇓[ k ] y = Rel k x (now y) _⇓_ : A ⊥ → A → Set _ x ⇓ y = x ⇓[ other weak ] y -- x ⇓ means that x terminates. infix 4 _⇓ _⇓ : A ⊥ → Set _ x ⇓ = ∃ λ v → x ⇓ v -- x ⇑ means that x does not terminate. infix 4 _⇑[_] _⇑ _⇑[_] : A ⊥ → Kind → Set _ x ⇑[ k ] = Rel k x never _⇑ : A ⊥ → Set _ x ⇑ = x ⇑[ other weak ] ------------------------------------------------------------------------ -- Lemmas relating the three relations module _ {a ℓ} {A : Set a} {_∼_ : A → A → Set ℓ} where open Equality _∼_ using (Rel; _≅_; _≳_; _≲_; _≈_; _⇓[_]_; _⇑[_]) open Equality.Rel -- All relations include strong equality. ≅⇒ : ∀ {k} {x y : A ⊥} → x ≅ y → Rel k x y ≅⇒ (now x∼y) = now x∼y ≅⇒ (later x≅y) = later (♯ ≅⇒ (♭ x≅y)) -- The weak equality includes the ordering. ≳⇒ : ∀ {k} {x y : A ⊥} → x ≳ y → Rel (other k) x y ≳⇒ (now x∼y) = now x∼y ≳⇒ (later x≳y) = later (♯ ≳⇒ (♭ x≳y)) ≳⇒ (laterˡ x≳y) = laterˡ (≳⇒ x≳y ) -- Weak equality includes the other relations. ⇒≈ : ∀ {k} {x y : A ⊥} → Rel k x y → x ≈ y ⇒≈ {strong} = ≅⇒ ⇒≈ {other geq} = ≳⇒ ⇒≈ {other weak} = id -- The relations agree for non-terminating computations. never⇒never : ∀ {k₁ k₂} {x : A ⊥} → Rel k₁ x never → Rel k₂ x never never⇒never (later x∼never) = later (♯ never⇒never (♭ x∼never)) never⇒never (laterʳ x≈never) = never⇒never x≈never never⇒never (laterˡ x∼never) = later (♯ never⇒never x∼never) -- The "other" relations agree when the right-hand side is a value. now⇒now : ∀ {k₁ k₂} {x} {y : A} → Rel (other k₁) x (now y) → Rel (other k₂) x (now y) now⇒now (now x∼y) = now x∼y now⇒now (laterˡ x∼now) = laterˡ (now⇒now x∼now) ------------------------------------------------------------------------ -- Later can be dropped laterʳ⁻¹ : ∀ {k} {x : A ⊥} {y} → Rel (other k) x (later y) → Rel (other k) x (♭ y) laterʳ⁻¹ (later x∼y) = laterˡ (♭ x∼y) laterʳ⁻¹ (laterʳ x≈y) = x≈y laterʳ⁻¹ (laterˡ x∼ly) = laterˡ (laterʳ⁻¹ x∼ly) laterˡ⁻¹ : ∀ {x} {y : A ⊥} → later x ≈ y → ♭ x ≈ y laterˡ⁻¹ (later x≈y) = laterʳ (♭ x≈y) laterˡ⁻¹ (laterʳ lx≈y) = laterʳ (laterˡ⁻¹ lx≈y) laterˡ⁻¹ (laterˡ x≈y) = x≈y later⁻¹ : ∀ {k} {x y : ∞ (A ⊥)} → Rel k (later x) (later y) → Rel k (♭ x) (♭ y) later⁻¹ (later x∼y) = ♭ x∼y later⁻¹ (laterʳ lx≈y) = laterˡ⁻¹ lx≈y later⁻¹ (laterˡ x∼ly) = laterʳ⁻¹ x∼ly ------------------------------------------------------------------------ -- The relations are equivalences or partial orders, given suitable -- assumptions about the underlying relation module Equivalence where -- Reflexivity. refl : Reflexive _∼_ → ∀ {k} → Reflexive (Rel k) refl refl-∼ {x = now v} = now refl-∼ refl refl-∼ {x = later x} = later (♯ refl refl-∼) -- Symmetry. sym : Symmetric _∼_ → ∀ {k} → Equality k → Symmetric (Rel k) sym sym-∼ eq (now x∼y) = now (sym-∼ x∼y) sym sym-∼ eq (later x∼y) = later (♯ sym sym-∼ eq (♭ x∼y)) sym sym-∼ eq (laterʳ x≈y) = laterˡ (sym sym-∼ eq x≈y ) sym sym-∼ eq (laterˡ {weak} x≈y) = laterʳ (sym sym-∼ eq x≈y ) sym sym-∼ () (laterˡ {geq} x≳y) -- Transitivity. private module Trans (trans-∼ : Transitive _∼_) where now-trans : ∀ {k x y} {v : A} → Rel k x y → Rel k y (now v) → Rel k x (now v) now-trans (now x∼y) (now y∼z) = now (trans-∼ x∼y y∼z) now-trans (laterˡ x∼y) y∼z = laterˡ (now-trans x∼y y∼z) now-trans x∼ly (laterˡ y∼z) = now-trans (laterʳ⁻¹ x∼ly) y∼z mutual later-trans : ∀ {k} {x y : A ⊥} {z} → Rel k x y → Rel k y (later z) → Rel k x (later z) later-trans (later x∼y) ly∼lz = later (♯ trans (♭ x∼y) (later⁻¹ ly∼lz)) later-trans (laterˡ x∼y) y∼lz = later (♯ trans x∼y (laterʳ⁻¹ y∼lz)) later-trans (laterʳ x≈y) ly≈lz = later-trans x≈y (laterˡ⁻¹ ly≈lz) later-trans x≈y (laterʳ y≈z) = laterʳ ( trans x≈y y≈z ) trans : ∀ {k} {x y z : A ⊥} → Rel k x y → Rel k y z → Rel k x z trans {z = now v} x∼y y∼v = now-trans x∼y y∼v trans {z = later z} x∼y y∼lz = later-trans x∼y y∼lz open Trans public using (trans) -- All the relations are preorders. preorder : IsPreorder _≡_ _∼_ → Kind → Preorder _ _ _ preorder pre k = record { Carrier = A ⊥ ; _≈_ = _≡_ ; _∼_ = Rel k ; isPreorder = record { isEquivalence = P.isEquivalence ; reflexive = refl′ ; trans = Equivalence.trans (IsPreorder.trans pre) } } where refl′ : ∀ {k} {x y : A ⊥} → x ≡ y → Rel k x y refl′ P.refl = Equivalence.refl (IsPreorder.refl pre) private preorder′ : IsEquivalence _∼_ → Kind → Preorder _ _ _ preorder′ equiv = preorder (Setoid.isPreorder (record { isEquivalence = equiv })) -- The two equalities are equivalence relations. setoid : IsEquivalence _∼_ → (k : Kind) {eq : Equality k} → Setoid _ _ setoid equiv k {eq} = record { Carrier = A ⊥ ; _≈_ = Rel k ; isEquivalence = record { refl = Pre.refl ; sym = Equivalence.sym (IsEquivalence.sym equiv) eq ; trans = Pre.trans } } where module Pre = Preorder (preorder′ equiv k) -- The order is a partial order, with strong equality as the -- underlying equality. ≳-poset : IsEquivalence _∼_ → Poset _ _ _ ≳-poset equiv = record { Carrier = A ⊥ ; _≈_ = _≅_ ; _≤_ = _≳_ ; isPartialOrder = record { antisym = antisym ; isPreorder = record { isEquivalence = S.isEquivalence ; reflexive = ≅⇒ ; trans = Pre.trans } } } where module S = Setoid (setoid equiv strong) module Pre = Preorder (preorder′ equiv (other geq)) antisym : {x y : A ⊥} → x ≳ y → x ≲ y → x ≅ y antisym (now x∼y) (now _) = now x∼y antisym (later x≳y) (later x≲y) = later (♯ antisym (♭ x≳y) (♭ x≲y)) antisym (later x≳y) (laterˡ x≲ly) = later (♯ antisym (♭ x≳y) (laterʳ⁻¹ x≲ly)) antisym (laterˡ x≳ly) (later x≲y) = later (♯ antisym (laterʳ⁻¹ x≳ly) (♭ x≲y)) antisym (laterˡ x≳ly) (laterˡ x≲ly) = later (♯ antisym (laterʳ⁻¹ x≳ly) (laterʳ⁻¹ x≲ly)) -- Equational reasoning. module Reasoning (isEquivalence : IsEquivalence _∼_) where private module Pre {k} = Preorder (preorder′ isEquivalence k) module S {k eq} = Setoid (setoid isEquivalence k {eq}) infix 2 _∎ infixr 2 _≡⟨_⟩_ _≅⟨_⟩_ _≳⟨_⟩_ _≈⟨_⟩_ _≡⟨_⟩_ : ∀ {k} x {y z : A ⊥} → x ≡ y → Rel k y z → Rel k x z _ ≡⟨ P.refl ⟩ y∼z = y∼z _≅⟨_⟩_ : ∀ {k} x {y z : A ⊥} → x ≅ y → Rel k y z → Rel k x z _ ≅⟨ x≅y ⟩ y∼z = Pre.trans (≅⇒ x≅y) y∼z _≳⟨_⟩_ : ∀ {k} x {y z : A ⊥} → x ≳ y → Rel (other k) y z → Rel (other k) x z _ ≳⟨ x≳y ⟩ y∼z = Pre.trans (≳⇒ x≳y) y∼z _≈⟨_⟩_ : ∀ x {y z : A ⊥} → x ≈ y → y ≈ z → x ≈ z _ ≈⟨ x≈y ⟩ y≈z = Pre.trans x≈y y≈z sym : ∀ {k} {eq : Equality k} {x y : A ⊥} → Rel k x y → Rel k y x sym {eq = eq} = S.sym {eq = eq} _∎ : ∀ {k} (x : A ⊥) → Rel k x x x ∎ = Pre.refl ------------------------------------------------------------------------ -- Lemmas related to now and never -- Now is not never. now≉never : ∀ {k} {x : A} → ¬ Rel k (now x) never now≉never (laterʳ hyp) = now≉never hyp -- A partial value is either now or never (classically, when the -- underlying relation is reflexive). now-or-never : Reflexive _∼_ → ∀ {k} (x : A ⊥) → ¬ ¬ ((∃ λ y → x ⇓[ other k ] y) ⊎ x ⇑[ other k ]) now-or-never refl x = helper <$> excluded-middle where open RawMonad ¬¬-Monad not-now-is-never : (x : A ⊥) → (∄ λ y → x ≳ now y) → x ≳ never not-now-is-never (now x) hyp with hyp (, now refl) ... | () not-now-is-never (later x) hyp = later (♯ not-now-is-never (♭ x) (hyp ∘ Prod.map id laterˡ)) helper : Dec (∃ λ y → x ≳ now y) → _ helper (yes ≳now) = inj₁ $ Prod.map id ≳⇒ ≳now helper (no ≵now) = inj₂ $ ≳⇒ $ not-now-is-never x ≵now ------------------------------------------------------------------------ -- Map-like results -- Map. map : ∀ {_∼′_ : A → A → Set a} {k} → _∼′_ ⇒ _∼_ → Equality.Rel _∼′_ k ⇒ Equality.Rel _∼_ k map ∼′⇒∼ (now x∼y) = now (∼′⇒∼ x∼y) map ∼′⇒∼ (later x∼y) = later (♯ map ∼′⇒∼ (♭ x∼y)) map ∼′⇒∼ (laterʳ x≈y) = laterʳ (map ∼′⇒∼ x≈y) map ∼′⇒∼ (laterˡ x∼y) = laterˡ (map ∼′⇒∼ x∼y) -- If a statement can be proved using propositional equality as the -- underlying relation, then it can also be proved for any other -- reflexive underlying relation. ≡⇒ : Reflexive _∼_ → ∀ {k x y} → Equality.Rel _≡_ k x y → Rel k x y ≡⇒ refl-∼ = map (flip (P.subst (_∼_ _)) refl-∼) ------------------------------------------------------------------------ -- Steps -- The number of later constructors (steps) in the terminating -- computation x. steps : ∀ {k} {x : A ⊥} {y} → x ⇓[ k ] y → ℕ steps (now _) = zero steps .{x = later x} (laterˡ {x = x} x⇓) = suc (steps {x = ♭ x} x⇓) module Steps {trans-∼ : Transitive _∼_} where left-identity : ∀ {k x y} {z : A} (x≅y : x ≅ y) (y⇓z : y ⇓[ k ] z) → steps (Equivalence.trans trans-∼ (≅⇒ x≅y) y⇓z) ≡ steps y⇓z left-identity (now _) (now _) = P.refl left-identity (later x≅y) (laterˡ y⇓z) = P.cong suc $ left-identity (♭ x≅y) y⇓z right-identity : ∀ {k x} {y z : A} (x⇓y : x ⇓[ k ] y) (y≈z : now y ⇓[ k ] z) → steps (Equivalence.trans trans-∼ x⇓y y≈z) ≡ steps x⇓y right-identity (now x∼y) (now y∼z) = P.refl right-identity (laterˡ x∼y) (now y∼z) = P.cong suc $ right-identity x∼y (now y∼z) ------------------------------------------------------------------------ -- Laws related to bind -- Never is a left and right "zero" of bind. left-zero : {B : Set a} (f : B → A ⊥) → let open M in (never >>= f) ≅ never left-zero f = later (♯ left-zero f) right-zero : ∀ {B} (x : B ⊥) → let open M in (x >>= λ _ → never) ≅ never right-zero (later x) = later (♯ right-zero (♭ x)) right-zero (now x) = never≅never where never≅never : never ≅ never never≅never = later (♯ never≅never) -- Now is a left and right identity of bind (for a reflexive -- underlying relation). left-identity : Reflexive _∼_ → ∀ {B} (x : B) (f : B → A ⊥) → let open M in (now x >>= f) ≅ f x left-identity refl-∼ x f = Equivalence.refl refl-∼ right-identity : Reflexive _∼_ → (x : A ⊥) → let open M in (x >>= now) ≅ x right-identity refl (now x) = now refl right-identity refl (later x) = later (♯ right-identity refl (♭ x)) -- Bind is associative (for a reflexive underlying relation). associative : Reflexive _∼_ → ∀ {B C} (x : C ⊥) (f : C → B ⊥) (g : B → A ⊥) → let open M in (x >>= f >>= g) ≅ (x >>= λ y → f y >>= g) associative refl-∼ (now x) f g = Equivalence.refl refl-∼ associative refl-∼ (later x) f g = later (♯ associative refl-∼ (♭ x) f g) module _ {s ℓ} {A B : Set s} {_∼A_ : A → A → Set ℓ} {_∼B_ : B → B → Set ℓ} where open Equality private open module EqA = Equality _∼A_ using () renaming (_⇓[_]_ to _⇓[_]A_; _⇑[_] to _⇑[_]A) open module EqB = Equality _∼B_ using () renaming (_⇓[_]_ to _⇓[_]B_; _⇑[_] to _⇑[_]B) -- Bind preserves all the relations. _>>=-cong_ : ∀ {k} {x₁ x₂ : A ⊥} {f₁ f₂ : A → B ⊥} → let open M in Rel _∼A_ k x₁ x₂ → (∀ {x₁ x₂} → x₁ ∼A x₂ → Rel _∼B_ k (f₁ x₁) (f₂ x₂)) → Rel _∼B_ k (x₁ >>= f₁) (x₂ >>= f₂) now x₁∼x₂ >>=-cong f₁∼f₂ = f₁∼f₂ x₁∼x₂ later x₁∼x₂ >>=-cong f₁∼f₂ = later (♯ (♭ x₁∼x₂ >>=-cong f₁∼f₂)) laterʳ x₁≈x₂ >>=-cong f₁≈f₂ = laterʳ (x₁≈x₂ >>=-cong f₁≈f₂) laterˡ x₁∼x₂ >>=-cong f₁∼f₂ = laterˡ (x₁∼x₂ >>=-cong f₁∼f₂) -- Inversion lemmas for bind. >>=-inversion-⇓ : Reflexive _∼A_ → ∀ {k} x {f : A → B ⊥} {y} → let open M in (x>>=f⇓ : (x >>= f) ⇓[ k ]B y) → ∃ λ z → ∃₂ λ (x⇓ : x ⇓[ k ]A z) (fz⇓ : f z ⇓[ k ]B y) → steps x⇓ + steps fz⇓ ≡ steps x>>=f⇓ >>=-inversion-⇓ refl (now x) fx⇓ = (x , now refl , fx⇓ , P.refl) >>=-inversion-⇓ refl (later x) (laterˡ x>>=f⇓) = Prod.map id (Prod.map laterˡ (Prod.map id (P.cong suc))) $ >>=-inversion-⇓ refl (♭ x) x>>=f⇓ >>=-inversion-⇑ : IsEquivalence _∼A_ → ∀ {k} x {f : A → B ⊥} → let open M in Rel _∼B_ (other k) (x >>= f) never → ¬ ¬ (x ⇑[ other k ]A ⊎ ∃ λ y → x ⇓[ other k ]A y × f y ⇑[ other k ]B) >>=-inversion-⇑ eqA {k} x {f} ∼never = helper <$> now-or-never IsEqA.refl x where open RawMonad ¬¬-Monad using (_<$>_) open M using (_>>=_) open Reasoning eqA module IsEqA = IsEquivalence eqA k≳ = other geq is-never : ∀ {x y} → x ⇓[ k≳ ]A y → (x >>= f) ⇑[ k≳ ]B → ∃ λ z → y ∼A z × f z ⇑[ k≳ ]B is-never (now x∼y) = λ fx⇑ → (_ , IsEqA.sym x∼y , fx⇑) is-never (laterˡ ≳now) = is-never ≳now ∘ later⁻¹ helper : (∃ λ y → x ⇓[ k≳ ]A y) ⊎ x ⇑[ k≳ ]A → x ⇑[ other k ]A ⊎ ∃ λ y → x ⇓[ other k ]A y × f y ⇑[ other k ]B helper (inj₂ ≳never) = inj₁ (≳⇒ ≳never) helper (inj₁ (y , ≳now)) with is-never ≳now (never⇒never ∼never) ... | (z , y∼z , fz⇑) = inj₂ (z , ≳⇒ (x ≳⟨ ≳now ⟩ now y ≅⟨ now y∼z ⟩ now z ∎) , ≳⇒ fz⇑) module _ {ℓ} {A B : Set ℓ} {_∼_ : B → B → Set ℓ} where open Equality -- A variant of _>>=-cong_. _≡->>=-cong_ : ∀ {k} {x₁ x₂ : A ⊥} {f₁ f₂ : A → B ⊥} → let open M in Rel _≡_ k x₁ x₂ → (∀ x → Rel _∼_ k (f₁ x) (f₂ x)) → Rel _∼_ k (x₁ >>= f₁) (x₂ >>= f₂) _≡->>=-cong_ {k} {f₁ = f₁} {f₂} x₁≈x₂ f₁≈f₂ = x₁≈x₂ >>=-cong λ {x} x≡x′ → P.subst (λ y → Rel _∼_ k (f₁ x) (f₂ y)) x≡x′ (f₁≈f₂ x) ------------------------------------------------------------------------ -- Productivity checker workaround -- The monad can be awkward to use, due to the limitations of guarded -- coinduction. The following code provides a (limited) workaround. module Workaround {a} where infixl 1 _>>=_ data _⊥P : Set a → Set (Level.suc a) where now : ∀ {A} (x : A) → A ⊥P later : ∀ {A} (x : ∞ (A ⊥P)) → A ⊥P _>>=_ : ∀ {A B} (x : A ⊥P) (f : A → B ⊥P) → B ⊥P private data _⊥W : Set a → Set (Level.suc a) where now : ∀ {A} (x : A) → A ⊥W later : ∀ {A} (x : A ⊥P) → A ⊥W mutual _>>=W_ : ∀ {A B} → A ⊥W → (A → B ⊥P) → B ⊥W now x >>=W f = whnf (f x) later x >>=W f = later (x >>= f) whnf : ∀ {A} → A ⊥P → A ⊥W whnf (now x) = now x whnf (later x) = later (♭ x) whnf (x >>= f) = whnf x >>=W f mutual private ⟦_⟧W : ∀ {A} → A ⊥W → A ⊥ ⟦ now x ⟧W = now x ⟦ later x ⟧W = later (♯ ⟦ x ⟧P) ⟦_⟧P : ∀ {A} → A ⊥P → A ⊥ ⟦ x ⟧P = ⟦ whnf x ⟧W -- The definitions above make sense. ⟦_⟧P is homomorphic with -- respect to now, later and _>>=_. module Correct where private open module Eq {A : Set a} = Equality {A = A} _≡_ open module R {A : Set a} = Reasoning (P.isEquivalence {A = A}) now-hom : ∀ {A} (x : A) → ⟦ now x ⟧P ≅ now x now-hom x = now x ∎ later-hom : ∀ {A} (x : ∞ (A ⊥P)) → ⟦ later x ⟧P ≅ later (♯ ⟦ ♭ x ⟧P) later-hom x = later (♯ (⟦ ♭ x ⟧P ∎)) mutual private >>=-homW : ∀ {A B} (x : B ⊥W) (f : B → A ⊥P) → ⟦ x >>=W f ⟧W ≅ M._>>=_ ⟦ x ⟧W (λ y → ⟦ f y ⟧P) >>=-homW (now x) f = ⟦ f x ⟧P ∎ >>=-homW (later x) f = later (♯ >>=-hom x f) >>=-hom : ∀ {A B} (x : B ⊥P) (f : B → A ⊥P) → ⟦ x >>= f ⟧P ≅ M._>>=_ ⟦ x ⟧P (λ y → ⟦ f y ⟧P) >>=-hom x f = >>=-homW (whnf x) f ------------------------------------------------------------------------ -- An alternative, but equivalent, formulation of equality/ordering module AlternativeEquality {a ℓ} where private El : Setoid a ℓ → Set _ El = Setoid.Carrier Eq : ∀ S → B.Rel (El S) _ Eq = Setoid._≈_ open Equality using (Rel) open Equality.Rel infix 4 _∣_≅P_ _∣_≳P_ _∣_≈P_ infix 2 _∎ infixr 2 _≡⟨_⟩_ _≅⟨_⟩_ _≳⟨_⟩_ _≳⟨_⟩≅_ _≳⟨_⟩≈_ _≈⟨_⟩≅_ _≈⟨_⟩≲_ infixl 1 _>>=_ mutual -- Proof "programs". _∣_≅P_ : ∀ S → B.Rel (El S ⊥) _ _∣_≅P_ = flip RelP strong _∣_≳P_ : ∀ S → B.Rel (El S ⊥) _ _∣_≳P_ = flip RelP (other geq) _∣_≈P_ : ∀ S → B.Rel (El S ⊥) _ _∣_≈P_ = flip RelP (other weak) data RelP S : Kind → B.Rel (El S ⊥) (Level.suc (a ⊔ ℓ)) where -- Congruences. now : ∀ {k x y} (xRy : x ⟨ Eq S ⟩ y) → RelP S k (now x) (now y) later : ∀ {k x y} (x∼y : ∞ (RelP S k (♭ x) (♭ y))) → RelP S k (later x) (later y) _>>=_ : ∀ {S′ : Setoid a ℓ} {k} {x₁ x₂} {f₁ f₂ : El S′ → El S ⊥} → let open M in (x₁∼x₂ : RelP S′ k x₁ x₂) (f₁∼f₂ : ∀ {x y} → x ⟨ Eq S′ ⟩ y → RelP S k (f₁ x) (f₂ y)) → RelP S k (x₁ >>= f₁) (x₂ >>= f₂) -- Ordering/weak equality. laterʳ : ∀ {x y} (x≈y : RelP S (other weak) x (♭ y)) → RelP S (other weak) x (later y) laterˡ : ∀ {k x y} (x∼y : RelP S (other k) (♭ x) y) → RelP S (other k) (later x) y -- Equational reasoning. Note that including full transitivity -- for weak equality would make _∣_≈P_ trivial; a similar -- problem applies to _∣_≳P_ (A ∣ never ≳P now x would be -- provable). Instead the definition of RelP includes limited -- notions of transitivity, similar to weak bisimulation up-to -- various things. _∎ : ∀ {k} x → RelP S k x x sym : ∀ {k x y} {eq : Equality k} (x∼y : RelP S k x y) → RelP S k y x _≡⟨_⟩_ : ∀ {k} x {y z} (x≡y : x ≡ y) (y∼z : RelP S k y z) → RelP S k x z _≅⟨_⟩_ : ∀ {k} x {y z} (x≅y : S ∣ x ≅P y) (y∼z : RelP S k y z) → RelP S k x z _≳⟨_⟩_ : let open Equality (Eq S) in ∀ x {y z} (x≳y : x ≳ y) (y≳z : S ∣ y ≳P z) → S ∣ x ≳P z _≳⟨_⟩≅_ : ∀ x {y z} (x≳y : S ∣ x ≳P y) (y≅z : S ∣ y ≅P z) → S ∣ x ≳P z _≳⟨_⟩≈_ : ∀ x {y z} (x≳y : S ∣ x ≳P y) (y≈z : S ∣ y ≈P z) → S ∣ x ≈P z _≈⟨_⟩≅_ : ∀ x {y z} (x≈y : S ∣ x ≈P y) (y≅z : S ∣ y ≅P z) → S ∣ x ≈P z _≈⟨_⟩≲_ : ∀ x {y z} (x≈y : S ∣ x ≈P y) (y≲z : S ∣ z ≳P y) → S ∣ x ≈P z -- If any of the following transitivity-like rules were added to -- RelP, then RelP and Rel would no longer be equivalent: -- -- x ≳P y → y ≳P z → x ≳P z -- x ≳P y → y ≳ z → x ≳P z -- x ≲P y → y ≈P z → x ≈P z -- x ≈P y → y ≳P z → x ≈P z -- x ≲ y → y ≈P z → x ≈P z -- x ≈P y → y ≳ z → x ≈P z -- x ≈P y → y ≈P z → x ≈P z -- x ≈P y → y ≈ z → x ≈P z -- x ≈ y → y ≈P z → x ≈P z -- -- The reason is that any of these rules would make it possible -- to derive that never and now x are related. -- RelP is complete with respect to Rel. complete : ∀ {S k} {x y : El S ⊥} → Equality.Rel (Eq S) k x y → RelP S k x y complete (now xRy) = now xRy complete (later x∼y) = later (♯ complete (♭ x∼y)) complete (laterʳ x≈y) = laterʳ (complete x≈y) complete (laterˡ x∼y) = laterˡ (complete x∼y) -- RelP is sound with respect to Rel. private -- Proof WHNFs. data RelW S : Kind → B.Rel (El S ⊥) (Level.suc (a ⊔ ℓ)) where now : ∀ {k x y} (xRy : x ⟨ Eq S ⟩ y) → RelW S k (now x) (now y) later : ∀ {k x y} (x∼y : RelP S k (♭ x) (♭ y)) → RelW S k (later x) (later y) laterʳ : ∀ {x y} (x≈y : RelW S (other weak) x (♭ y)) → RelW S (other weak) x (later y) laterˡ : ∀ {k x y} (x∼y : RelW S (other k) (♭ x) y) → RelW S (other k) (later x) y -- WHNFs can be turned into programs. program : ∀ {S k x y} → RelW S k x y → RelP S k x y program (now xRy) = now xRy program (later x∼y) = later (♯ x∼y) program (laterˡ x∼y) = laterˡ (program x∼y) program (laterʳ x≈y) = laterʳ (program x≈y) -- Lemmas for WHNFs. _>>=W_ : ∀ {A B k x₁ x₂} {f₁ f₂ : El A → El B ⊥} → RelW A k x₁ x₂ → (∀ {x y} → x ⟨ Eq A ⟩ y → RelW B k (f₁ x) (f₂ y)) → RelW B k (M._>>=_ x₁ f₁) (M._>>=_ x₂ f₂) now xRy >>=W f₁∼f₂ = f₁∼f₂ xRy later x∼y >>=W f₁∼f₂ = later (x∼y >>= program ∘ f₁∼f₂) laterʳ x≈y >>=W f₁≈f₂ = laterʳ (x≈y >>=W f₁≈f₂) laterˡ x∼y >>=W f₁∼f₂ = laterˡ (x∼y >>=W f₁∼f₂) reflW : ∀ {S k} x → RelW S k x x reflW {S} (now x) = now (Setoid.refl S) reflW (later x) = later (♭ x ∎) symW : ∀ {S k x y} → Equality k → RelW S k x y → RelW S k y x symW {S} eq (now xRy) = now (Setoid.sym S xRy) symW eq (later x≈y) = later (sym {eq = eq} x≈y) symW eq (laterʳ x≈y) = laterˡ (symW eq x≈y) symW eq (laterˡ {weak} x≈y) = laterʳ (symW eq x≈y) symW () (laterˡ {geq} x≈y) trans≅W : ∀ {S x y z} → RelW S strong x y → RelW S strong y z → RelW S strong x z trans≅W {S} (now xRy) (now yRz) = now (Setoid.trans S xRy yRz) trans≅W (later x≅y) (later y≅z) = later (_ ≅⟨ x≅y ⟩ y≅z) trans≳-W : ∀ {S x y z} → let open Equality (Eq S) in x ≳ y → RelW S (other geq) y z → RelW S (other geq) x z trans≳-W {S} (now xRy) (now yRz) = now (Setoid.trans S xRy yRz) trans≳-W (later x≳y) (later y≳z) = later (_ ≳⟨ ♭ x≳y ⟩ y≳z) trans≳-W (later x≳y) (laterˡ y≳z) = laterˡ (trans≳-W (♭ x≳y) y≳z) trans≳-W (laterˡ x≳y) y≳z = laterˡ (trans≳-W x≳y y≳z) -- Strong equality programs can be turned into WHNFs. whnf≅ : ∀ {S x y} → S ∣ x ≅P y → RelW S strong x y whnf≅ (now xRy) = now xRy whnf≅ (later x≅y) = later (♭ x≅y) whnf≅ (x₁≅x₂ >>= f₁≅f₂) = whnf≅ x₁≅x₂ >>=W λ xRy → whnf≅ (f₁≅f₂ xRy) whnf≅ (x ∎) = reflW x whnf≅ (sym x≅y) = symW _ (whnf≅ x≅y) whnf≅ (x ≡⟨ P.refl ⟩ y≅z) = whnf≅ y≅z whnf≅ (x ≅⟨ x≅y ⟩ y≅z) = trans≅W (whnf≅ x≅y) (whnf≅ y≅z) -- More transitivity lemmas. _⟨_⟩≅_ : ∀ {S k} x {y z} → RelP S k x y → S ∣ y ≅P z → RelP S k x z _⟨_⟩≅_ {k = strong} x x≅y y≅z = x ≅⟨ x≅y ⟩ y≅z _⟨_⟩≅_ {k = other geq} x x≳y y≅z = x ≳⟨ x≳y ⟩≅ y≅z _⟨_⟩≅_ {k = other weak} x x≈y y≅z = x ≈⟨ x≈y ⟩≅ y≅z trans∼≅W : ∀ {S k x y z} → RelW S k x y → RelW S strong y z → RelW S k x z trans∼≅W {S} (now xRy) (now yRz) = now (Setoid.trans S xRy yRz) trans∼≅W (later x∼y) (later y≅z) = later (_ ⟨ x∼y ⟩≅ y≅z) trans∼≅W (laterʳ x≈y) (later y≅z) = laterʳ (trans∼≅W x≈y (whnf≅ y≅z)) trans∼≅W (laterˡ x∼y) y≅z = laterˡ (trans∼≅W x∼y y≅z) trans≅∼W : ∀ {S k x y z} → RelW S strong x y → RelW S k y z → RelW S k x z trans≅∼W {S} (now xRy) (now yRz) = now (Setoid.trans S xRy yRz) trans≅∼W (later x≅y) (later y∼z) = later (_ ≅⟨ x≅y ⟩ y∼z) trans≅∼W (later x≅y) (laterˡ y∼z) = laterˡ (trans≅∼W (whnf≅ x≅y) y∼z) trans≅∼W x≅y (laterʳ ly≈z) = laterʳ (trans≅∼W x≅y ly≈z) -- Order programs can be turned into WHNFs. whnf≳ : ∀ {S x y} → S ∣ x ≳P y → RelW S (other geq) x y whnf≳ (now xRy) = now xRy whnf≳ (later x∼y) = later (♭ x∼y) whnf≳ (laterˡ x≲y) = laterˡ (whnf≳ x≲y) whnf≳ (x₁∼x₂ >>= f₁∼f₂) = whnf≳ x₁∼x₂ >>=W λ xRy → whnf≳ (f₁∼f₂ xRy) whnf≳ (x ∎) = reflW x whnf≳ (sym {eq = ()} x≅y) whnf≳ (x ≡⟨ P.refl ⟩ y≳z) = whnf≳ y≳z whnf≳ (x ≅⟨ x≅y ⟩ y≳z) = trans≅∼W (whnf≅ x≅y) (whnf≳ y≳z) whnf≳ (x ≳⟨ x≳y ⟩ y≳z) = trans≳-W x≳y (whnf≳ y≳z) whnf≳ (x ≳⟨ x≳y ⟩≅ y≅z) = trans∼≅W (whnf≳ x≳y) (whnf≅ y≅z) -- Another transitivity lemma. trans≳≈W : ∀ {S x y z} → RelW S (other geq) x y → RelW S (other weak) y z → RelW S (other weak) x z trans≳≈W {S} (now xRy) (now yRz) = now (Setoid.trans S xRy yRz) trans≳≈W (later x≳y) (later y≈z) = later (_ ≳⟨ x≳y ⟩≈ y≈z) trans≳≈W (laterˡ x≳y) y≈z = laterˡ (trans≳≈W x≳y y≈z) trans≳≈W x≳y (laterʳ y≈z) = laterʳ (trans≳≈W x≳y y≈z) trans≳≈W (later x≳y) (laterˡ y≈z) = laterˡ (trans≳≈W (whnf≳ x≳y) y≈z) -- All programs can be turned into WHNFs. whnf : ∀ {S k x y} → RelP S k x y → RelW S k x y whnf (now xRy) = now xRy whnf (later x∼y) = later (♭ x∼y) whnf (laterʳ x≈y) = laterʳ (whnf x≈y) whnf (laterˡ x∼y) = laterˡ (whnf x∼y) whnf (x₁∼x₂ >>= f₁∼f₂) = whnf x₁∼x₂ >>=W λ xRy → whnf (f₁∼f₂ xRy) whnf (x ∎) = reflW x whnf (sym {eq = eq} x≈y) = symW eq (whnf x≈y) whnf (x ≡⟨ P.refl ⟩ y∼z) = whnf y∼z whnf (x ≅⟨ x≅y ⟩ y∼z) = trans≅∼W (whnf x≅y) (whnf y∼z) whnf (x ≳⟨ x≳y ⟩ y≳z) = trans≳-W x≳y (whnf y≳z) whnf (x ≳⟨ x≳y ⟩≅ y≅z) = trans∼≅W (whnf x≳y) (whnf y≅z) whnf (x ≳⟨ x≳y ⟩≈ y≈z) = trans≳≈W (whnf x≳y) (whnf y≈z) whnf (x ≈⟨ x≈y ⟩≅ y≅z) = trans∼≅W (whnf x≈y) (whnf y≅z) whnf (x ≈⟨ x≈y ⟩≲ y≲z) = symW _ (trans≳≈W (whnf y≲z) (symW _ (whnf x≈y))) mutual -- Soundness. private soundW : ∀ {S k x y} → RelW S k x y → Rel (Eq S) k x y soundW (now xRy) = now xRy soundW (later x∼y) = later (♯ sound x∼y) soundW (laterʳ x≈y) = laterʳ (soundW x≈y) soundW (laterˡ x∼y) = laterˡ (soundW x∼y) sound : ∀ {S k x y} → RelP S k x y → Rel (Eq S) k x y sound x∼y = soundW (whnf x∼y) -- RelP and Rel are equivalent (when the underlying relation is an -- equivalence). correct : ∀ {S k x y} → RelP S k x y ⇔ Rel (Eq S) k x y correct = equivalence sound complete ------------------------------------------------------------------------ -- Another lemma -- Bind is "idempotent". idempotent : ∀ {ℓ} {A : Set ℓ} (B : Setoid ℓ ℓ) → let open M; open Setoid B using (_≈_; Carrier); open Equality _≈_ in (x : A ⊥) (f : A → A → Carrier ⊥) → (x >>= λ y′ → x >>= λ y″ → f y′ y″) ≳ (x >>= λ y′ → f y′ y′) idempotent {A = A} B x f = sound (idem x) where open AlternativeEquality hiding (_>>=_) open M open Equality.Rel using (laterˡ) open Equivalence using (refl) idem : (x : A ⊥) → B ∣ (x >>= λ y′ → x >>= λ y″ → f y′ y″) ≳P (x >>= λ y′ → f y′ y′) idem (now x) = f x x ∎ idem (later x) = later (♯ ( (♭ x >>= λ y′ → later x >>= λ y″ → f y′ y″) ≳⟨ (refl P.refl {x = ♭ x} ≡->>=-cong λ _ → laterˡ (refl (Setoid.refl B))) ⟩ (♭ x >>= λ y′ → ♭ x >>= λ y″ → f y′ y″) ≳⟨ idem (♭ x) ⟩≅ (♭ x >>= λ y′ → f y′ y′) ∎)) ------------------------------------------------------------------------ -- Example private module Example where open Data.Nat open Workaround -- McCarthy's f91: f91′ : ℕ → ℕ ⊥P f91′ n with n ≤? 100 ... | yes _ = later (♯ (f91′ (11 + n) >>= f91′)) ... | no _ = now (n ∸ 10) f91 : ℕ → ℕ ⊥ f91 n = ⟦ f91′ n ⟧P lib-0.7/src/Category/Monad/Continuation.agda0000644000000000000000000000417712101774706017224 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- A delimited continuation monad ------------------------------------------------------------------------ module Category.Monad.Continuation where open import Category.Applicative open import Category.Applicative.Indexed open import Category.Monad open import Category.Monad.Identity open import Category.Monad.Indexed open import Function open import Level ------------------------------------------------------------------------ -- Delimited continuation monads DContT : ∀ {i f} {I : Set i} → (I → Set f) → (Set f → Set f) → IFun I f DContT K M r₂ r₁ a = (a → M (K r₁)) → M (K r₂) DCont : ∀ {i f} {I : Set i} → (I → Set f) → IFun I f DCont K = DContT K Identity DContTIMonad : ∀ {i f} {I : Set i} (K : I → Set f) {M} → RawMonad M → RawIMonad (DContT K M) DContTIMonad K Mon = record { return = λ a k → k a ; _>>=_ = λ c f k → c (flip f k) } where open RawMonad Mon DContIMonad : ∀ {i f} {I : Set i} (K : I → Set f) → RawIMonad (DCont K) DContIMonad K = DContTIMonad K IdentityMonad ------------------------------------------------------------------------ -- Delimited continuation operations record RawIMonadDCont {i f} {I : Set i} (K : I → Set f) (M : IFun I f) : Set (i ⊔ suc f) where field monad : RawIMonad M reset : ∀ {r₁ r₂ r₃} → M r₁ r₂ (K r₂) → M r₃ r₃ (K r₁) shift : ∀ {a r₁ r₂ r₃ r₄} → ((a → M r₁ r₁ (K r₂)) → M r₃ r₄ (K r₄)) → M r₃ r₂ a open RawIMonad monad public DContTIMonadDCont : ∀ {i f} {I : Set i} (K : I → Set f) {M} → RawMonad M → RawIMonadDCont K (DContT K M) DContTIMonadDCont K Mon = record { monad = DContTIMonad K Mon ; reset = λ e k → e return >>= k ; shift = λ e k → e (λ a k' → (k a) >>= k') return } where open RawIMonad Mon DContIMonadDCont : ∀ {i f} {I : Set i} (K : I → Set f) → RawIMonadDCont K (DCont K) DContIMonadDCont K = DContTIMonadDCont K IdentityMonad lib-0.7/src/Category/Monad/State.agda0000644000000000000000000000776512101774706015640 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- The state monad ------------------------------------------------------------------------ module Category.Monad.State where open import Category.Applicative.Indexed open import Category.Monad open import Category.Monad.Identity open import Category.Monad.Indexed open import Data.Product open import Data.Unit open import Function open import Level ------------------------------------------------------------------------ -- Indexed state monads IStateT : ∀ {i f} {I : Set i} → (I → Set f) → (Set f → Set f) → IFun I f IStateT S M i j A = S i → M (A × S j) StateTIMonad : ∀ {i f} {I : Set i} (S : I → Set f) {M} → RawMonad M → RawIMonad (IStateT S M) StateTIMonad S Mon = record { return = λ x s → return (x , s) ; _>>=_ = λ m f s → m s >>= uncurry f } where open RawMonad Mon StateTIMonadZero : ∀ {i f} {I : Set i} (S : I → Set f) {M} → RawMonadZero M → RawIMonadZero (IStateT S M) StateTIMonadZero S Mon = record { monad = StateTIMonad S (RawMonadZero.monad Mon) ; ∅ = const ∅ } where open RawMonadZero Mon StateTIMonadPlus : ∀ {i f} {I : Set i} (S : I → Set f) {M} → RawMonadPlus M → RawIMonadPlus (IStateT S M) StateTIMonadPlus S Mon = record { monadZero = StateTIMonadZero S (RawIMonadPlus.monadZero Mon) ; _∣_ = λ m₁ m₂ s → m₁ s ∣ m₂ s } where open RawMonadPlus Mon ------------------------------------------------------------------------ -- State monad operations record RawIMonadState {i f} {I : Set i} (S : I → Set f) (M : IFun I f) : Set (i ⊔ suc f) where field monad : RawIMonad M get : ∀ {i} → M i i (S i) put : ∀ {i j} → S j → M i j (Lift ⊤) open RawIMonad monad public modify : ∀ {i j} → (S i → S j) → M i j (Lift ⊤) modify f = get >>= put ∘ f StateTIMonadState : ∀ {i f} {I : Set i} (S : I → Set f) {M} → RawMonad M → RawIMonadState S (IStateT S M) StateTIMonadState S Mon = record { monad = StateTIMonad S Mon ; get = λ s → return (s , s) ; put = λ s _ → return (_ , s) } where open RawIMonad Mon ------------------------------------------------------------------------ -- Ordinary state monads RawMonadState : ∀ {f} → Set f → (Set f → Set f) → Set _ RawMonadState S M = RawIMonadState {I = ⊤} (λ _ → S) (λ _ _ → M) module RawMonadState {f} {S : Set f} {M : Set f → Set f} (Mon : RawMonadState S M) where open RawIMonadState Mon public StateT : ∀ {f} → Set f → (Set f → Set f) → Set f → Set f StateT S M = IStateT {I = ⊤} (λ _ → S) M _ _ StateTMonad : ∀ {f} (S : Set f) {M} → RawMonad M → RawMonad (StateT S M) StateTMonad S = StateTIMonad (λ _ → S) StateTMonadZero : ∀ {f} (S : Set f) {M} → RawMonadZero M → RawMonadZero (StateT S M) StateTMonadZero S = StateTIMonadZero (λ _ → S) StateTMonadPlus : ∀ {f} (S : Set f) {M} → RawMonadPlus M → RawMonadPlus (StateT S M) StateTMonadPlus S = StateTIMonadPlus (λ _ → S) StateTMonadState : ∀ {f} (S : Set f) {M} → RawMonad M → RawMonadState S (StateT S M) StateTMonadState S = StateTIMonadState (λ _ → S) State : ∀ {f} → Set f → Set f → Set f State S = StateT S Identity StateMonad : ∀ {f} (S : Set f) → RawMonad (State S) StateMonad S = StateTMonad S IdentityMonad StateMonadState : ∀ {f} (S : Set f) → RawMonadState S (State S) StateMonadState S = StateTMonadState S IdentityMonad LiftMonadState : ∀ {f S₁} (S₂ : Set f) {M} → RawMonadState S₁ M → RawMonadState S₁ (StateT S₂ M) LiftMonadState S₂ Mon = record { monad = StateTIMonad (λ _ → S₂) monad ; get = λ s → get >>= λ x → return (x , s) ; put = λ s′ s → put s′ >> return (_ , s) } where open RawIMonadState Mon lib-0.7/src/Category/Monad/Partiality/0000755000000000000000000000000012101774706016045 5ustar0000000000000000lib-0.7/src/Category/Monad/Partiality/All.agda0000644000000000000000000001444512101774706017403 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- An All predicate for the partiality monad ------------------------------------------------------------------------ module Category.Monad.Partiality.All where open import Category.Monad open import Category.Monad.Partiality as Partiality using (_⊥; ⇒≈) open import Coinduction open import Function open import Level open import Relation.Binary using (_Respects_; IsEquivalence) open import Relation.Binary.PropositionalEquality as P using (_≡_) open Partiality._⊥ open Partiality.Equality using (Rel) open Partiality.Equality.Rel private open module E {a} {A : Set a} = Partiality.Equality (_≡_ {A = A}) using (_≅_; _≳_) open module M {f} = RawMonad (Partiality.monad {f = f}) using (_>>=_) ------------------------------------------------------------------------ -- All, along with some lemmas -- All P x means that if x terminates with the value v, then P v -- holds. data All {a p} {A : Set a} (P : A → Set p) : A ⊥ → Set (a ⊔ p) where now : ∀ {v} (p : P v) → All P (now v) later : ∀ {x} (p : ∞ (All P (♭ x))) → All P (later x) -- Bind preserves All in the following way: _>>=-cong_ : ∀ {ℓ p q} {A B : Set ℓ} {P : A → Set p} {Q : B → Set q} {x : A ⊥} {f : A → B ⊥} → All P x → (∀ {x} → P x → All Q (f x)) → All Q (x >>= f) now p >>=-cong f = f p later p >>=-cong f = later (♯ (♭ p >>=-cong f)) -- All respects all the relations, given that the predicate respects -- the underlying relation. respects : ∀ {k a p ℓ} {A : Set a} {P : A → Set p} {_∼_ : A → A → Set ℓ} → P Respects _∼_ → All P Respects Rel _∼_ k respects resp (now x∼y) (now p) = now (resp x∼y p) respects resp (later x∼y) (later p) = later (♯ respects resp (♭ x∼y) (♭ p)) respects resp (laterˡ x∼y) (later p) = respects resp x∼y (♭ p) respects resp (laterʳ x≈y) p = later (♯ respects resp x≈y p) respects-flip : ∀ {k a p ℓ} {A : Set a} {P : A → Set p} {_∼_ : A → A → Set ℓ} → P Respects flip _∼_ → All P Respects flip (Rel _∼_ k) respects-flip resp (now x∼y) (now p) = now (resp x∼y p) respects-flip resp (later x∼y) (later p) = later (♯ respects-flip resp (♭ x∼y) (♭ p)) respects-flip resp (laterˡ x∼y) p = later (♯ respects-flip resp x∼y p) respects-flip resp (laterʳ x≈y) (later p) = respects-flip resp x≈y (♭ p) -- "Equational" reasoning. module Reasoning {a p ℓ} {A : Set a} {P : A → Set p} {_∼_ : A → A → Set ℓ} (resp : P Respects flip _∼_) where infix 2 finally infixr 2 _≡⟨_⟩_ _∼⟨_⟩_ _≡⟨_⟩_ : ∀ x {y} → x ≡ y → All P y → All P x _ ≡⟨ P.refl ⟩ p = p _∼⟨_⟩_ : ∀ {k} x {y} → Rel _∼_ k x y → All P y → All P x _ ∼⟨ x∼y ⟩ p = respects-flip resp (⇒≈ x∼y) p -- A cosmetic combinator. finally : (x : A ⊥) → All P x → All P x finally _ p = p syntax finally x p = x ⟨ p ⟩ -- "Equational" reasoning with _∼_ instantiated to propositional -- equality. module Reasoning-≡ {a p} {A : Set a} {P : A → Set p} = Reasoning {P = P} {_∼_ = _≡_} (P.subst P ∘ P.sym) ------------------------------------------------------------------------ -- An alternative, but equivalent, formulation of All module Alternative {a p : Level} where infixr 2 _≅⟨_⟩P_ _≳⟨_⟩P_ infix 2 _⟨_⟩P -- All "programs". data AllP {A : Set a} (P : A → Set p) : A ⊥ → Set (suc (a ⊔ p)) where now : ∀ {x} (p : P x) → AllP P (now x) later : ∀ {x} (p : ∞ (AllP P (♭ x))) → AllP P (later x) _>>=-congP_ : ∀ {B : Set a} {Q : B → Set p} {x f} (p-x : AllP Q x) (p-f : ∀ {v} → Q v → AllP P (f v)) → AllP P (x >>= f) _≅⟨_⟩P_ : ∀ x {y} (x≅y : x ≅ y) (p : AllP P y) → AllP P x _≳⟨_⟩P_ : ∀ x {y} (x≳y : x ≳ y) (p : AllP P y) → AllP P x _⟨_⟩P : ∀ x (p : AllP P x) → AllP P x private -- WHNFs. data AllW {A} (P : A → Set p) : A ⊥ → Set (suc (a ⊔ p)) where now : ∀ {x} (p : P x) → AllW P (now x) later : ∀ {x} (p : AllP P (♭ x)) → AllW P (later x) -- A function which turns WHNFs into programs. program : ∀ {A} {P : A → Set p} {x} → AllW P x → AllP P x program (now p) = now p program (later p) = later (♯ p) -- Functions which turn programs into WHNFs. trans-≅ : ∀ {A} {P : A → Set p} {x y : A ⊥} → x ≅ y → AllW P y → AllW P x trans-≅ (now P.refl) (now p) = now p trans-≅ (later x≅y) (later p) = later (_ ≅⟨ ♭ x≅y ⟩P p) trans-≳ : ∀ {A} {P : A → Set p} {x y : A ⊥} → x ≳ y → AllW P y → AllW P x trans-≳ (now P.refl) (now p) = now p trans-≳ (later x≳y) (later p) = later (_ ≳⟨ ♭ x≳y ⟩P p) trans-≳ (laterˡ x≳y) p = later (_ ≳⟨ x≳y ⟩P program p) mutual _>>=-congW_ : ∀ {A B} {P : A → Set p} {Q : B → Set p} {x f} → AllW P x → (∀ {v} → P v → AllP Q (f v)) → AllW Q (x >>= f) now p >>=-congW p-f = whnf (p-f p) later p >>=-congW p-f = later (p >>=-congP p-f) whnf : ∀ {A} {P : A → Set p} {x} → AllP P x → AllW P x whnf (now p) = now p whnf (later p) = later (♭ p) whnf (p-x >>=-congP p-f) = whnf p-x >>=-congW p-f whnf (_ ≅⟨ x≅y ⟩P p) = trans-≅ x≅y (whnf p) whnf (_ ≳⟨ x≳y ⟩P p) = trans-≳ x≳y (whnf p) whnf (_ ⟨ p ⟩P) = whnf p -- AllP P is sound and complete with respect to All P. sound : ∀ {A} {P : A → Set p} {x} → AllP P x → All P x sound = λ p → soundW (whnf p) where soundW : ∀ {A} {P : A → Set p} {x} → AllW P x → All P x soundW (now p) = now p soundW (later p) = later (♯ sound p) complete : ∀ {A} {P : A → Set p} {x} → All P x → AllP P x complete (now p) = now p complete (later p) = later (♯ complete (♭ p)) lib-0.7/src/Function/0000755000000000000000000000000012101774706012675 5ustar0000000000000000lib-0.7/src/Function/Bijection.agda0000644000000000000000000000525712101774706015432 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Bijections ------------------------------------------------------------------------ module Function.Bijection where open import Data.Product open import Level open import Relation.Binary open import Function.Equality as F using (_⟶_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_) open import Function.Injection as Inj hiding (id; _∘_) open import Function.Surjection as Surj hiding (id; _∘_) open import Function.LeftInverse as Left hiding (id; _∘_) -- Bijective functions. record Bijective {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} (to : From ⟶ To) : Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where field injective : Injective to surjective : Surjective to open Surjective surjective public left-inverse-of : from LeftInverseOf to left-inverse-of x = injective (right-inverse-of (to ⟨$⟩ x)) -- The set of all bijections between two setoids. record Bijection {f₁ f₂ t₁ t₂} (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) : Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where field to : From ⟶ To bijective : Bijective to open Bijective bijective public injection : Injection From To injection = record { to = to ; injective = injective } surjection : Surjection From To surjection = record { to = to ; surjective = surjective } open Surjection surjection public using (equivalence; right-inverse; from-to) left-inverse : LeftInverse From To left-inverse = record { to = to ; from = from ; left-inverse-of = left-inverse-of } open LeftInverse left-inverse public using (to-from) -- Identity and composition. (Note that these proofs are superfluous, -- given that Bijection is equivalent to Function.Inverse.Inverse.) id : ∀ {s₁ s₂} {S : Setoid s₁ s₂} → Bijection S S id {S = S} = record { to = F.id ; bijective = record { injective = Injection.injective (Inj.id {S = S}) ; surjective = Surjection.surjective (Surj.id {S = S}) } } infixr 9 _∘_ _∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂} {F : Setoid f₁ f₂} {M : Setoid m₁ m₂} {T : Setoid t₁ t₂} → Bijection M T → Bijection F M → Bijection F T f ∘ g = record { to = to f ⟪∘⟫ to g ; bijective = record { injective = Injection.injective (Inj._∘_ (injection f) (injection g)) ; surjective = Surjection.surjective (Surj._∘_ (surjection f) (surjection g)) } } where open Bijection lib-0.7/src/Function/LeftInverse.agda0000644000000000000000000000702412101774706015744 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Left inverses ------------------------------------------------------------------------ module Function.LeftInverse where open import Data.Product open import Level import Relation.Binary.EqReasoning as EqReasoning open import Relation.Binary open import Function.Equality as Eq using (_⟶_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_) open import Function.Equivalence using (Equivalence) open import Function.Injection using (Injective; Injection) import Relation.Binary.PropositionalEquality as P -- Left and right inverses. _LeftInverseOf_ : ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} → To ⟶ From → From ⟶ To → Set _ _LeftInverseOf_ {From = From} f g = ∀ x → f ⟨$⟩ (g ⟨$⟩ x) ≈ x where open Setoid From _RightInverseOf_ : ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} → To ⟶ From → From ⟶ To → Set _ f RightInverseOf g = g LeftInverseOf f -- The set of all left inverses between two setoids. record LeftInverse {f₁ f₂ t₁ t₂} (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) : Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where field to : From ⟶ To from : To ⟶ From left-inverse-of : from LeftInverseOf to private open module F = Setoid From open module T = Setoid To open EqReasoning From injective : Injective to injective {x} {y} eq = begin x ≈⟨ F.sym (left-inverse-of x) ⟩ from ⟨$⟩ (to ⟨$⟩ x) ≈⟨ Eq.cong from eq ⟩ from ⟨$⟩ (to ⟨$⟩ y) ≈⟨ left-inverse-of y ⟩ y ∎ injection : Injection From To injection = record { to = to; injective = injective } equivalence : Equivalence From To equivalence = record { to = to ; from = from } to-from : ∀ {x y} → to ⟨$⟩ x T.≈ y → from ⟨$⟩ y F.≈ x to-from {x} {y} to-x≈y = begin from ⟨$⟩ y ≈⟨ Eq.cong from (T.sym to-x≈y) ⟩ from ⟨$⟩ (to ⟨$⟩ x) ≈⟨ left-inverse-of x ⟩ x ∎ -- The set of all right inverses between two setoids. RightInverse : ∀ {f₁ f₂ t₁ t₂} (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) → Set _ RightInverse From To = LeftInverse To From -- The set of all left inverses from one set to another. (Read A ↞ B -- as "surjection from B to A".) infix 3 _↞_ _↞_ : ∀ {f t} → Set f → Set t → Set _ From ↞ To = LeftInverse (P.setoid From) (P.setoid To) -- Identity and composition. id : ∀ {s₁ s₂} {S : Setoid s₁ s₂} → LeftInverse S S id {S = S} = record { to = Eq.id ; from = Eq.id ; left-inverse-of = λ _ → Setoid.refl S } infixr 9 _∘_ _∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂} {F : Setoid f₁ f₂} {M : Setoid m₁ m₂} {T : Setoid t₁ t₂} → LeftInverse M T → LeftInverse F M → LeftInverse F T _∘_ {F = F} f g = record { to = to f ⟪∘⟫ to g ; from = from g ⟪∘⟫ from f ; left-inverse-of = λ x → begin from g ⟨$⟩ (from f ⟨$⟩ (to f ⟨$⟩ (to g ⟨$⟩ x))) ≈⟨ Eq.cong (from g) (left-inverse-of f (to g ⟨$⟩ x)) ⟩ from g ⟨$⟩ (to g ⟨$⟩ x) ≈⟨ left-inverse-of g x ⟩ x ∎ } where open LeftInverse open EqReasoning F lib-0.7/src/Function/Equality.agda0000644000000000000000000000633012101774706015312 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Function setoids and related constructions ------------------------------------------------------------------------ module Function.Equality where open import Function as Fun using (_on_) open import Level import Relation.Binary as B import Relation.Binary.Indexed as I ------------------------------------------------------------------------ -- Functions which preserve equality record Π {f₁ f₂ t₁ t₂} (From : B.Setoid f₁ f₂) (To : I.Setoid (B.Setoid.Carrier From) t₁ t₂) : Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where open I using (_=[_]⇒_) infixl 5 _⟨$⟩_ field _⟨$⟩_ : (x : B.Setoid.Carrier From) → I.Setoid.Carrier To x cong : B.Setoid._≈_ From =[ _⟨$⟩_ ]⇒ I.Setoid._≈_ To open Π public infixr 0 _⟶_ _⟶_ : ∀ {f₁ f₂ t₁ t₂} → B.Setoid f₁ f₂ → B.Setoid t₁ t₂ → Set _ From ⟶ To = Π From (B.Setoid.indexedSetoid To) -- Identity and composition. id : ∀ {a₁ a₂} {A : B.Setoid a₁ a₂} → A ⟶ A id = record { _⟨$⟩_ = Fun.id; cong = Fun.id } infixr 9 _∘_ _∘_ : ∀ {a₁ a₂} {A : B.Setoid a₁ a₂} {b₁ b₂} {B : B.Setoid b₁ b₂} {c₁ c₂} {C : B.Setoid c₁ c₂} → B ⟶ C → A ⟶ B → A ⟶ C f ∘ g = record { _⟨$⟩_ = Fun._∘_ (_⟨$⟩_ f) (_⟨$⟩_ g) ; cong = Fun._∘_ (cong f) (cong g) } -- Constant equality-preserving function. const : ∀ {a₁ a₂} {A : B.Setoid a₁ a₂} {b₁ b₂} {B : B.Setoid b₁ b₂} → B.Setoid.Carrier B → A ⟶ B const {B = B} b = record { _⟨$⟩_ = Fun.const b ; cong = Fun.const (B.Setoid.refl B) } ------------------------------------------------------------------------ -- Function setoids -- Dependent. setoid : ∀ {f₁ f₂ t₁ t₂} (From : B.Setoid f₁ f₂) → I.Setoid (B.Setoid.Carrier From) t₁ t₂ → B.Setoid _ _ setoid From To = record { Carrier = Π From To ; _≈_ = λ f g → ∀ {x y} → x ≈₁ y → f ⟨$⟩ x ≈₂ g ⟨$⟩ y ; isEquivalence = record { refl = λ {f} → cong f ; sym = λ f∼g x∼y → To.sym (f∼g (From.sym x∼y)) ; trans = λ f∼g g∼h x∼y → To.trans (f∼g From.refl) (g∼h x∼y) } } where open module From = B.Setoid From using () renaming (_≈_ to _≈₁_) open module To = I.Setoid To using () renaming (_≈_ to _≈₂_) -- Non-dependent. infixr 0 _⇨_ _⇨_ : ∀ {f₁ f₂ t₁ t₂} → B.Setoid f₁ f₂ → B.Setoid t₁ t₂ → B.Setoid _ _ From ⇨ To = setoid From (B.Setoid.indexedSetoid To) -- A variant of setoid which uses the propositional equality setoid -- for the domain, and a more convenient definition of _≈_. ≡-setoid : ∀ {f t₁ t₂} (From : Set f) → I.Setoid From t₁ t₂ → B.Setoid _ _ ≡-setoid From To = record { Carrier = (x : From) → Carrier x ; _≈_ = λ f g → ∀ x → f x ≈ g x ; isEquivalence = record { refl = λ {f} x → refl ; sym = λ f∼g x → sym (f∼g x) ; trans = λ f∼g g∼h x → trans (f∼g x) (g∼h x) } } where open I.Setoid To lib-0.7/src/Function/Related.agda0000644000000000000000000003212712101774706015100 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- A universe which includes several kinds of "relatedness" for sets, -- such as equivalences, surjections and bijections ------------------------------------------------------------------------ module Function.Related where open import Level open import Function open import Function.Equality using (_⟨$⟩_) open import Function.Equivalence as Eq using (Equivalence) open import Function.Injection as Inj using (Injection; _↣_) open import Function.Inverse as Inv using (Inverse; _↔_) open import Function.LeftInverse as LeftInv using (LeftInverse) open import Function.Surjection as Surj using (Surjection) open import Relation.Binary open import Relation.Binary.PropositionalEquality as P using (_≡_) ------------------------------------------------------------------------ -- Wrapper types -- Synonyms which are used to make _∼[_]_ below "constructor-headed" -- (which implies that Agda can deduce the universe code from an -- expression matching any of the right-hand sides). record _←_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where constructor lam field app-← : B → A open _←_ public record _↢_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where constructor lam field app-↢ : B ↣ A open _↢_ public ------------------------------------------------------------------------ -- Relatedness -- There are several kinds of "relatedness". -- The idea to include kinds other than equivalence and bijection came -- from Simon Thompson and Bengt Nordström. /NAD data Kind : Set where implication reverse-implication equivalence injection reverse-injection left-inverse surjection bijection : Kind -- Interpretation of the codes above. The code "bijection" is -- interpreted as Inverse rather than Bijection; the two types are -- equivalent. infix 4 _∼[_]_ _∼[_]_ : ∀ {ℓ₁ ℓ₂} → Set ℓ₁ → Kind → Set ℓ₂ → Set _ A ∼[ implication ] B = A → B A ∼[ reverse-implication ] B = A ← B A ∼[ equivalence ] B = Equivalence (P.setoid A) (P.setoid B) A ∼[ injection ] B = Injection (P.setoid A) (P.setoid B) A ∼[ reverse-injection ] B = A ↢ B A ∼[ left-inverse ] B = LeftInverse (P.setoid A) (P.setoid B) A ∼[ surjection ] B = Surjection (P.setoid A) (P.setoid B) A ∼[ bijection ] B = Inverse (P.setoid A) (P.setoid B) -- A non-infix synonym. Related : Kind → ∀ {ℓ₁ ℓ₂} → Set ℓ₁ → Set ℓ₂ → Set _ Related k A B = A ∼[ k ] B -- The bijective equality implies any kind of relatedness. ↔⇒ : ∀ {k x y} {X : Set x} {Y : Set y} → X ∼[ bijection ] Y → X ∼[ k ] Y ↔⇒ {implication} = _⟨$⟩_ ∘ Inverse.to ↔⇒ {reverse-implication} = lam ∘′ _⟨$⟩_ ∘ Inverse.from ↔⇒ {equivalence} = Inverse.equivalence ↔⇒ {injection} = Inverse.injection ↔⇒ {reverse-injection} = lam ∘′ Inverse.injection ∘ Inv.sym ↔⇒ {left-inverse} = Inverse.left-inverse ↔⇒ {surjection} = Inverse.surjection ↔⇒ {bijection} = id -- Actual equality also implies any kind of relatedness. ≡⇒ : ∀ {k ℓ} {X Y : Set ℓ} → X ≡ Y → X ∼[ k ] Y ≡⇒ P.refl = ↔⇒ Inv.id ------------------------------------------------------------------------ -- Special kinds of kinds -- Kinds whose interpretation is symmetric. data Symmetric-kind : Set where equivalence bijection : Symmetric-kind -- Forgetful map. ⌊_⌋ : Symmetric-kind → Kind ⌊ equivalence ⌋ = equivalence ⌊ bijection ⌋ = bijection -- The proof of symmetry can be found below. -- Kinds whose interpretation include a function which "goes in the -- forward direction". data Forward-kind : Set where implication equivalence injection left-inverse surjection bijection : Forward-kind -- Forgetful map. ⌊_⌋→ : Forward-kind → Kind ⌊ implication ⌋→ = implication ⌊ equivalence ⌋→ = equivalence ⌊ injection ⌋→ = injection ⌊ left-inverse ⌋→ = left-inverse ⌊ surjection ⌋→ = surjection ⌊ bijection ⌋→ = bijection -- The function. ⇒→ : ∀ {k x y} {X : Set x} {Y : Set y} → X ∼[ ⌊ k ⌋→ ] Y → X → Y ⇒→ {implication} = id ⇒→ {equivalence} = _⟨$⟩_ ∘ Equivalence.to ⇒→ {injection} = _⟨$⟩_ ∘ Injection.to ⇒→ {left-inverse} = _⟨$⟩_ ∘ LeftInverse.to ⇒→ {surjection} = _⟨$⟩_ ∘ Surjection.to ⇒→ {bijection} = _⟨$⟩_ ∘ Inverse.to -- Kinds whose interpretation include a function which "goes backwards". data Backward-kind : Set where reverse-implication equivalence reverse-injection left-inverse surjection bijection : Backward-kind -- Forgetful map. ⌊_⌋← : Backward-kind → Kind ⌊ reverse-implication ⌋← = reverse-implication ⌊ equivalence ⌋← = equivalence ⌊ reverse-injection ⌋← = reverse-injection ⌊ left-inverse ⌋← = left-inverse ⌊ surjection ⌋← = surjection ⌊ bijection ⌋← = bijection -- The function. ⇒← : ∀ {k x y} {X : Set x} {Y : Set y} → X ∼[ ⌊ k ⌋← ] Y → Y → X ⇒← {reverse-implication} = app-← ⇒← {equivalence} = _⟨$⟩_ ∘ Equivalence.from ⇒← {reverse-injection} = _⟨$⟩_ ∘ Injection.to ∘ app-↢ ⇒← {left-inverse} = _⟨$⟩_ ∘ LeftInverse.from ⇒← {surjection} = _⟨$⟩_ ∘ Surjection.from ⇒← {bijection} = _⟨$⟩_ ∘ Inverse.from -- Kinds whose interpretation include functions going in both -- directions. data Equivalence-kind : Set where equivalence left-inverse surjection bijection : Equivalence-kind -- Forgetful map. ⌊_⌋⇔ : Equivalence-kind → Kind ⌊ equivalence ⌋⇔ = equivalence ⌊ left-inverse ⌋⇔ = left-inverse ⌊ surjection ⌋⇔ = surjection ⌊ bijection ⌋⇔ = bijection -- The functions. ⇒⇔ : ∀ {k x y} {X : Set x} {Y : Set y} → X ∼[ ⌊ k ⌋⇔ ] Y → X ∼[ equivalence ] Y ⇒⇔ {equivalence} = id ⇒⇔ {left-inverse} = LeftInverse.equivalence ⇒⇔ {surjection} = Surjection.equivalence ⇒⇔ {bijection} = Inverse.equivalence -- Conversions between special kinds. ⇔⌊_⌋ : Symmetric-kind → Equivalence-kind ⇔⌊ equivalence ⌋ = equivalence ⇔⌊ bijection ⌋ = bijection →⌊_⌋ : Equivalence-kind → Forward-kind →⌊ equivalence ⌋ = equivalence →⌊ left-inverse ⌋ = left-inverse →⌊ surjection ⌋ = surjection →⌊ bijection ⌋ = bijection ←⌊_⌋ : Equivalence-kind → Backward-kind ←⌊ equivalence ⌋ = equivalence ←⌊ left-inverse ⌋ = left-inverse ←⌊ surjection ⌋ = surjection ←⌊ bijection ⌋ = bijection ------------------------------------------------------------------------ -- Opposites -- For every kind there is an opposite kind. _op : Kind → Kind implication op = reverse-implication reverse-implication op = implication equivalence op = equivalence injection op = reverse-injection reverse-injection op = injection left-inverse op = surjection surjection op = left-inverse bijection op = bijection -- For every morphism there is a corresponding reverse morphism of the -- opposite kind. reverse : ∀ {k a b} {A : Set a} {B : Set b} → A ∼[ k ] B → B ∼[ k op ] A reverse {implication} = lam reverse {reverse-implication} = app-← reverse {equivalence} = Eq.sym reverse {injection} = lam reverse {reverse-injection} = app-↢ reverse {left-inverse} = Surj.fromRightInverse reverse {surjection} = Surjection.right-inverse reverse {bijection} = Inv.sym ------------------------------------------------------------------------ -- Equational reasoning -- Equational reasoning for related things. module EquationalReasoning where private refl : ∀ {k ℓ} → Reflexive (Related k {ℓ}) refl {implication} = id refl {reverse-implication} = lam id refl {equivalence} = Eq.id refl {injection} = Inj.id refl {reverse-injection} = lam Inj.id refl {left-inverse} = LeftInv.id refl {surjection} = Surj.id refl {bijection} = Inv.id trans : ∀ {k ℓ₁ ℓ₂ ℓ₃} → Trans (Related k {ℓ₁} {ℓ₂}) (Related k {ℓ₂} {ℓ₃}) (Related k {ℓ₁} {ℓ₃}) trans {implication} = flip _∘′_ trans {reverse-implication} = λ f g → lam (app-← f ∘ app-← g) trans {equivalence} = flip Eq._∘_ trans {injection} = flip Inj._∘_ trans {reverse-injection} = λ f g → lam (Inj._∘_ (app-↢ f) (app-↢ g)) trans {left-inverse} = flip LeftInv._∘_ trans {surjection} = flip Surj._∘_ trans {bijection} = flip Inv._∘_ sym : ∀ {k ℓ₁ ℓ₂} → Sym (Related ⌊ k ⌋ {ℓ₁} {ℓ₂}) (Related ⌊ k ⌋ {ℓ₂} {ℓ₁}) sym {equivalence} = Eq.sym sym {bijection} = Inv.sym infix 2 _∎ infixr 2 _∼⟨_⟩_ _↔⟨_⟩_ _≡⟨_⟩_ _∼⟨_⟩_ : ∀ {k x y z} (X : Set x) {Y : Set y} {Z : Set z} → X ∼[ k ] Y → Y ∼[ k ] Z → X ∼[ k ] Z _ ∼⟨ X↝Y ⟩ Y↝Z = trans X↝Y Y↝Z -- Isomorphisms can be combined with any other kind of relatedness. _↔⟨_⟩_ : ∀ {k x y z} (X : Set x) {Y : Set y} {Z : Set z} → X ↔ Y → Y ∼[ k ] Z → X ∼[ k ] Z X ↔⟨ X↔Y ⟩ Y⇔Z = X ∼⟨ ↔⇒ X↔Y ⟩ Y⇔Z _≡⟨_⟩_ : ∀ {k ℓ z} (X : Set ℓ) {Y : Set ℓ} {Z : Set z} → X ≡ Y → Y ∼[ k ] Z → X ∼[ k ] Z X ≡⟨ X≡Y ⟩ Y⇔Z = X ∼⟨ ≡⇒ X≡Y ⟩ Y⇔Z _∎ : ∀ {k x} (X : Set x) → X ∼[ k ] X X ∎ = refl -- For a symmetric kind and a fixed universe level we can construct a -- setoid. setoid : Symmetric-kind → (ℓ : Level) → Setoid _ _ setoid k ℓ = record { Carrier = Set ℓ ; _≈_ = Related ⌊ k ⌋ ; isEquivalence = record {refl = _ ∎; sym = sym; trans = _∼⟨_⟩_ _} } where open EquationalReasoning -- For an arbitrary kind and a fixed universe level we can construct a -- preorder. preorder : Kind → (ℓ : Level) → Preorder _ _ _ preorder k ℓ = record { Carrier = Set ℓ ; _≈_ = _↔_ ; _∼_ = Related k ; isPreorder = record { isEquivalence = Setoid.isEquivalence (setoid bijection ℓ) ; reflexive = ↔⇒ ; trans = _∼⟨_⟩_ _ } } where open EquationalReasoning ------------------------------------------------------------------------ -- Some induced relations -- Every unary relation induces a preorder and, for symmetric kinds, -- an equivalence. (No claim is made that these relations are unique.) InducedRelation₁ : Kind → ∀ {a s} {A : Set a} → (A → Set s) → A → A → Set _ InducedRelation₁ k S = λ x y → S x ∼[ k ] S y InducedPreorder₁ : Kind → ∀ {a s} {A : Set a} → (A → Set s) → Preorder _ _ _ InducedPreorder₁ k S = record { _≈_ = P._≡_ ; _∼_ = InducedRelation₁ k S ; isPreorder = record { isEquivalence = P.isEquivalence ; reflexive = reflexive ∘ Setoid.reflexive (setoid bijection _) ∘ P.cong S ; trans = trans } } where open Preorder (preorder _ _) InducedEquivalence₁ : Symmetric-kind → ∀ {a s} {A : Set a} → (A → Set s) → Setoid _ _ InducedEquivalence₁ k S = record { _≈_ = InducedRelation₁ ⌊ k ⌋ S ; isEquivalence = record {refl = refl; sym = sym; trans = trans} } where open Setoid (setoid _ _) -- Every binary relation induces a preorder and, for symmetric kinds, -- an equivalence. (No claim is made that these relations are unique.) InducedRelation₂ : Kind → ∀ {a b s} {A : Set a} {B : Set b} → (A → B → Set s) → B → B → Set _ InducedRelation₂ k _S_ = λ x y → ∀ {z} → (z S x) ∼[ k ] (z S y) InducedPreorder₂ : Kind → ∀ {a b s} {A : Set a} {B : Set b} → (A → B → Set s) → Preorder _ _ _ InducedPreorder₂ k _S_ = record { _≈_ = P._≡_ ; _∼_ = InducedRelation₂ k _S_ ; isPreorder = record { isEquivalence = P.isEquivalence ; reflexive = λ x≡y {z} → reflexive $ Setoid.reflexive (setoid bijection _) $ P.cong (_S_ z) x≡y ; trans = λ i↝j j↝k → trans i↝j j↝k } } where open Preorder (preorder _ _) InducedEquivalence₂ : Symmetric-kind → ∀ {a b s} {A : Set a} {B : Set b} → (A → B → Set s) → Setoid _ _ InducedEquivalence₂ k _S_ = record { _≈_ = InducedRelation₂ ⌊ k ⌋ _S_ ; isEquivalence = record { refl = refl ; sym = λ i↝j → sym i↝j ; trans = λ i↝j j↝k → trans i↝j j↝k } } where open Setoid (setoid _ _) lib-0.7/src/Function/Inverse.agda0000644000000000000000000001254612101774706015136 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Inverses ------------------------------------------------------------------------ module Function.Inverse where open import Level open import Function using (flip) open import Function.Bijection hiding (id; _∘_) 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 -- Inverses. record _InverseOf_ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} (from : To ⟶ From) (to : From ⟶ To) : Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where field left-inverse-of : from LeftInverseOf to right-inverse-of : from RightInverseOf to -- The set of all inverses between two setoids. record Inverse {f₁ f₂ t₁ t₂} (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) : Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where field to : From ⟶ To from : To ⟶ From inverse-of : from InverseOf to open _InverseOf_ inverse-of public left-inverse : LeftInverse From To left-inverse = record { to = to ; from = from ; left-inverse-of = left-inverse-of } open LeftInverse left-inverse public using (injective; injection) bijection : Bijection From To bijection = record { to = to ; bijective = record { injective = injective ; surjective = record { from = from ; right-inverse-of = right-inverse-of } } } open Bijection bijection public using (equivalence; surjective; surjection; right-inverse; to-from; from-to) -- The set of all inverses between two sets. infix 3 _↔_ _↔_ : ∀ {f t} → Set f → Set t → Set _ From ↔ To = Inverse (P.setoid From) (P.setoid To) -- If two setoids are in bijective correspondence, then there is an -- inverse between them. fromBijection : ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} → Bijection From To → Inverse From To fromBijection b = record { to = Bijection.to b ; from = Bijection.from b ; inverse-of = record { left-inverse-of = Bijection.left-inverse-of b ; right-inverse-of = Bijection.right-inverse-of b } } ------------------------------------------------------------------------ -- Map and zip map : ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} {f₁′ f₂′ t₁′ t₂′} {From′ : Setoid f₁′ f₂′} {To′ : Setoid t₁′ t₂′} → (t : (From ⟶ To) → (From′ ⟶ To′)) → (f : (To ⟶ From) → (To′ ⟶ From′)) → (∀ {to from} → from InverseOf to → f from InverseOf t to) → Inverse From To → Inverse From′ To′ map t f pres eq = record { to = t to ; from = f from ; inverse-of = pres inverse-of } where open Inverse eq zip : ∀ {f₁₁ f₂₁ t₁₁ t₂₁} {From₁ : Setoid f₁₁ f₂₁} {To₁ : Setoid t₁₁ t₂₁} {f₁₂ f₂₂ t₁₂ t₂₂} {From₂ : Setoid f₁₂ f₂₂} {To₂ : Setoid t₁₂ t₂₂} {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} → (t : (From₁ ⟶ To₁) → (From₂ ⟶ To₂) → (From ⟶ To)) → (f : (To₁ ⟶ From₁) → (To₂ ⟶ From₂) → (To ⟶ From)) → (∀ {to₁ from₁ to₂ from₂} → from₁ InverseOf to₁ → from₂ InverseOf to₂ → f from₁ from₂ InverseOf t to₁ to₂) → Inverse From₁ To₁ → Inverse From₂ To₂ → Inverse From To zip t f pres eq₁ eq₂ = record { to = t (to eq₁) (to eq₂) ; from = f (from eq₁) (from eq₂) ; inverse-of = pres (inverse-of eq₁) (inverse-of eq₂) } where open Inverse ------------------------------------------------------------------------ -- Inverse is an equivalence relation -- Identity and composition (reflexivity and transitivity). id : ∀ {s₁ s₂} → Reflexive (Inverse {s₁} {s₂}) id {x = S} = record { to = F.id ; from = F.id ; inverse-of = record { left-inverse-of = LeftInverse.left-inverse-of id′ ; right-inverse-of = LeftInverse.left-inverse-of id′ } } where id′ = Left.id {S = S} infixr 9 _∘_ _∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂} → TransFlip (Inverse {f₁} {f₂} {m₁} {m₂}) (Inverse {m₁} {m₂} {t₁} {t₂}) (Inverse {f₁} {f₂} {t₁} {t₂}) f ∘ g = record { to = to f ⟪∘⟫ to g ; from = from g ⟪∘⟫ from f ; inverse-of = record { left-inverse-of = LeftInverse.left-inverse-of (Left._∘_ (left-inverse f) (left-inverse g)) ; right-inverse-of = LeftInverse.left-inverse-of (Left._∘_ (right-inverse g) (right-inverse f)) } } where open Inverse -- Symmetry. sym : ∀ {f₁ f₂ t₁ t₂} → Sym (Inverse {f₁} {f₂} {t₁} {t₂}) (Inverse {t₁} {t₂} {f₁} {f₂}) sym inv = record { from = to ; to = from ; inverse-of = record { left-inverse-of = right-inverse-of ; right-inverse-of = left-inverse-of } } where open Inverse inv lib-0.7/src/Function/Surjection.agda0000644000000000000000000000610612101774706015643 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Surjections ------------------------------------------------------------------------ module Function.Surjection where open import Level open import Function.Equality as F using (_⟶_) renaming (_∘_ to _⟪∘⟫_) open import Function.Equivalence using (Equivalence) open import Function.Injection hiding (id; _∘_) open import Function.LeftInverse as Left hiding (id; _∘_) open import Data.Product open import Relation.Binary import Relation.Binary.PropositionalEquality as P -- Surjective functions. record Surjective {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} (to : From ⟶ To) : Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where field from : To ⟶ From right-inverse-of : from RightInverseOf to -- The set of all surjections from one setoid to another. record Surjection {f₁ f₂ t₁ t₂} (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) : Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where field to : From ⟶ To surjective : Surjective to open Surjective surjective public right-inverse : RightInverse From To right-inverse = record { to = from ; from = to ; left-inverse-of = right-inverse-of } open LeftInverse right-inverse public using () renaming (to-from to from-to) injective : Injective from injective = LeftInverse.injective right-inverse injection : Injection To From injection = LeftInverse.injection right-inverse equivalence : Equivalence From To equivalence = record { to = to ; from = from } -- Right inverses can be turned into surjections. fromRightInverse : ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} → RightInverse From To → Surjection From To fromRightInverse r = record { to = from ; surjective = record { from = to ; right-inverse-of = left-inverse-of } } where open LeftInverse r -- The set of all surjections from one set to another. infix 3 _↠_ _↠_ : ∀ {f t} → Set f → Set t → Set _ From ↠ To = Surjection (P.setoid From) (P.setoid To) -- Identity and composition. id : ∀ {s₁ s₂} {S : Setoid s₁ s₂} → Surjection S S id {S = S} = record { to = F.id ; surjective = record { from = LeftInverse.to id′ ; right-inverse-of = LeftInverse.left-inverse-of id′ } } where id′ = Left.id {S = S} infixr 9 _∘_ _∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂} {F : Setoid f₁ f₂} {M : Setoid m₁ m₂} {T : Setoid t₁ t₂} → Surjection M T → Surjection F M → Surjection F T f ∘ g = record { to = to f ⟪∘⟫ to g ; surjective = record { from = LeftInverse.to g∘f ; right-inverse-of = LeftInverse.left-inverse-of g∘f } } where open Surjection g∘f = Left._∘_ (right-inverse g) (right-inverse f) lib-0.7/src/Function/Equivalence.agda0000644000000000000000000000714412101774706015762 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Equivalence (coinhabitance) ------------------------------------------------------------------------ module Function.Equivalence where open import Function using (flip) open import Function.Equality as F using (_⟶_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_) open import Level open import Relation.Binary import Relation.Binary.PropositionalEquality as P -- Setoid equivalence. record Equivalence {f₁ f₂ t₁ t₂} (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) : Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where field to : From ⟶ To from : To ⟶ From -- Set equivalence. infix 3 _⇔_ _⇔_ : ∀ {f t} → Set f → Set t → Set _ From ⇔ To = Equivalence (P.setoid From) (P.setoid To) equivalence : ∀ {f t} {From : Set f} {To : Set t} → (From → To) → (To → From) → From ⇔ To equivalence to from = record { to = P.→-to-⟶ to; from = P.→-to-⟶ from } ------------------------------------------------------------------------ -- Map and zip map : ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} {f₁′ f₂′ t₁′ t₂′} {From′ : Setoid f₁′ f₂′} {To′ : Setoid t₁′ t₂′} → ((From ⟶ To) → (From′ ⟶ To′)) → ((To ⟶ From) → (To′ ⟶ From′)) → Equivalence From To → Equivalence From′ To′ map t f eq = record { to = t to; from = f from } where open Equivalence eq zip : ∀ {f₁₁ f₂₁ t₁₁ t₂₁} {From₁ : Setoid f₁₁ f₂₁} {To₁ : Setoid t₁₁ t₂₁} {f₁₂ f₂₂ t₁₂ t₂₂} {From₂ : Setoid f₁₂ f₂₂} {To₂ : Setoid t₁₂ t₂₂} {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} → ((From₁ ⟶ To₁) → (From₂ ⟶ To₂) → (From ⟶ To)) → ((To₁ ⟶ From₁) → (To₂ ⟶ From₂) → (To ⟶ From)) → Equivalence From₁ To₁ → Equivalence From₂ To₂ → Equivalence From To zip t f eq₁ eq₂ = record { to = t (to eq₁) (to eq₂); from = f (from eq₁) (from eq₂) } where open Equivalence ------------------------------------------------------------------------ -- Equivalence is an equivalence relation -- Identity and composition (reflexivity and transitivity). id : ∀ {s₁ s₂} → Reflexive (Equivalence {s₁} {s₂}) id {x = S} = record { to = F.id ; from = F.id } infixr 9 _∘_ _∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂} → TransFlip (Equivalence {f₁} {f₂} {m₁} {m₂}) (Equivalence {m₁} {m₂} {t₁} {t₂}) (Equivalence {f₁} {f₂} {t₁} {t₂}) f ∘ g = record { to = to f ⟪∘⟫ to g ; from = from g ⟪∘⟫ from f } where open Equivalence -- Symmetry. sym : ∀ {f₁ f₂ t₁ t₂} → Sym (Equivalence {f₁} {f₂} {t₁} {t₂}) (Equivalence {t₁} {t₂} {f₁} {f₂}) sym eq = record { from = to ; to = from } where open Equivalence eq -- For fixed universe levels we can construct setoids. setoid : (s₁ s₂ : Level) → Setoid (suc (s₁ ⊔ s₂)) (s₁ ⊔ s₂) setoid s₁ s₂ = record { Carrier = Setoid s₁ s₂ ; _≈_ = Equivalence ; isEquivalence = record {refl = id; sym = sym; trans = flip _∘_} } ⇔-setoid : (ℓ : Level) → Setoid (suc ℓ) ℓ ⇔-setoid ℓ = record { Carrier = Set ℓ ; _≈_ = _⇔_ ; isEquivalence = record {refl = id; sym = sym; trans = flip _∘_} } lib-0.7/src/Function/Injection.agda0000644000000000000000000000335612101774706015444 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Injections ------------------------------------------------------------------------ module Function.Injection where open import Function as Fun using () renaming (_∘_ to _⟨∘⟩_) open import Level open import Relation.Binary open import Function.Equality as F using (_⟶_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_) import Relation.Binary.PropositionalEquality as P -- Injective functions. Injective : ∀ {a₁ a₂ b₁ b₂} {A : Setoid a₁ a₂} {B : Setoid b₁ b₂} → A ⟶ B → Set _ Injective {A = A} {B} f = ∀ {x y} → f ⟨$⟩ x ≈₂ f ⟨$⟩ y → x ≈₁ y where open Setoid A renaming (_≈_ to _≈₁_) open Setoid B renaming (_≈_ to _≈₂_) -- The set of all injections between two setoids. record Injection {f₁ f₂ t₁ t₂} (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) : Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where field to : From ⟶ To injective : Injective to -- The set of all injections from one set to another. infix 3 _↣_ _↣_ : ∀ {f t} → Set f → Set t → Set _ From ↣ To = Injection (P.setoid From) (P.setoid To) -- Identity and composition. infixr 9 _∘_ id : ∀ {s₁ s₂} {S : Setoid s₁ s₂} → Injection S S id = record { to = F.id; injective = Fun.id } _∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂} {F : Setoid f₁ f₂} {M : Setoid m₁ m₂} {T : Setoid t₁ t₂} → Injection M T → Injection F M → Injection F T f ∘ g = record { to = to f ⟪∘⟫ to g ; injective = (λ {_} → injective g) ⟨∘⟩ injective f } where open Injection lib-0.7/src/Function/Related/0000755000000000000000000000000012101774706014255 5ustar0000000000000000lib-0.7/src/Function/Related/TypeIsomorphisms.agda0000644000000000000000000004077112101774706020442 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Basic lemmas showing that various types are related (isomorphic or -- equivalent or…) ------------------------------------------------------------------------ module Function.Related.TypeIsomorphisms where open import Algebra import Algebra.FunctionProperties as FP import Algebra.Operations import Algebra.RingSolver.Natural-coefficients open import Algebra.Structures open import Data.Empty open import Data.Nat as Nat using (zero; suc) open import Data.Product as Prod hiding (swap) open import Data.Sum as Sum open import Data.Unit open import Level hiding (zero; suc) open import Function open import Function.Equality using (_⟨$⟩_) open import Function.Equivalence as Eq using (_⇔_; module Equivalence) open import Function.Inverse as Inv using (_↔_; module Inverse) open import Function.Related as Related open import Relation.Binary open import Relation.Binary.Product.Pointwise open import Relation.Binary.PropositionalEquality as P using (_≡_; _≗_) open import Relation.Binary.Sum open import Relation.Nullary import Relation.Nullary.Decidable as Dec ------------------------------------------------------------------------ -- Σ is "associative" Σ-assoc : ∀ {a b c} {A : Set a} {B : A → Set b} {C : (a : A) → B a → Set c} → Σ (Σ A B) (uncurry C) ↔ Σ A (λ a → Σ (B a) (C a)) Σ-assoc = record { to = P.→-to-⟶ λ p → proj₁ (proj₁ p) , (proj₂ (proj₁ p) , proj₂ p) ; from = P.→-to-⟶ _ ; inverse-of = record { left-inverse-of = λ _ → P.refl ; right-inverse-of = λ _ → P.refl } } ------------------------------------------------------------------------ -- ⊥, ⊤, _×_ and _⊎_ form a commutative semiring ×-CommutativeMonoid : Symmetric-kind → (ℓ : Level) → CommutativeMonoid _ _ ×-CommutativeMonoid k ℓ = record { Carrier = Set ℓ ; _≈_ = Related ⌊ k ⌋ ; _∙_ = _×_ ; ε = Lift ⊤ ; isCommutativeMonoid = record { isSemigroup = record { isEquivalence = Setoid.isEquivalence $ Related.setoid k ℓ ; assoc = λ _ _ _ → ↔⇒ Σ-assoc ; ∙-cong = _×-cong_ } ; identityˡ = λ A → ↔⇒ $ left-identity A ; comm = λ A B → ↔⇒ $ comm A B } } where open FP _↔_ left-identity : LeftIdentity (Lift {ℓ = ℓ} ⊤) _×_ left-identity _ = record { to = P.→-to-⟶ proj₂ ; from = P.→-to-⟶ λ y → _ , y ; inverse-of = record { left-inverse-of = λ _ → P.refl ; right-inverse-of = λ _ → P.refl } } comm : Commutative _×_ comm _ _ = record { to = P.→-to-⟶ Prod.swap ; from = P.→-to-⟶ Prod.swap ; inverse-of = record { left-inverse-of = λ _ → P.refl ; right-inverse-of = λ _ → P.refl } } ⊎-CommutativeMonoid : Symmetric-kind → (ℓ : Level) → CommutativeMonoid _ _ ⊎-CommutativeMonoid k ℓ = record { Carrier = Set ℓ ; _≈_ = Related ⌊ k ⌋ ; _∙_ = _⊎_ ; ε = Lift ⊥ ; isCommutativeMonoid = record { isSemigroup = record { isEquivalence = Setoid.isEquivalence $ Related.setoid k ℓ ; assoc = λ A B C → ↔⇒ $ assoc A B C ; ∙-cong = _⊎-cong_ } ; identityˡ = λ A → ↔⇒ $ left-identity A ; comm = λ A B → ↔⇒ $ comm A B } } where open FP _↔_ left-identity : LeftIdentity (Lift ⊥) (_⊎_ {a = ℓ} {b = ℓ}) left-identity A = record { to = P.→-to-⟶ [ (λ ()) ∘′ lower , id ] ; from = P.→-to-⟶ inj₂ ; inverse-of = record { right-inverse-of = λ _ → P.refl ; left-inverse-of = [ ⊥-elim ∘ lower , (λ _ → P.refl) ] } } assoc : Associative _⊎_ assoc A B C = record { to = P.→-to-⟶ [ [ inj₁ , inj₂ ∘ inj₁ ] , inj₂ ∘ inj₂ ] ; from = P.→-to-⟶ [ inj₁ ∘ inj₁ , [ inj₁ ∘ inj₂ , inj₂ ] ] ; inverse-of = record { left-inverse-of = [ [ (λ _ → P.refl) , (λ _ → P.refl) ] , (λ _ → P.refl) ] ; right-inverse-of = [ (λ _ → P.refl) , [ (λ _ → P.refl) , (λ _ → P.refl) ] ] } } comm : Commutative _⊎_ comm _ _ = record { to = P.→-to-⟶ swap ; from = P.→-to-⟶ swap ; inverse-of = record { left-inverse-of = inv ; right-inverse-of = inv } } where swap : {A B : Set ℓ} → A ⊎ B → B ⊎ A swap = [ inj₂ , inj₁ ] inv : ∀ {A B} → swap ∘ swap {A} {B} ≗ id inv = [ (λ _ → P.refl) , (λ _ → P.refl) ] ×⊎-CommutativeSemiring : Symmetric-kind → (ℓ : Level) → CommutativeSemiring (Level.suc ℓ) ℓ ×⊎-CommutativeSemiring k ℓ = record { Carrier = Set ℓ ; _≈_ = Related ⌊ k ⌋ ; _+_ = _⊎_ ; _*_ = _×_ ; 0# = Lift ⊥ ; 1# = Lift ⊤ ; isCommutativeSemiring = isCommutativeSemiring } where open CommutativeMonoid open FP _↔_ left-zero : LeftZero (Lift ⊥) (_×_ {a = ℓ} {b = ℓ}) left-zero A = record { to = P.→-to-⟶ proj₁ ; from = P.→-to-⟶ (⊥-elim ∘′ lower) ; inverse-of = record { left-inverse-of = λ p → ⊥-elim (lower $ proj₁ p) ; right-inverse-of = λ x → ⊥-elim (lower x) } } right-distrib : _×_ DistributesOverʳ _⊎_ right-distrib A B C = record { to = P.→-to-⟶ $ uncurry [ curry inj₁ , curry inj₂ ] ; from = P.→-to-⟶ from ; inverse-of = record { right-inverse-of = [ (λ _ → P.refl) , (λ _ → P.refl) ] ; left-inverse-of = uncurry [ (λ _ _ → P.refl) , (λ _ _ → P.refl) ] } } where from : B × A ⊎ C × A → (B ⊎ C) × A from = [ Prod.map inj₁ id , Prod.map inj₂ id ] abstract -- If isCommutativeSemiring is made concrete, then it takes much -- more time to type-check coefficient-dec (at the time of -- writing, on a given system, using certain Agda options). isCommutativeSemiring : IsCommutativeSemiring {ℓ = ℓ} (Related ⌊ k ⌋) _⊎_ _×_ (Lift ⊥) (Lift ⊤) isCommutativeSemiring = record { +-isCommutativeMonoid = isCommutativeMonoid $ ⊎-CommutativeMonoid k ℓ ; *-isCommutativeMonoid = isCommutativeMonoid $ ×-CommutativeMonoid k ℓ ; distribʳ = λ A B C → ↔⇒ $ right-distrib A B C ; zeroˡ = λ A → ↔⇒ $ left-zero A } private -- A decision procedure used by the solver below. coefficient-dec : ∀ s ℓ → let open CommutativeSemiring (×⊎-CommutativeSemiring s ℓ) open Algebra.Operations semiring renaming (_×_ to Times) in ∀ m n → Dec (Times m 1# ∼[ ⌊ s ⌋ ] Times n 1#) coefficient-dec equivalence ℓ m n with m | n ... | zero | zero = yes (Eq.equivalence id id) ... | zero | suc _ = no (λ eq → lower (Equivalence.from eq ⟨$⟩ inj₁ _)) ... | suc _ | zero = no (λ eq → lower (Equivalence.to eq ⟨$⟩ inj₁ _)) ... | suc _ | suc _ = yes (Eq.equivalence (λ _ → inj₁ _) (λ _ → inj₁ _)) coefficient-dec bijection ℓ m n = Dec.map′ to (from m n) (Nat._≟_ m n) where open CommutativeSemiring (×⊎-CommutativeSemiring bijection ℓ) using (1#; semiring) open Algebra.Operations semiring renaming (_×_ to Times) to : ∀ {m n} → m ≡ n → Times m 1# ↔ Times n 1# to {m} P.refl = Times m 1# ∎ where open Related.EquationalReasoning from : ∀ m n → Times m 1# ↔ Times n 1# → m ≡ n from zero zero _ = P.refl from zero (suc n) 0↔+ = ⊥-elim $ lower $ Inverse.from 0↔+ ⟨$⟩ inj₁ _ from (suc m) zero +↔0 = ⊥-elim $ lower $ Inverse.to +↔0 ⟨$⟩ inj₁ _ from (suc m) (suc n) +↔+ = P.cong suc $ from m n (pred↔pred +↔+) where open P.≡-Reasoning ↑⊤ : Set ℓ ↑⊤ = Lift ⊤ inj₁≢inj₂ : ∀ {A : Set ℓ} {x : ↑⊤ ⊎ A} {y} → x ≡ inj₂ y → x ≡ inj₁ _ → ⊥ inj₁≢inj₂ {x = x} {y} eq₁ eq₂ = P.subst [ const ⊥ , const ⊤ ] (begin inj₂ y ≡⟨ P.sym eq₁ ⟩ x ≡⟨ eq₂ ⟩ inj₁ _ ∎) _ g′ : {A B : Set ℓ} (f : (↑⊤ ⊎ A) ↔ (↑⊤ ⊎ B)) (x : A) (y z : ↑⊤ ⊎ B) → Inverse.to f ⟨$⟩ inj₂ x ≡ y → Inverse.to f ⟨$⟩ inj₁ _ ≡ z → B g′ _ _ (inj₂ y) _ _ _ = y g′ _ _ (inj₁ _) (inj₂ z) _ _ = z g′ f _ (inj₁ _) (inj₁ _) eq₁ eq₂ = ⊥-elim $ inj₁≢inj₂ (Inverse.to-from f eq₁) (Inverse.to-from f eq₂) g : {A B : Set ℓ} → (↑⊤ ⊎ A) ↔ (↑⊤ ⊎ B) → A → B g f x = g′ f x _ _ P.refl P.refl g′∘g′ : ∀ {A B} (f : (↑⊤ ⊎ A) ↔ (↑⊤ ⊎ B)) x y₁ z₁ y₂ z₂ eq₁₁ eq₂₁ eq₁₂ eq₂₂ → g′ (reverse f) (g′ f x y₁ z₁ eq₁₁ eq₂₁) y₂ z₂ eq₁₂ eq₂₂ ≡ x g′∘g′ f x (inj₂ y₁) _ (inj₂ y₂) _ eq₁₁ _ eq₁₂ _ = P.cong [ const y₂ , id ] (begin inj₂ y₂ ≡⟨ P.sym eq₁₂ ⟩ Inverse.from f ⟨$⟩ inj₂ y₁ ≡⟨ Inverse.to-from f eq₁₁ ⟩ inj₂ x ∎) g′∘g′ f x (inj₁ _) (inj₂ _) (inj₁ _) (inj₂ z₂) eq₁₁ _ _ eq₂₂ = P.cong [ const z₂ , id ] (begin inj₂ z₂ ≡⟨ P.sym eq₂₂ ⟩ Inverse.from f ⟨$⟩ inj₁ _ ≡⟨ Inverse.to-from f eq₁₁ ⟩ inj₂ x ∎) g′∘g′ f _ (inj₂ y₁) _ (inj₁ _) _ eq₁₁ _ eq₁₂ _ = ⊥-elim $ inj₁≢inj₂ (Inverse.to-from f eq₁₁) eq₁₂ g′∘g′ f _ (inj₁ _) (inj₂ z₁) (inj₂ y₂) _ _ eq₂₁ eq₁₂ _ = ⊥-elim $ inj₁≢inj₂ eq₁₂ (Inverse.to-from f eq₂₁) g′∘g′ f _ (inj₁ _) (inj₂ _) (inj₁ _) (inj₁ _) eq₁₁ _ _ eq₂₂ = ⊥-elim $ inj₁≢inj₂ (Inverse.to-from f eq₁₁) eq₂₂ g′∘g′ f _ (inj₁ _) (inj₁ _) _ _ eq₁₁ eq₂₁ _ _ = ⊥-elim $ inj₁≢inj₂ (Inverse.to-from f eq₁₁) (Inverse.to-from f eq₂₁) g∘g : ∀ {A B} (f : (↑⊤ ⊎ A) ↔ (↑⊤ ⊎ B)) x → g (reverse f) (g f x) ≡ x g∘g f x = g′∘g′ f x _ _ _ _ P.refl P.refl P.refl P.refl pred↔pred : {A B : Set ℓ} → (↑⊤ ⊎ A) ↔ (↑⊤ ⊎ B) → A ↔ B pred↔pred X⊎↔X⊎ = record { to = P.→-to-⟶ $ g X⊎↔X⊎ ; from = P.→-to-⟶ $ g (reverse X⊎↔X⊎) ; inverse-of = record { left-inverse-of = g∘g X⊎↔X⊎ ; right-inverse-of = g∘g (reverse X⊎↔X⊎) } } module Solver s {ℓ} = Algebra.RingSolver.Natural-coefficients (×⊎-CommutativeSemiring s ℓ) (coefficient-dec s ℓ) private -- A test of the solver above. test : (A B C : Set) → (Lift ⊤ × A × (B ⊎ C)) ↔ (A × B ⊎ C × (Lift ⊥ ⊎ A)) test = solve 3 (λ A B C → con 1 :* (A :* (B :+ C)) := A :* B :+ C :* (con 0 :+ A)) Inv.id where open Solver bijection ------------------------------------------------------------------------ -- Some reordering lemmas ΠΠ↔ΠΠ : ∀ {a b p} {A : Set a} {B : Set b} (P : A → B → Set p) → ((x : A) (y : B) → P x y) ↔ ((y : B) (x : A) → P x y) ΠΠ↔ΠΠ _ = record { to = P.→-to-⟶ λ f x y → f y x ; from = P.→-to-⟶ λ f y x → f x y ; inverse-of = record { left-inverse-of = λ _ → P.refl ; right-inverse-of = λ _ → P.refl } } ∃∃↔∃∃ : ∀ {a b p} {A : Set a} {B : Set b} (P : A → B → Set p) → (∃₂ λ x y → P x y) ↔ (∃₂ λ y x → P x y) ∃∃↔∃∃ {a} {b} {p} _ = record { to = P.→-to-⟶ λ p → (proj₁ (proj₂ p) , proj₁ p , proj₂ (proj₂ p)) ; from = P.→-to-⟶ λ p → (proj₁ (proj₂ p) , proj₁ p , proj₂ (proj₂ p)) ; inverse-of = record { left-inverse-of = λ _ → P.refl ; right-inverse-of = λ _ → P.refl } } ------------------------------------------------------------------------ -- Implicit and explicit function spaces are isomorphic Π↔Π : ∀ {a b} {A : Set a} {B : A → Set b} → ((x : A) → B x) ↔ ({x : A} → B x) Π↔Π = record { to = P.→-to-⟶ λ f {x} → f x ; from = P.→-to-⟶ λ f x → f {x} ; inverse-of = record { left-inverse-of = λ _ → P.refl ; right-inverse-of = λ _ → P.refl } } ------------------------------------------------------------------------ -- _→_ preserves the symmetric relations _→-cong-⇔_ : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → A ⇔ B → C ⇔ D → (A → C) ⇔ (B → D) A⇔B →-cong-⇔ C⇔D = record { to = P.→-to-⟶ λ f x → Equivalence.to C⇔D ⟨$⟩ f (Equivalence.from A⇔B ⟨$⟩ x) ; from = P.→-to-⟶ λ f x → Equivalence.from C⇔D ⟨$⟩ f (Equivalence.to A⇔B ⟨$⟩ x) } →-cong : ∀ {a b c d} → P.Extensionality a c → P.Extensionality b d → ∀ {k} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → A ∼[ ⌊ k ⌋ ] B → C ∼[ ⌊ k ⌋ ] D → (A → C) ∼[ ⌊ k ⌋ ] (B → D) →-cong extAC extBD {equivalence} A⇔B C⇔D = A⇔B →-cong-⇔ C⇔D →-cong extAC extBD {bijection} A↔B C↔D = record { to = Equivalence.to A→C⇔B→D ; from = Equivalence.from A→C⇔B→D ; inverse-of = record { left-inverse-of = λ f → extAC λ x → begin Inverse.from C↔D ⟨$⟩ (Inverse.to C↔D ⟨$⟩ f (Inverse.from A↔B ⟨$⟩ (Inverse.to A↔B ⟨$⟩ x))) ≡⟨ Inverse.left-inverse-of C↔D _ ⟩ f (Inverse.from A↔B ⟨$⟩ (Inverse.to A↔B ⟨$⟩ x)) ≡⟨ P.cong f $ Inverse.left-inverse-of A↔B x ⟩ f x ∎ ; right-inverse-of = λ f → extBD λ x → begin Inverse.to C↔D ⟨$⟩ (Inverse.from C↔D ⟨$⟩ f (Inverse.to A↔B ⟨$⟩ (Inverse.from A↔B ⟨$⟩ x))) ≡⟨ Inverse.right-inverse-of C↔D _ ⟩ f (Inverse.to A↔B ⟨$⟩ (Inverse.from A↔B ⟨$⟩ x)) ≡⟨ P.cong f $ Inverse.right-inverse-of A↔B x ⟩ f x ∎ } } where open P.≡-Reasoning A→C⇔B→D = ↔⇒ A↔B →-cong-⇔ ↔⇒ C↔D ------------------------------------------------------------------------ -- ¬_ preserves the symmetric relations ¬-cong-⇔ : ∀ {a b} {A : Set a} {B : Set b} → A ⇔ B → (¬ A) ⇔ (¬ B) ¬-cong-⇔ A⇔B = A⇔B →-cong-⇔ (⊥ ∎) where open Related.EquationalReasoning ¬-cong : ∀ {a b} → P.Extensionality a Level.zero → P.Extensionality b Level.zero → ∀ {k} {A : Set a} {B : Set b} → A ∼[ ⌊ k ⌋ ] B → (¬ A) ∼[ ⌊ k ⌋ ] (¬ B) ¬-cong extA extB A≈B = →-cong extA extB A≈B (⊥ ∎) where open Related.EquationalReasoning ------------------------------------------------------------------------ -- _⇔_ preserves _⇔_ -- The type of the following proof is a bit more general. Related-cong : ∀ {k a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → A ∼[ ⌊ k ⌋ ] B → C ∼[ ⌊ k ⌋ ] D → (A ∼[ ⌊ k ⌋ ] C) ⇔ (B ∼[ ⌊ k ⌋ ] D) Related-cong {A = A} {B} {C} {D} A≈B C≈D = Eq.equivalence (λ A≈C → B ∼⟨ sym A≈B ⟩ A ∼⟨ A≈C ⟩ C ∼⟨ C≈D ⟩ D ∎) (λ B≈D → A ∼⟨ A≈B ⟩ B ∼⟨ B≈D ⟩ D ∼⟨ sym C≈D ⟩ C ∎) where open Related.EquationalReasoning lib-0.7/src/Relation/0000755000000000000000000000000012101774706012665 5ustar0000000000000000lib-0.7/src/Relation/Nullary.agda0000644000000000000000000000117212101774706015132 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Operations on nullary relations (like negation and decidability) ------------------------------------------------------------------------ -- Some operations on/properties of nullary relations, i.e. sets. module Relation.Nullary where import Relation.Nullary.Core as Core ------------------------------------------------------------------------ -- Negation open Core public using (¬_) ------------------------------------------------------------------------ -- Decidable relations open Core public using (Dec; yes; no) lib-0.7/src/Relation/Binary.agda0000644000000000000000000002466512101774706014744 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Properties of homogeneous binary relations ------------------------------------------------------------------------ module Relation.Binary where open import Data.Product open import Data.Sum open import Function open import Level import Relation.Binary.PropositionalEquality.Core as PropEq open import Relation.Binary.Consequences open import Relation.Binary.Core as Core using (_≡_) import Relation.Binary.Indexed.Core as I ------------------------------------------------------------------------ -- Simple properties and equivalence relations open Core public hiding (_≡_; refl; _≢_) ------------------------------------------------------------------------ -- Preorders record IsPreorder {a ℓ₁ ℓ₂} {A : Set a} (_≈_ : Rel A ℓ₁) -- The underlying equality. (_∼_ : Rel A ℓ₂) -- The relation. : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where field isEquivalence : IsEquivalence _≈_ -- Reflexivity is expressed in terms of an underlying equality: reflexive : _≈_ ⇒ _∼_ trans : Transitive _∼_ module Eq = IsEquivalence isEquivalence refl : Reflexive _∼_ refl = reflexive Eq.refl ∼-resp-≈ : _∼_ Respects₂ _≈_ ∼-resp-≈ = (λ x≈y z∼x → trans z∼x (reflexive x≈y)) , (λ x≈y x∼z → trans (reflexive $ Eq.sym x≈y) x∼z) record Preorder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where infix 4 _≈_ _∼_ field Carrier : Set c _≈_ : Rel Carrier ℓ₁ -- The underlying equality. _∼_ : Rel Carrier ℓ₂ -- The relation. isPreorder : IsPreorder _≈_ _∼_ open IsPreorder isPreorder public ------------------------------------------------------------------------ -- Setoids -- Equivalence relations are defined in Relation.Binary.Core. record Setoid c ℓ : Set (suc (c ⊔ ℓ)) where infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ isEquivalence : IsEquivalence _≈_ open IsEquivalence isEquivalence public isPreorder : IsPreorder _≡_ _≈_ isPreorder = record { isEquivalence = PropEq.isEquivalence ; reflexive = reflexive ; trans = trans } preorder : Preorder c c ℓ preorder = record { isPreorder = isPreorder } -- A trivially indexed setoid. indexedSetoid : ∀ {i} {I : Set i} → I.Setoid I c _ indexedSetoid = record { Carrier = λ _ → Carrier ; _≈_ = _≈_ ; isEquivalence = record { refl = refl ; sym = sym ; trans = trans } } ------------------------------------------------------------------------ -- Decidable equivalence relations record IsDecEquivalence {a ℓ} {A : Set a} (_≈_ : Rel A ℓ) : Set (a ⊔ ℓ) where infix 4 _≟_ field isEquivalence : IsEquivalence _≈_ _≟_ : Decidable _≈_ open IsEquivalence isEquivalence public record DecSetoid c ℓ : Set (suc (c ⊔ ℓ)) where infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ isDecEquivalence : IsDecEquivalence _≈_ open IsDecEquivalence isDecEquivalence public setoid : Setoid c ℓ setoid = record { isEquivalence = isEquivalence } open Setoid setoid public using (preorder) ------------------------------------------------------------------------ -- Partial orders record IsPartialOrder {a ℓ₁ ℓ₂} {A : Set a} (_≈_ : Rel A ℓ₁) (_≤_ : Rel A ℓ₂) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where field isPreorder : IsPreorder _≈_ _≤_ antisym : Antisymmetric _≈_ _≤_ open IsPreorder isPreorder public renaming (∼-resp-≈ to ≤-resp-≈) record Poset c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where infix 4 _≈_ _≤_ field Carrier : Set c _≈_ : Rel Carrier ℓ₁ _≤_ : Rel Carrier ℓ₂ isPartialOrder : IsPartialOrder _≈_ _≤_ open IsPartialOrder isPartialOrder public preorder : Preorder c ℓ₁ ℓ₂ preorder = record { isPreorder = isPreorder } ------------------------------------------------------------------------ -- Decidable partial orders record IsDecPartialOrder {a ℓ₁ ℓ₂} {A : Set a} (_≈_ : Rel A ℓ₁) (_≤_ : Rel A ℓ₂) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where infix 4 _≟_ _≤?_ field isPartialOrder : IsPartialOrder _≈_ _≤_ _≟_ : Decidable _≈_ _≤?_ : Decidable _≤_ private module PO = IsPartialOrder isPartialOrder open PO public hiding (module Eq) module Eq where isDecEquivalence : IsDecEquivalence _≈_ isDecEquivalence = record { isEquivalence = PO.isEquivalence ; _≟_ = _≟_ } open IsDecEquivalence isDecEquivalence public record DecPoset c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where infix 4 _≈_ _≤_ field Carrier : Set c _≈_ : Rel Carrier ℓ₁ _≤_ : Rel Carrier ℓ₂ isDecPartialOrder : IsDecPartialOrder _≈_ _≤_ private module DPO = IsDecPartialOrder isDecPartialOrder open DPO public hiding (module Eq) poset : Poset c ℓ₁ ℓ₂ poset = record { isPartialOrder = isPartialOrder } open Poset poset public using (preorder) module Eq where decSetoid : DecSetoid c ℓ₁ decSetoid = record { isDecEquivalence = DPO.Eq.isDecEquivalence } open DecSetoid decSetoid public ------------------------------------------------------------------------ -- Strict partial orders record IsStrictPartialOrder {a ℓ₁ ℓ₂} {A : Set a} (_≈_ : Rel A ℓ₁) (_<_ : Rel A ℓ₂) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where field isEquivalence : IsEquivalence _≈_ irrefl : Irreflexive _≈_ _<_ trans : Transitive _<_ <-resp-≈ : _<_ Respects₂ _≈_ module Eq = IsEquivalence isEquivalence record StrictPartialOrder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where infix 4 _≈_ _<_ field Carrier : Set c _≈_ : Rel Carrier ℓ₁ _<_ : Rel Carrier ℓ₂ isStrictPartialOrder : IsStrictPartialOrder _≈_ _<_ open IsStrictPartialOrder isStrictPartialOrder public ------------------------------------------------------------------------ -- Total orders record IsTotalOrder {a ℓ₁ ℓ₂} {A : Set a} (_≈_ : Rel A ℓ₁) (_≤_ : Rel A ℓ₂) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where field isPartialOrder : IsPartialOrder _≈_ _≤_ total : Total _≤_ open IsPartialOrder isPartialOrder public record TotalOrder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where infix 4 _≈_ _≤_ field Carrier : Set c _≈_ : Rel Carrier ℓ₁ _≤_ : Rel Carrier ℓ₂ isTotalOrder : IsTotalOrder _≈_ _≤_ open IsTotalOrder isTotalOrder public poset : Poset c ℓ₁ ℓ₂ poset = record { isPartialOrder = isPartialOrder } open Poset poset public using (preorder) ------------------------------------------------------------------------ -- Decidable total orders record IsDecTotalOrder {a ℓ₁ ℓ₂} {A : Set a} (_≈_ : Rel A ℓ₁) (_≤_ : Rel A ℓ₂) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where infix 4 _≟_ _≤?_ field isTotalOrder : IsTotalOrder _≈_ _≤_ _≟_ : Decidable _≈_ _≤?_ : Decidable _≤_ private module TO = IsTotalOrder isTotalOrder open TO public hiding (module Eq) module Eq where isDecEquivalence : IsDecEquivalence _≈_ isDecEquivalence = record { isEquivalence = TO.isEquivalence ; _≟_ = _≟_ } open IsDecEquivalence isDecEquivalence public record DecTotalOrder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where infix 4 _≈_ _≤_ field Carrier : Set c _≈_ : Rel Carrier ℓ₁ _≤_ : Rel Carrier ℓ₂ isDecTotalOrder : IsDecTotalOrder _≈_ _≤_ private module DTO = IsDecTotalOrder isDecTotalOrder open DTO public hiding (module Eq) totalOrder : TotalOrder c ℓ₁ ℓ₂ totalOrder = record { isTotalOrder = isTotalOrder } open TotalOrder totalOrder public using (poset; preorder) module Eq where decSetoid : DecSetoid c ℓ₁ decSetoid = record { isDecEquivalence = DTO.Eq.isDecEquivalence } open DecSetoid decSetoid public ------------------------------------------------------------------------ -- Strict total orders -- Note that these orders are decidable (see compare). record IsStrictTotalOrder {a ℓ₁ ℓ₂} {A : Set a} (_≈_ : Rel A ℓ₁) (_<_ : Rel A ℓ₂) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where field isEquivalence : IsEquivalence _≈_ trans : Transitive _<_ compare : Trichotomous _≈_ _<_ <-resp-≈ : _<_ Respects₂ _≈_ infix 4 _≟_ _ (λ()) (λ()) (₁∼₂ _) tri (inj₁ x) (inj₁ y) with tri₁ x y ... | tri< x x≮y x≉y x>y = tri> (x≮y ∘ drop-inj₁) (x≉y ∘ drop-inj₁) (₁∼₁ x>y) tri (inj₂ x) (inj₂ y) with tri₂ x y ... | tri< x x≮y x≉y x>y = tri> (x≮y ∘ drop-inj₂) (x≉y ∘ drop-inj₂) (₂∼₂ x>y) ---------------------------------------------------------------------- -- Some collections of properties which are preserved _⊎-isEquivalence_ : ∀ {ℓ₁ ℓ₂} {≈₁ : Rel A₁ ℓ₁} {≈₂ : Rel A₂ ℓ₂} → IsEquivalence ≈₁ → IsEquivalence ≈₂ → IsEquivalence (≈₁ ⊎-Rel ≈₂) eq₁ ⊎-isEquivalence eq₂ = record { refl = refl eq₁ ⊎-refl refl eq₂ ; sym = sym eq₁ ⊎-symmetric sym eq₂ ; trans = trans eq₁ ⊎-transitive trans eq₂ } where open IsEquivalence _⊎-isPreorder_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {∼₁ : Rel A₁ ℓ₁′} {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {∼₂ : Rel A₂ ℓ₂′} → IsPreorder ≈₁ ∼₁ → IsPreorder ≈₂ ∼₂ → ∀ {P} → IsPreorder (≈₁ ⊎-Rel ≈₂) (⊎ʳ P ∼₁ ∼₂) pre₁ ⊎-isPreorder pre₂ = record { isEquivalence = isEquivalence pre₁ ⊎-isEquivalence isEquivalence pre₂ ; reflexive = reflexive pre₁ ⊎-reflexive reflexive pre₂ ; trans = trans pre₁ ⊎-transitive trans pre₂ } where open IsPreorder _⊎-isDecEquivalence_ : ∀ {ℓ₁ ℓ₂} {≈₁ : Rel A₁ ℓ₁} {≈₂ : Rel A₂ ℓ₂} → IsDecEquivalence ≈₁ → IsDecEquivalence ≈₂ → IsDecEquivalence (≈₁ ⊎-Rel ≈₂) eq₁ ⊎-isDecEquivalence eq₂ = record { isEquivalence = isEquivalence eq₁ ⊎-isEquivalence isEquivalence eq₂ ; _≟_ = ⊎-decidable (_≟_ eq₁) (_≟_ eq₂) (no ₁≁₂) } where open IsDecEquivalence _⊎-isPartialOrder_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {≤₁ : Rel A₁ ℓ₁′} {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {≤₂ : Rel A₂ ℓ₂′} → IsPartialOrder ≈₁ ≤₁ → IsPartialOrder ≈₂ ≤₂ → ∀ {P} → IsPartialOrder (≈₁ ⊎-Rel ≈₂) (⊎ʳ P ≤₁ ≤₂) po₁ ⊎-isPartialOrder po₂ = record { isPreorder = isPreorder po₁ ⊎-isPreorder isPreorder po₂ ; antisym = antisym po₁ ⊎-antisymmetric antisym po₂ } where open IsPartialOrder _⊎-isStrictPartialOrder_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {<₁ : Rel A₁ ℓ₁′} {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {<₂ : Rel A₂ ℓ₂′} → IsStrictPartialOrder ≈₁ <₁ → IsStrictPartialOrder ≈₂ <₂ → ∀ {P} → IsStrictPartialOrder (≈₁ ⊎-Rel ≈₂) (⊎ʳ P <₁ <₂) spo₁ ⊎-isStrictPartialOrder spo₂ = record { isEquivalence = isEquivalence spo₁ ⊎-isEquivalence isEquivalence spo₂ ; irrefl = irrefl spo₁ ⊎-irreflexive irrefl spo₂ ; trans = trans spo₁ ⊎-transitive trans spo₂ ; <-resp-≈ = <-resp-≈ spo₁ ⊎-≈-respects₂ <-resp-≈ spo₂ } where open IsStrictPartialOrder _⊎-<-isTotalOrder_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {≤₁ : Rel A₁ ℓ₁′} {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {≤₂ : Rel A₂ ℓ₂′} → IsTotalOrder ≈₁ ≤₁ → IsTotalOrder ≈₂ ≤₂ → IsTotalOrder (≈₁ ⊎-Rel ≈₂) (≤₁ ⊎-< ≤₂) to₁ ⊎-<-isTotalOrder to₂ = record { isPartialOrder = isPartialOrder to₁ ⊎-isPartialOrder isPartialOrder to₂ ; total = total to₁ ⊎-<-total total to₂ } where open IsTotalOrder _⊎-<-isDecTotalOrder_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {≤₁ : Rel A₁ ℓ₁′} {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {≤₂ : Rel A₂ ℓ₂′} → IsDecTotalOrder ≈₁ ≤₁ → IsDecTotalOrder ≈₂ ≤₂ → IsDecTotalOrder (≈₁ ⊎-Rel ≈₂) (≤₁ ⊎-< ≤₂) to₁ ⊎-<-isDecTotalOrder to₂ = record { isTotalOrder = isTotalOrder to₁ ⊎-<-isTotalOrder isTotalOrder to₂ ; _≟_ = ⊎-decidable (_≟_ to₁) (_≟_ to₂) (no ₁≁₂) ; _≤?_ = ⊎-decidable (_≤?_ to₁) (_≤?_ to₂) (yes (₁∼₂ _)) } where open IsDecTotalOrder _⊎-<-isStrictTotalOrder_ : ∀ {ℓ₁ ℓ₁′} {≈₁ : Rel A₁ ℓ₁} {<₁ : Rel A₁ ℓ₁′} {ℓ₂ ℓ₂′} {≈₂ : Rel A₂ ℓ₂} {<₂ : Rel A₂ ℓ₂′} → IsStrictTotalOrder ≈₁ <₁ → IsStrictTotalOrder ≈₂ <₂ → IsStrictTotalOrder (≈₁ ⊎-Rel ≈₂) (<₁ ⊎-< <₂) sto₁ ⊎-<-isStrictTotalOrder sto₂ = record { isEquivalence = isEquivalence sto₁ ⊎-isEquivalence isEquivalence sto₂ ; trans = trans sto₁ ⊎-transitive trans sto₂ ; compare = compare sto₁ ⊎-<-trichotomous compare sto₂ ; <-resp-≈ = <-resp-≈ sto₁ ⊎-≈-respects₂ <-resp-≈ sto₂ } where open IsStrictTotalOrder ------------------------------------------------------------------------ -- The game can be taken even further... _⊎-setoid_ : ∀ {s₁ s₂ s₃ s₄} → Setoid s₁ s₂ → Setoid s₃ s₄ → Setoid _ _ s₁ ⊎-setoid s₂ = record { isEquivalence = isEquivalence s₁ ⊎-isEquivalence isEquivalence s₂ } where open Setoid _⊎-preorder_ : ∀ {p₁ p₂ p₃ p₄ p₅ p₆} → Preorder p₁ p₂ p₃ → Preorder p₄ p₅ p₆ → Preorder _ _ _ p₁ ⊎-preorder p₂ = record { _∼_ = _∼_ p₁ ⊎-Rel _∼_ p₂ ; isPreorder = isPreorder p₁ ⊎-isPreorder isPreorder p₂ } where open Preorder _⊎-decSetoid_ : ∀ {s₁ s₂ s₃ s₄} → DecSetoid s₁ s₂ → DecSetoid s₃ s₄ → DecSetoid _ _ ds₁ ⊎-decSetoid ds₂ = record { isDecEquivalence = isDecEquivalence ds₁ ⊎-isDecEquivalence isDecEquivalence ds₂ } where open DecSetoid _⊎-poset_ : ∀ {p₁ p₂ p₃ p₄ p₅ p₆} → Poset p₁ p₂ p₃ → Poset p₄ p₅ p₆ → Poset _ _ _ po₁ ⊎-poset po₂ = record { _≤_ = _≤_ po₁ ⊎-Rel _≤_ po₂ ; isPartialOrder = isPartialOrder po₁ ⊎-isPartialOrder isPartialOrder po₂ } where open Poset _⊎-<-poset_ : ∀ {p₁ p₂ p₃ p₄ p₅ p₆} → Poset p₁ p₂ p₃ → Poset p₄ p₅ p₆ → Poset _ _ _ po₁ ⊎-<-poset po₂ = record { _≤_ = _≤_ po₁ ⊎-< _≤_ po₂ ; isPartialOrder = isPartialOrder po₁ ⊎-isPartialOrder isPartialOrder po₂ } where open Poset _⊎-<-strictPartialOrder_ : ∀ {p₁ p₂ p₃ p₄ p₅ p₆} → StrictPartialOrder p₁ p₂ p₃ → StrictPartialOrder p₄ p₅ p₆ → StrictPartialOrder _ _ _ spo₁ ⊎-<-strictPartialOrder spo₂ = record { _<_ = _<_ spo₁ ⊎-< _<_ spo₂ ; isStrictPartialOrder = isStrictPartialOrder spo₁ ⊎-isStrictPartialOrder isStrictPartialOrder spo₂ } where open StrictPartialOrder _⊎-<-totalOrder_ : ∀ {t₁ t₂ t₃ t₄ t₅ t₆} → TotalOrder t₁ t₂ t₃ → TotalOrder t₄ t₅ t₆ → TotalOrder _ _ _ to₁ ⊎-<-totalOrder to₂ = record { isTotalOrder = isTotalOrder to₁ ⊎-<-isTotalOrder isTotalOrder to₂ } where open TotalOrder _⊎-<-decTotalOrder_ : ∀ {t₁ t₂ t₃ t₄ t₅ t₆} → DecTotalOrder t₁ t₂ t₃ → DecTotalOrder t₄ t₅ t₆ → DecTotalOrder _ _ _ to₁ ⊎-<-decTotalOrder to₂ = record { isDecTotalOrder = isDecTotalOrder to₁ ⊎-<-isDecTotalOrder isDecTotalOrder to₂ } where open DecTotalOrder _⊎-<-strictTotalOrder_ : ∀ {p₁ p₂ p₃ p₄ p₅ p₆} → StrictTotalOrder p₁ p₂ p₃ → StrictTotalOrder p₄ p₅ p₆ → StrictTotalOrder _ _ _ sto₁ ⊎-<-strictTotalOrder sto₂ = record { _<_ = _<_ sto₁ ⊎-< _<_ sto₂ ; isStrictTotalOrder = isStrictTotalOrder sto₁ ⊎-<-isStrictTotalOrder isStrictTotalOrder sto₂ } where open StrictTotalOrder ------------------------------------------------------------------------ -- Some properties related to "relatedness" ⊎-Rel↔≡ : ∀ {a b} (A : Set a) (B : Set b) → Inverse (P.setoid A ⊎-setoid P.setoid B) (P.setoid (A ⊎ B)) ⊎-Rel↔≡ _ _ = record { to = record { _⟨$⟩_ = id; cong = to-cong } ; from = record { _⟨$⟩_ = id; cong = from-cong } ; inverse-of = record { left-inverse-of = λ _ → P.refl ⊎-refl P.refl ; right-inverse-of = λ _ → P.refl } } where to-cong : (P._≡_ ⊎-Rel P._≡_) ⇒ P._≡_ to-cong (₁∼₂ ()) to-cong (₁∼₁ P.refl) = P.refl to-cong (₂∼₂ P.refl) = P.refl from-cong : P._≡_ ⇒ (P._≡_ ⊎-Rel P._≡_) from-cong P.refl = P.refl ⊎-refl P.refl _⊎-⟶_ : ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈} {A : Setoid s₁ s₂} {B : Setoid s₃ s₄} {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} → A ⟶ B → C ⟶ D → (A ⊎-setoid C) ⟶ (B ⊎-setoid D) _⊎-⟶_ {A = A} {B} {C} {D} f g = record { _⟨$⟩_ = fg ; cong = fg-cong } where open Setoid (A ⊎-setoid C) using () renaming (_≈_ to _≈AC_) open Setoid (B ⊎-setoid D) using () renaming (_≈_ to _≈BD_) fg = Sum.map (_⟨$⟩_ f) (_⟨$⟩_ g) fg-cong : _≈AC_ =[ fg ]⇒ _≈BD_ fg-cong (₁∼₂ ()) fg-cong (₁∼₁ x∼₁y) = ₁∼₁ $ F.cong f x∼₁y fg-cong (₂∼₂ x∼₂y) = ₂∼₂ $ F.cong g x∼₂y _⊎-equivalence_ : ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈} {A : Setoid s₁ s₂} {B : Setoid s₃ s₄} {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} → Equivalence A B → Equivalence C D → Equivalence (A ⊎-setoid C) (B ⊎-setoid D) A⇔B ⊎-equivalence C⇔D = record { to = to A⇔B ⊎-⟶ to C⇔D ; from = from A⇔B ⊎-⟶ from C⇔D } where open Equivalence _⊎-⇔_ : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → A ⇔ B → C ⇔ D → (A ⊎ C) ⇔ (B ⊎ D) _⊎-⇔_ {A = A} {B} {C} {D} A⇔B C⇔D = Inverse.equivalence (⊎-Rel↔≡ B D) ⟨∘⟩ A⇔B ⊎-equivalence C⇔D ⟨∘⟩ Eq.sym (Inverse.equivalence (⊎-Rel↔≡ A C)) where open Eq using () renaming (_∘_ to _⟨∘⟩_) _⊎-injection_ : ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈} {A : Setoid s₁ s₂} {B : Setoid s₃ s₄} {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} → Injection A B → Injection C D → Injection (A ⊎-setoid C) (B ⊎-setoid D) _⊎-injection_ {A = A} {B} {C} {D} A↣B C↣D = record { to = to A↣B ⊎-⟶ to C↣D ; injective = inj _ _ } where open Injection open Setoid (A ⊎-setoid C) using () renaming (_≈_ to _≈AC_) open Setoid (B ⊎-setoid D) using () renaming (_≈_ to _≈BD_) inj : ∀ x y → (to A↣B ⊎-⟶ to C↣D) ⟨$⟩ x ≈BD (to A↣B ⊎-⟶ to C↣D) ⟨$⟩ y → x ≈AC y inj (inj₁ x) (inj₁ y) (₁∼₁ x∼₁y) = ₁∼₁ (injective A↣B x∼₁y) inj (inj₂ x) (inj₂ y) (₂∼₂ x∼₂y) = ₂∼₂ (injective C↣D x∼₂y) inj (inj₁ x) (inj₂ y) (₁∼₂ ()) inj (inj₂ x) (inj₁ y) () _⊎-↣_ : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → A ↣ B → C ↣ D → (A ⊎ C) ↣ (B ⊎ D) _⊎-↣_ {A = A} {B} {C} {D} A↣B C↣D = Inverse.injection (⊎-Rel↔≡ B D) ⟨∘⟩ A↣B ⊎-injection C↣D ⟨∘⟩ Inverse.injection (Inv.sym (⊎-Rel↔≡ A C)) where open Inj using () renaming (_∘_ to _⟨∘⟩_) _⊎-left-inverse_ : ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈} {A : Setoid s₁ s₂} {B : Setoid s₃ s₄} {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} → LeftInverse A B → LeftInverse C D → LeftInverse (A ⊎-setoid C) (B ⊎-setoid D) A↞B ⊎-left-inverse C↞D = record { to = Equivalence.to eq ; from = Equivalence.from eq ; left-inverse-of = [ ₁∼₁ ∘ left-inverse-of A↞B , ₂∼₂ ∘ left-inverse-of C↞D ] } where open LeftInverse eq = LeftInverse.equivalence A↞B ⊎-equivalence LeftInverse.equivalence C↞D _⊎-↞_ : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → A ↞ B → C ↞ D → (A ⊎ C) ↞ (B ⊎ D) _⊎-↞_ {A = A} {B} {C} {D} A↞B C↞D = Inverse.left-inverse (⊎-Rel↔≡ B D) ⟨∘⟩ A↞B ⊎-left-inverse C↞D ⟨∘⟩ Inverse.left-inverse (Inv.sym (⊎-Rel↔≡ A C)) where open LeftInv using () renaming (_∘_ to _⟨∘⟩_) _⊎-surjection_ : ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈} {A : Setoid s₁ s₂} {B : Setoid s₃ s₄} {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} → Surjection A B → Surjection C D → Surjection (A ⊎-setoid C) (B ⊎-setoid D) A↠B ⊎-surjection C↠D = record { to = LeftInverse.from inv ; surjective = record { from = LeftInverse.to inv ; right-inverse-of = LeftInverse.left-inverse-of inv } } where open Surjection inv = right-inverse A↠B ⊎-left-inverse right-inverse C↠D _⊎-↠_ : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → A ↠ B → C ↠ D → (A ⊎ C) ↠ (B ⊎ D) _⊎-↠_ {A = A} {B} {C} {D} A↠B C↠D = Inverse.surjection (⊎-Rel↔≡ B D) ⟨∘⟩ A↠B ⊎-surjection C↠D ⟨∘⟩ Inverse.surjection (Inv.sym (⊎-Rel↔≡ A C)) where open Surj using () renaming (_∘_ to _⟨∘⟩_) _⊎-inverse_ : ∀ {s₁ s₂ s₃ s₄ s₅ s₆ s₇ s₈} {A : Setoid s₁ s₂} {B : Setoid s₃ s₄} {C : Setoid s₅ s₆} {D : Setoid s₇ s₈} → Inverse A B → Inverse C D → Inverse (A ⊎-setoid C) (B ⊎-setoid D) A↔B ⊎-inverse C↔D = record { to = Surjection.to surj ; from = Surjection.from surj ; inverse-of = record { left-inverse-of = LeftInverse.left-inverse-of inv ; right-inverse-of = Surjection.right-inverse-of surj } } where open Inverse surj = Inverse.surjection A↔B ⊎-surjection Inverse.surjection C↔D inv = Inverse.left-inverse A↔B ⊎-left-inverse Inverse.left-inverse C↔D _⊎-↔_ : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → A ↔ B → C ↔ D → (A ⊎ C) ↔ (B ⊎ D) _⊎-↔_ {A = A} {B} {C} {D} A↔B C↔D = ⊎-Rel↔≡ B D ⟨∘⟩ A↔B ⊎-inverse C↔D ⟨∘⟩ Inv.sym (⊎-Rel↔≡ A C) where open Inv using () renaming (_∘_ to _⟨∘⟩_) _⊎-cong_ : ∀ {k a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → A ∼[ k ] B → C ∼[ k ] D → (A ⊎ C) ∼[ k ] (B ⊎ D) _⊎-cong_ {implication} = Sum.map _⊎-cong_ {reverse-implication} = λ f g → lam (Sum.map (app-← f) (app-← g)) _⊎-cong_ {equivalence} = _⊎-⇔_ _⊎-cong_ {injection} = _⊎-↣_ _⊎-cong_ {reverse-injection} = λ f g → lam (app-↢ f ⊎-↣ app-↢ g) _⊎-cong_ {left-inverse} = _⊎-↞_ _⊎-cong_ {surjection} = _⊎-↠_ _⊎-cong_ {bijection} = _⊎-↔_ lib-0.7/src/Relation/Binary/Core.agda0000644000000000000000000001370312101774706015623 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Properties of homogeneous binary relations ------------------------------------------------------------------------ -- This file contains some core definitions which are reexported by -- Relation.Binary or Relation.Binary.PropositionalEquality. module Relation.Binary.Core where open import Data.Product open import Data.Sum open import Function open import Level open import Relation.Nullary ------------------------------------------------------------------------ -- Binary relations -- Heterogeneous binary relations REL : ∀ {a b} → Set a → Set b → (ℓ : Level) → Set (a ⊔ b ⊔ suc ℓ) REL A B ℓ = A → B → Set ℓ -- Homogeneous binary relations Rel : ∀ {a} → Set a → (ℓ : Level) → Set (a ⊔ suc ℓ) Rel A ℓ = REL A A ℓ ------------------------------------------------------------------------ -- Simple properties of binary relations infixr 4 _⇒_ _=[_]⇒_ -- Implication/containment. Could also be written ⊆. _⇒_ : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b} → REL A B ℓ₁ → REL A B ℓ₂ → Set _ P ⇒ Q = ∀ {i j} → P i j → Q i j -- Generalised implication. If P ≡ Q it can be read as "f preserves -- P". _=[_]⇒_ : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b} → Rel A ℓ₁ → (A → B) → Rel B ℓ₂ → Set _ P =[ f ]⇒ Q = P ⇒ (Q on f) -- A synonym, along with a binary variant. _Preserves_⟶_ : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b} → (A → B) → Rel A ℓ₁ → Rel B ℓ₂ → Set _ f Preserves P ⟶ Q = P =[ f ]⇒ Q _Preserves₂_⟶_⟶_ : ∀ {a b c ℓ₁ ℓ₂ ℓ₃} {A : Set a} {B : Set b} {C : Set c} → (A → B → C) → Rel A ℓ₁ → Rel B ℓ₂ → Rel C ℓ₃ → Set _ _+_ Preserves₂ P ⟶ Q ⟶ R = ∀ {x y u v} → P x y → Q u v → R (x + u) (y + v) -- Reflexivity of _∼_ can be expressed as _≈_ ⇒ _∼_, for some -- underlying equality _≈_. However, the following variant is often -- easier to use. Reflexive : ∀ {a ℓ} {A : Set a} → Rel A ℓ → Set _ Reflexive _∼_ = ∀ {x} → x ∼ x -- Irreflexivity is defined using an underlying equality. Irreflexive : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b} → REL A B ℓ₁ → REL A B ℓ₂ → Set _ Irreflexive _≈_ _<_ = ∀ {x y} → x ≈ y → ¬ (x < y) -- Generalised symmetry. Sym : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b} → REL A B ℓ₁ → REL B A ℓ₂ → Set _ Sym P Q = P ⇒ flip Q Symmetric : ∀ {a ℓ} {A : Set a} → Rel A ℓ → Set _ Symmetric _∼_ = Sym _∼_ _∼_ -- Generalised transitivity. Trans : ∀ {a b c ℓ₁ ℓ₂ ℓ₃} {A : Set a} {B : Set b} {C : Set c} → REL A B ℓ₁ → REL B C ℓ₂ → REL A C ℓ₃ → Set _ Trans P Q R = ∀ {i j k} → P i j → Q j k → R i k -- A variant of Trans. TransFlip : ∀ {a b c ℓ₁ ℓ₂ ℓ₃} {A : Set a} {B : Set b} {C : Set c} → REL A B ℓ₁ → REL B C ℓ₂ → REL A C ℓ₃ → Set _ TransFlip P Q R = ∀ {i j k} → Q j k → P i j → R i k Transitive : ∀ {a ℓ} {A : Set a} → Rel A ℓ → Set _ Transitive _∼_ = Trans _∼_ _∼_ _∼_ Antisymmetric : ∀ {a ℓ₁ ℓ₂} {A : Set a} → Rel A ℓ₁ → Rel A ℓ₂ → Set _ Antisymmetric _≈_ _≤_ = ∀ {x y} → x ≤ y → y ≤ x → x ≈ y Asymmetric : ∀ {a ℓ} {A : Set a} → Rel A ℓ → Set _ Asymmetric _<_ = ∀ {x y} → x < y → ¬ (y < x) _Respects_ : ∀ {a ℓ₁ ℓ₂} {A : Set a} → (A → Set ℓ₁) → Rel A ℓ₂ → Set _ P Respects _∼_ = ∀ {x y} → x ∼ y → P x → P y _Respects₂_ : ∀ {a ℓ₁ ℓ₂} {A : Set a} → Rel A ℓ₁ → Rel A ℓ₂ → Set _ P Respects₂ _∼_ = (∀ {x} → P x Respects _∼_) × (∀ {y} → flip P y Respects _∼_) Substitutive : ∀ {a ℓ₁} {A : Set a} → Rel A ℓ₁ → (ℓ₂ : Level) → Set _ Substitutive {A = A} _∼_ p = (P : A → Set p) → P Respects _∼_ Decidable : ∀ {a b ℓ} {A : Set a} {B : Set b} → REL A B ℓ → Set _ Decidable _∼_ = ∀ x y → Dec (x ∼ y) Total : ∀ {a ℓ} {A : Set a} → Rel A ℓ → Set _ Total _∼_ = ∀ x y → (x ∼ y) ⊎ (y ∼ x) data Tri {a b c} (A : Set a) (B : Set b) (C : Set c) : Set (a ⊔ b ⊔ c) where tri< : ( a : A) (¬b : ¬ B) (¬c : ¬ C) → Tri A B C tri≈ : (¬a : ¬ A) ( b : B) (¬c : ¬ C) → Tri A B C tri> : (¬a : ¬ A) (¬b : ¬ B) ( c : C) → Tri A B C Trichotomous : ∀ {a ℓ₁ ℓ₂} {A : Set a} → Rel A ℓ₁ → Rel A ℓ₂ → Set _ Trichotomous _≈_ _<_ = ∀ x y → Tri (x < y) (x ≈ y) (x > y) where _>_ = flip _<_ record NonEmpty {a b ℓ} {A : Set a} {B : Set b} (T : REL A B ℓ) : Set (a ⊔ b ⊔ ℓ) where constructor nonEmpty field {x} : A {y} : B proof : T x y ------------------------------------------------------------------------ -- Propositional equality -- This dummy module is used to avoid shadowing of the field named -- refl defined in IsEquivalence below. The module is opened publicly -- at the end of this file. private module Dummy where infix 4 _≡_ _≢_ data _≡_ {a} {A : Set a} (x : A) : A → Set a where refl : x ≡ x {-# BUILTIN EQUALITY _≡_ #-} {-# BUILTIN REFL refl #-} -- Nonequality. _≢_ : ∀ {a} {A : Set a} → A → A → Set a x ≢ y = ¬ x ≡ y ------------------------------------------------------------------------ -- Equivalence relations -- The preorders of this library are defined in terms of an underlying -- equivalence relation, and hence equivalence relations are not -- defined in terms of preorders. record IsEquivalence {a ℓ} {A : Set a} (_≈_ : Rel A ℓ) : Set (a ⊔ ℓ) where field refl : Reflexive _≈_ sym : Symmetric _≈_ trans : Transitive _≈_ reflexive : Dummy._≡_ ⇒ _≈_ reflexive Dummy.refl = refl open Dummy public lib-0.7/src/Relation/Binary/EqReasoning.agda0000644000000000000000000000215312101774706017143 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Convenient syntax for equational reasoning ------------------------------------------------------------------------ -- Example use: -- n*0≡0 : ∀ n → n * 0 ≡ 0 -- n*0≡0 zero = refl -- n*0≡0 (suc n) = -- begin -- suc n * 0 -- ≈⟨ refl ⟩ -- n * 0 + 0 -- ≈⟨ ... ⟩ -- n * 0 -- ≈⟨ n*0≡0 n ⟩ -- 0 -- ∎ -- Note that some modules contain generalised versions of specific -- instantiations of this module. For instance, the module ≡-Reasoning -- in Relation.Binary.PropositionalEquality is recommended for -- equational reasoning when the underlying equality is -- Relation.Binary.PropositionalEquality._≡_. open import Relation.Binary module Relation.Binary.EqReasoning {s₁ s₂} (S : Setoid s₁ s₂) where open Setoid S import Relation.Binary.PreorderReasoning as PreR open PreR preorder public renaming ( _∼⟨_⟩_ to _≈⟨_⟩_ ; _≈⟨_⟩_ to _≡⟨_⟩_ ; _≈⟨⟩_ to _≡⟨⟩_ ) lib-0.7/src/Relation/Binary/HeterogeneousEquality.agda0000644000000000000000000001557512101774706021276 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Heterogeneous equality ------------------------------------------------------------------------ module Relation.Binary.HeterogeneousEquality where open import Data.Product open import Function open import Function.Inverse using (Inverse) open import Level open import Relation.Nullary open import Relation.Binary open import Relation.Binary.Consequences open import Relation.Binary.Indexed as I using (_at_) open import Relation.Binary.PropositionalEquality as P using (_≡_; refl) import Relation.Binary.HeterogeneousEquality.Core as Core ------------------------------------------------------------------------ -- Heterogeneous equality infix 4 _≇_ open Core public using (_≅_; refl) -- Nonequality. _≇_ : ∀ {a} {A : Set a} → A → ∀ {b} {B : Set b} → B → Set x ≇ y = ¬ x ≅ y ------------------------------------------------------------------------ -- Conversion ≡-to-≅ : ∀ {a} {A : Set a} {x y : A} → x ≡ y → x ≅ y ≡-to-≅ refl = refl open Core public using (≅-to-≡) ------------------------------------------------------------------------ -- Some properties 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 refl = refl trans : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} {x : A} {y : B} {z : C} → x ≅ y → y ≅ z → x ≅ z trans refl eq = eq subst : ∀ {a} {A : Set a} {p} → Substitutive {A = A} (λ x y → x ≅ y) p subst P refl p = p subst₂ : ∀ {a b p} {A : Set a} {B : Set b} (P : A → B → Set p) → ∀ {x₁ x₂ y₁ y₂} → x₁ ≅ x₂ → y₁ ≅ y₂ → P x₁ y₁ → P x₂ y₂ subst₂ P refl refl p = p subst-removable : ∀ {a p} {A : Set a} (P : A → Set p) {x y} (eq : x ≅ y) z → subst P eq z ≅ z subst-removable P refl z = refl ≡-subst-removable : ∀ {a p} {A : Set a} (P : A → Set p) {x y} (eq : x ≡ y) z → P.subst P eq z ≅ z ≡-subst-removable P refl z = refl cong : ∀ {a b} {A : Set a} {B : A → Set b} {x y} (f : (x : A) → B x) → x ≅ y → f x ≅ f y cong f refl = 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 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} (p q : x ≅ y) → p ≡ q proof-irrelevance refl refl = refl isEquivalence : ∀ {a} {A : Set a} → IsEquivalence {A = A} (λ x y → x ≅ y) isEquivalence = record { refl = refl ; sym = sym ; trans = trans } setoid : ∀ {a} → Set a → Setoid _ _ setoid A = record { Carrier = A ; _≈_ = λ x y → x ≅ y ; isEquivalence = isEquivalence } indexedSetoid : ∀ {a b} {A : Set a} → (A → Set b) → I.Setoid A _ _ indexedSetoid B = record { Carrier = B ; _≈_ = λ x y → x ≅ y ; isEquivalence = record { refl = refl ; sym = sym ; trans = trans } } ≡↔≅ : ∀ {a b} {A : Set a} (B : A → Set b) {x : A} → Inverse (P.setoid (B x)) (indexedSetoid B at x) ≡↔≅ B = record { to = record { _⟨$⟩_ = id; cong = ≡-to-≅ } ; from = record { _⟨$⟩_ = id; cong = ≅-to-≡ } ; inverse-of = record { left-inverse-of = λ _ → refl ; right-inverse-of = λ _ → refl } } decSetoid : ∀ {a} {A : Set a} → Decidable {A = A} {B = A} (λ x y → x ≅ y) → DecSetoid _ _ decSetoid dec = record { _≈_ = λ x y → x ≅ y ; isDecEquivalence = record { isEquivalence = isEquivalence ; _≟_ = dec } } isPreorder : ∀ {a} {A : Set a} → IsPreorder {A = A} (λ x y → x ≅ y) (λ x y → x ≅ y) isPreorder = record { isEquivalence = isEquivalence ; reflexive = id ; trans = trans } isPreorder-≡ : ∀ {a} {A : Set a} → IsPreorder {A = A} _≡_ (λ x y → x ≅ y) isPreorder-≡ = record { isEquivalence = P.isEquivalence ; reflexive = reflexive ; trans = trans } preorder : ∀ {a} → Set a → Preorder _ _ _ preorder A = record { Carrier = A ; _≈_ = _≡_ ; _∼_ = λ x y → x ≅ y ; isPreorder = isPreorder-≡ } ------------------------------------------------------------------------ -- Convenient syntax for equational reasoning module ≅-Reasoning where -- The code in Relation.Binary.EqReasoning cannot handle -- heterogeneous equalities, hence the code duplication here. infix 4 _IsRelatedTo_ infix 2 _∎ infixr 2 _≅⟨_⟩_ _≡⟨_⟩_ _≡⟨⟩_ infix 1 begin_ data _IsRelatedTo_ {a} {A : Set a} (x : A) {b} {B : Set b} (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} → 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} → 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} → 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} → x IsRelatedTo y → x IsRelatedTo y _ ≡⟨⟩ x≅y = x≅y _∎ : ∀ {a} {A : Set a} (x : A) → x IsRelatedTo x _∎ _ = relTo refl ------------------------------------------------------------------------ -- Functional extensionality -- A form of functional extensionality for _≅_. Extensionality : (a b : Level) → Set _ Extensionality a b = {A : Set a} {B₁ B₂ : A → Set b} {f₁ : (x : A) → B₁ x} {f₂ : (x : A) → B₂ x} → (∀ x → B₁ x ≡ B₂ x) → (∀ x → f₁ x ≅ f₂ x) → f₁ ≅ f₂ -- This form of extensionality follows from extensionality for _≡_. ≡-ext-to-≅-ext : ∀ {ℓ₁ ℓ₂} → P.Extensionality ℓ₁ (suc ℓ₂) → Extensionality ℓ₁ ℓ₂ ≡-ext-to-≅-ext ext B₁≡B₂ f₁≅f₂ with ext B₁≡B₂ ≡-ext-to-≅-ext {ℓ₁} {ℓ₂} ext B₁≡B₂ f₁≅f₂ | P.refl = ≡-to-≅ $ ext′ (≅-to-≡ ∘ f₁≅f₂) where ext′ : P.Extensionality ℓ₁ ℓ₂ ext′ = P.extensionality-for-lower-levels ℓ₁ (suc ℓ₂) ext lib-0.7/src/Relation/Binary/NonStrictToStrict.agda0000644000000000000000000000772112101774706020355 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Conversion of ≤ to <, along with a number of properties ------------------------------------------------------------------------ -- Possible TODO: Prove that a conversion ≤ → < → ≤ returns a -- relation equivalent to the original one (and similarly for -- < → ≤ → <). open import Relation.Binary module Relation.Binary.NonStrictToStrict {a ℓ₁ ℓ₂} {A : Set a} (_≈_ : Rel A ℓ₁) (_≤_ : Rel A ℓ₂) where open import Relation.Nullary open import Relation.Binary.Consequences open import Function open import Data.Product open import Data.Sum ------------------------------------------------------------------------ -- Conversion -- _≤_ can be turned into _<_ as follows: _<_ : Rel A _ x < y = (x ≤ y) × ¬ (x ≈ y) ------------------------------------------------------------------------ -- The converted relations have certain properties -- (if the original relations have certain other properties) irrefl : Irreflexive _≈_ _<_ irrefl x≈y x (x≉y ∘ flip antisym x≥y ∘ proj₁) x≉y (x≥y , x≉y ∘ ≈-sym) decidable : Decidable _≈_ → Decidable _≤_ → Decidable _<_ decidable ≈-dec ≤-dec x y with ≈-dec x y | ≤-dec x y ... | yes x≈y | _ = no (flip proj₂ x≈y) ... | no x≉y | yes x≤y = yes (x≤y , x≉y) ... | no x≉y | no x≰y = no (x≰y ∘ proj₁) isPartialOrder⟶isStrictPartialOrder : IsPartialOrder _≈_ _≤_ → IsStrictPartialOrder _≈_ _<_ isPartialOrder⟶isStrictPartialOrder po = record { isEquivalence = PO.isEquivalence ; irrefl = irrefl ; trans = trans po ; <-resp-≈ = <-resp-≈ PO.isEquivalence PO.≤-resp-≈ } where module PO = IsPartialOrder po isTotalOrder⟶isStrictTotalOrder : Decidable _≈_ → IsTotalOrder _≈_ _≤_ → IsStrictTotalOrder _≈_ _<_ isTotalOrder⟶isStrictTotalOrder dec-≈ tot = record { isEquivalence = TO.isEquivalence ; trans = trans TO.isPartialOrder ; compare = trichotomous TO.Eq.sym dec-≈ TO.antisym TO.total ; <-resp-≈ = <-resp-≈ TO.isEquivalence TO.≤-resp-≈ } where module TO = IsTotalOrder tot isDecTotalOrder⟶isStrictTotalOrder : IsDecTotalOrder _≈_ _≤_ → IsStrictTotalOrder _≈_ _<_ isDecTotalOrder⟶isStrictTotalOrder dtot = isTotalOrder⟶isStrictTotalOrder DTO._≟_ DTO.isTotalOrder where module DTO = IsDecTotalOrder dtot lib-0.7/src/Relation/Binary/Indexed.agda0000644000000000000000000000203312101774706016305 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Indexed binary relations ------------------------------------------------------------------------ module Relation.Binary.Indexed where import Relation.Binary as B open import Relation.Binary.Indexed.Core public -- By "instantiating" an indexed setoid one gets an ordinary setoid. _at_ : ∀ {i s₁ s₂} {I : Set i} → Setoid I s₁ s₂ → I → B.Setoid _ _ S at i = record { Carrier = S.Carrier i ; _≈_ = S._≈_ ; isEquivalence = record { refl = S.refl ; sym = S.sym ; trans = S.trans } } where module S = Setoid S ------------------------------------------------------------------------ -- Simple properties of indexed binary relations -- Generalised implication. infixr 4 _=[_]⇒_ _=[_]⇒_ : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : A → Set b} → B.Rel A ℓ₁ → ((x : A) → B x) → Rel B ℓ₂ → Set _ P =[ f ]⇒ Q = ∀ {i j} → P i j → Q (f i) (f j) lib-0.7/src/Relation/Binary/Reflection.agda0000644000000000000000000000753412101774706017032 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Helpers intended to ease the development of "tactics" which use -- proof by reflection ------------------------------------------------------------------------ open import Data.Fin open import Data.Nat open import Data.Vec as Vec open import Function open import Function.Equality using (_⟨$⟩_) open import Function.Equivalence using (module Equivalence) open import Level open import Relation.Binary import Relation.Binary.PropositionalEquality as P -- Think of the parameters as follows: -- -- * Expr: A representation of code. -- * var: The Expr type should support a notion of variables. -- * ⟦_⟧: Computes the semantics of an expression. Takes an -- environment mapping variables to something. -- * ⟦_⇓⟧: Computes the semantics of the normal form of the -- expression. -- * correct: Normalisation preserves the semantics. -- -- Given these parameters two "tactics" are returned, prove and solve. -- -- For an example of the use of this module, see Algebra.RingSolver. module Relation.Binary.Reflection {e a s} {Expr : ℕ → Set e} {A : Set a} (Sem : Setoid a s) (var : ∀ {n} → Fin n → Expr n) (⟦_⟧ ⟦_⇓⟧ : ∀ {n} → Expr n → Vec A n → Setoid.Carrier Sem) (correct : ∀ {n} (e : Expr n) ρ → ⟦ e ⇓⟧ ρ ⟨ Setoid._≈_ Sem ⟩ ⟦ e ⟧ ρ) where open import Data.Vec.N-ary open import Data.Product import Relation.Binary.EqReasoning as Eq open Setoid Sem open Eq Sem -- If two normalised expressions are semantically equal, then their -- non-normalised forms are also equal. prove : ∀ {n} (ρ : Vec A n) e₁ e₂ → ⟦ e₁ ⇓⟧ ρ ≈ ⟦ e₂ ⇓⟧ ρ → ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ prove ρ e₁ e₂ hyp = begin ⟦ e₁ ⟧ ρ ≈⟨ sym (correct e₁ ρ) ⟩ ⟦ e₁ ⇓⟧ ρ ≈⟨ hyp ⟩ ⟦ e₂ ⇓⟧ ρ ≈⟨ correct e₂ ρ ⟩ ⟦ e₂ ⟧ ρ ∎ -- Applies the function to all possible "variables". close : ∀ {A : Set e} n → N-ary n (Expr n) A → A close n f = f $ⁿ Vec.map var (allFin n) -- A variant of prove which should in many cases be easier to use, -- because variables and environments are handled in a less explicit -- way. -- -- If the type signature of solve is a bit daunting, then it may be -- helpful to instantiate n with a small natural number and normalise -- the remainder of the type. solve : ∀ n (f : N-ary n (Expr n) (Expr n × Expr n)) → Eqʰ n _≈_ (curryⁿ ⟦ proj₁ (close n f) ⇓⟧) (curryⁿ ⟦ proj₂ (close n f) ⇓⟧) → Eq n _≈_ (curryⁿ ⟦ proj₁ (close n f) ⟧) (curryⁿ ⟦ proj₂ (close n f) ⟧) solve n f hyp = curryⁿ-cong _≈_ ⟦ proj₁ (close n f) ⟧ ⟦ proj₂ (close n f) ⟧ (λ ρ → prove ρ (proj₁ (close n f)) (proj₂ (close n f)) (curryⁿ-cong⁻¹ _≈_ ⟦ proj₁ (close n f) ⇓⟧ ⟦ proj₂ (close n f) ⇓⟧ (Eqʰ-to-Eq n _≈_ hyp) ρ)) -- A variant of solve which does not require that the normal form -- equality is proved for an arbitrary environment. solve₁ : ∀ n (f : N-ary n (Expr n) (Expr n × Expr n)) → ∀ⁿ n (curryⁿ λ ρ → ⟦ proj₁ (close n f) ⇓⟧ ρ ≈ ⟦ proj₂ (close n f) ⇓⟧ ρ → ⟦ proj₁ (close n f) ⟧ ρ ≈ ⟦ proj₂ (close n f) ⟧ ρ) solve₁ n f = Equivalence.from (uncurry-∀ⁿ n) ⟨$⟩ λ ρ → P.subst id (P.sym (left-inverse (λ _ → _ ≈ _ → _ ≈ _) ρ)) (prove ρ (proj₁ (close n f)) (proj₂ (close n f))) -- A variant of _,_ which is intended to make uses of solve and solve₁ -- look a bit nicer. infix 4 _⊜_ _⊜_ : ∀ {n} → Expr n → Expr n → Expr n × Expr n _⊜_ = _,_ lib-0.7/src/Relation/Binary/OrderMorphism.agda0000644000000000000000000000313312101774706017521 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Order morphisms ------------------------------------------------------------------------ module Relation.Binary.OrderMorphism where open import Relation.Binary open Poset import Function as F open import Level record _⇒-Poset_ {p₁ p₂ p₃ p₄ p₅ p₆} (P₁ : Poset p₁ p₂ p₃) (P₂ : Poset p₄ p₅ p₆) : Set (p₁ ⊔ p₃ ⊔ p₄ ⊔ p₆) where field fun : Carrier P₁ → Carrier P₂ monotone : _≤_ P₁ =[ fun ]⇒ _≤_ P₂ _⇒-DTO_ : ∀ {p₁ p₂ p₃ p₄ p₅ p₆} → DecTotalOrder p₁ p₂ p₃ → DecTotalOrder p₄ p₅ p₆ → Set _ DTO₁ ⇒-DTO DTO₂ = poset DTO₁ ⇒-Poset poset DTO₂ where open DecTotalOrder open _⇒-Poset_ id : ∀ {p₁ p₂ p₃} {P : Poset p₁ p₂ p₃} → P ⇒-Poset P id = record { fun = F.id ; monotone = F.id } _∘_ : ∀ {p₁ p₂ p₃ p₄ p₅ p₆ p₇ p₈ p₉} {P₁ : Poset p₁ p₂ p₃} {P₂ : Poset p₄ p₅ p₆} {P₃ : Poset p₇ p₈ p₉} → P₂ ⇒-Poset P₃ → P₁ ⇒-Poset P₂ → P₁ ⇒-Poset P₃ f ∘ g = record { fun = F._∘_ (fun f) (fun g) ; monotone = F._∘_ (monotone f) (monotone g) } const : ∀ {p₁ p₂ p₃ p₄ p₅ p₆} {P₁ : Poset p₁ p₂ p₃} {P₂ : Poset p₄ p₅ p₆} → Carrier P₂ → P₁ ⇒-Poset P₂ const {P₂ = P₂} x = record { fun = F.const x ; monotone = F.const (refl P₂) } lib-0.7/src/Relation/Binary/StrictToNonStrict.agda0000644000000000000000000000704612101774706020355 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Conversion of < to ≤, along with a number of properties ------------------------------------------------------------------------ -- Possible TODO: Prove that a conversion ≤ → < → ≤ returns a -- relation equivalent to the original one (and similarly for -- < → ≤ → <). open import Relation.Binary module Relation.Binary.StrictToNonStrict {a ℓ₁ ℓ₂} {A : Set a} (_≈_ : Rel A ℓ₁) (_<_ : Rel A ℓ₂) where open import Relation.Nullary open import Relation.Binary.Consequences open import Function open import Data.Product open import Data.Sum open import Data.Empty ------------------------------------------------------------------------ -- Conversion -- _<_ can be turned into _≤_ as follows: _≤_ : Rel A _ x ≤ y = (x < y) ⊎ (x ≈ y) ------------------------------------------------------------------------ -- The converted relations have certain properties -- (if the original relations have certain other properties) reflexive : _≈_ ⇒ _≤_ reflexive = inj₂ antisym : IsEquivalence _≈_ → Transitive _<_ → Irreflexive _≈_ _<_ → Antisymmetric _≈_ _≤_ antisym eq trans irrefl = as where module Eq = IsEquivalence eq as : Antisymmetric _≈_ _≤_ as (inj₂ x≈y) _ = x≈y as (inj₁ _) (inj₂ y≈x) = Eq.sym y≈x as (inj₁ x x≮y x≉y x>y = inj₂ (inj₁ x>y) decidable : Decidable _≈_ → Decidable _<_ → Decidable _≤_ decidable ≈-dec <-dec x y with ≈-dec x y | <-dec x y ... | yes x≈y | _ = yes (inj₂ x≈y) ... | no x≉y | yes x x≮y x≉y _ = no helper where helper : x ≤ y → ⊥ helper (inj₁ xy with tri x y ... | tri< _ _ x≯y = x≯y x>y ... | tri≈ _ _ x≯y = x≯y x>y ... | tri> x≮y _ _ = x≮y x _ x≉y _ = no x≉y tri⟶dec< : ∀ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} → Trichotomous _≈_ _<_ → Decidable _<_ tri⟶dec< compare x y with compare x y ... | tri< x x≮y _ _ = no x≮y map-NonEmpty : ∀ {a b p q} {A : Set a} {B : Set b} {P : REL A B p} {Q : REL A B q} → P ⇒ Q → NonEmpty P → NonEmpty Q map-NonEmpty f x = nonEmpty (f (NonEmpty.proof x)) lib-0.7/src/Relation/Binary/On.agda0000644000000000000000000002000012101774706015273 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Many properties which hold for _∼_ also hold for _∼_ on f ------------------------------------------------------------------------ open import Relation.Binary module Relation.Binary.On where open import Function open import Data.Product module _ {a b} {A : Set a} {B : Set b} (f : B → A) where implies : ∀ {ℓ₁ ℓ₂} (≈ : Rel A ℓ₁) (∼ : Rel A ℓ₂) → ≈ ⇒ ∼ → (≈ on f) ⇒ (∼ on f) implies _ _ impl = impl reflexive : ∀ {ℓ} (∼ : Rel A ℓ) → Reflexive ∼ → Reflexive (∼ on f) reflexive _ refl = refl irreflexive : ∀ {ℓ₁ ℓ₂} (≈ : Rel A ℓ₁) (∼ : Rel A ℓ₂) → Irreflexive ≈ ∼ → Irreflexive (≈ on f) (∼ on f) irreflexive _ _ irrefl = irrefl symmetric : ∀ {ℓ} (∼ : Rel A ℓ) → Symmetric ∼ → Symmetric (∼ on f) symmetric _ sym = sym transitive : ∀ {ℓ} (∼ : Rel A ℓ) → Transitive ∼ → Transitive (∼ on f) transitive _ trans = trans antisymmetric : ∀ {ℓ₁ ℓ₂} (≈ : Rel A ℓ₁) (≤ : Rel A ℓ₂) → Antisymmetric ≈ ≤ → Antisymmetric (≈ on f) (≤ on f) antisymmetric _ _ antisym = antisym asymmetric : ∀ {ℓ} (< : Rel A ℓ) → Asymmetric < → Asymmetric (< on f) asymmetric _ asym = asym respects : ∀ {ℓ p} (∼ : Rel A ℓ) (P : A → Set p) → P Respects ∼ → (P ∘ f) Respects (∼ on f) respects _ _ resp = resp respects₂ : ∀ {ℓ₁ ℓ₂} (∼₁ : Rel A ℓ₁) (∼₂ : Rel A ℓ₂) → ∼₁ Respects₂ ∼₂ → (∼₁ on f) Respects₂ (∼₂ on f) respects₂ _ _ (resp₁ , resp₂) = ((λ {_} {_} {_} → resp₁) , λ {_} {_} {_} → resp₂) decidable : ∀ {ℓ} (∼ : Rel A ℓ) → Decidable ∼ → Decidable (∼ on f) decidable _ dec = λ x y → dec (f x) (f y) total : ∀ {ℓ} (∼ : Rel A ℓ) → Total ∼ → Total (∼ on f) total _ tot = λ x y → tot (f x) (f y) trichotomous : ∀ {ℓ₁ ℓ₂} (≈ : Rel A ℓ₁) (< : Rel A ℓ₂) → Trichotomous ≈ < → Trichotomous (≈ on f) (< on f) trichotomous _ _ compare = λ x y → compare (f x) (f y) isEquivalence : ∀ {ℓ} {≈ : Rel A ℓ} → IsEquivalence ≈ → IsEquivalence (≈ on f) isEquivalence {≈ = ≈} eq = record { refl = reflexive ≈ Eq.refl ; sym = symmetric ≈ Eq.sym ; trans = transitive ≈ Eq.trans } where module Eq = IsEquivalence eq isPreorder : ∀ {ℓ₁ ℓ₂} {≈ : Rel A ℓ₁} {∼ : Rel A ℓ₂} → IsPreorder ≈ ∼ → IsPreorder (≈ on f) (∼ on f) isPreorder {≈ = ≈} {∼} pre = record { isEquivalence = isEquivalence Pre.isEquivalence ; reflexive = implies ≈ ∼ Pre.reflexive ; trans = transitive ∼ Pre.trans } where module Pre = IsPreorder pre isDecEquivalence : ∀ {ℓ} {≈ : Rel A ℓ} → IsDecEquivalence ≈ → IsDecEquivalence (≈ on f) isDecEquivalence {≈ = ≈} dec = record { isEquivalence = isEquivalence Dec.isEquivalence ; _≟_ = decidable ≈ Dec._≟_ } where module Dec = IsDecEquivalence dec isPartialOrder : ∀ {ℓ₁ ℓ₂} {≈ : Rel A ℓ₁} {≤ : Rel A ℓ₂} → IsPartialOrder ≈ ≤ → IsPartialOrder (≈ on f) (≤ on f) isPartialOrder {≈ = ≈} {≤} po = record { isPreorder = isPreorder Po.isPreorder ; antisym = antisymmetric ≈ ≤ Po.antisym } where module Po = IsPartialOrder po isDecPartialOrder : ∀ {ℓ₁ ℓ₂} {≈ : Rel A ℓ₁} {≤ : Rel A ℓ₂} → IsDecPartialOrder ≈ ≤ → IsDecPartialOrder (≈ on f) (≤ on f) isDecPartialOrder dpo = record { isPartialOrder = isPartialOrder DPO.isPartialOrder ; _≟_ = decidable _ DPO._≟_ ; _≤?_ = decidable _ DPO._≤?_ } where module DPO = IsDecPartialOrder dpo isStrictPartialOrder : ∀ {ℓ₁ ℓ₂} {≈ : Rel A ℓ₁} {< : Rel A ℓ₂} → IsStrictPartialOrder ≈ < → IsStrictPartialOrder (≈ on f) (< on f) isStrictPartialOrder {≈ = ≈} {<} spo = record { isEquivalence = isEquivalence Spo.isEquivalence ; irrefl = irreflexive ≈ < Spo.irrefl ; trans = transitive < Spo.trans ; <-resp-≈ = respects₂ < ≈ Spo.<-resp-≈ } where module Spo = IsStrictPartialOrder spo isTotalOrder : ∀ {ℓ₁ ℓ₂} {≈ : Rel A ℓ₁} {≤ : Rel A ℓ₂} → IsTotalOrder ≈ ≤ → IsTotalOrder (≈ on f) (≤ on f) isTotalOrder {≈ = ≈} {≤} to = record { isPartialOrder = isPartialOrder To.isPartialOrder ; total = total ≤ To.total } where module To = IsTotalOrder to isDecTotalOrder : ∀ {ℓ₁ ℓ₂} {≈ : Rel A ℓ₁} {≤ : Rel A ℓ₂} → IsDecTotalOrder ≈ ≤ → IsDecTotalOrder (≈ on f) (≤ on f) isDecTotalOrder {≈ = ≈} {≤} dec = record { isTotalOrder = isTotalOrder Dec.isTotalOrder ; _≟_ = decidable ≈ Dec._≟_ ; _≤?_ = decidable ≤ Dec._≤?_ } where module Dec = IsDecTotalOrder dec isStrictTotalOrder : ∀ {ℓ₁ ℓ₂} {≈ : Rel A ℓ₁} {< : Rel A ℓ₂} → IsStrictTotalOrder ≈ < → IsStrictTotalOrder (≈ on f) (< on f) isStrictTotalOrder {≈ = ≈} {<} sto = record { isEquivalence = isEquivalence Sto.isEquivalence ; trans = transitive < Sto.trans ; compare = trichotomous ≈ < Sto.compare ; <-resp-≈ = respects₂ < ≈ Sto.<-resp-≈ } where module Sto = IsStrictTotalOrder sto preorder : ∀ {p₁ p₂ p₃ b} {B : Set b} (P : Preorder p₁ p₂ p₃) → (B → Preorder.Carrier P) → Preorder _ _ _ preorder P f = record { isPreorder = isPreorder f (Preorder.isPreorder P) } setoid : ∀ {s₁ s₂ b} {B : Set b} (S : Setoid s₁ s₂) → (B → Setoid.Carrier S) → Setoid _ _ setoid S f = record { isEquivalence = isEquivalence f (Setoid.isEquivalence S) } decSetoid : ∀ {d₁ d₂ b} {B : Set b} (D : DecSetoid d₁ d₂) → (B → DecSetoid.Carrier D) → DecSetoid _ _ decSetoid D f = record { isDecEquivalence = isDecEquivalence f (DecSetoid.isDecEquivalence D) } poset : ∀ {p₁ p₂ p₃ b} {B : Set b} (P : Poset p₁ p₂ p₃) → (B → Poset.Carrier P) → Poset _ _ _ poset P f = record { isPartialOrder = isPartialOrder f (Poset.isPartialOrder P) } decPoset : ∀ {d₁ d₂ d₃ b} {B : Set b} (D : DecPoset d₁ d₂ d₃) → (B → DecPoset.Carrier D) → DecPoset _ _ _ decPoset D f = record { isDecPartialOrder = isDecPartialOrder f (DecPoset.isDecPartialOrder D) } strictPartialOrder : ∀ {s₁ s₂ s₃ b} {B : Set b} (S : StrictPartialOrder s₁ s₂ s₃) → (B → StrictPartialOrder.Carrier S) → StrictPartialOrder _ _ _ strictPartialOrder S f = record { isStrictPartialOrder = isStrictPartialOrder f (StrictPartialOrder.isStrictPartialOrder S) } totalOrder : ∀ {t₁ t₂ t₃ b} {B : Set b} (T : TotalOrder t₁ t₂ t₃) → (B → TotalOrder.Carrier T) → TotalOrder _ _ _ totalOrder T f = record { isTotalOrder = isTotalOrder f (TotalOrder.isTotalOrder T) } decTotalOrder : ∀ {d₁ d₂ d₃ b} {B : Set b} (D : DecTotalOrder d₁ d₂ d₃) → (B → DecTotalOrder.Carrier D) → DecTotalOrder _ _ _ decTotalOrder D f = record { isDecTotalOrder = isDecTotalOrder f (DecTotalOrder.isDecTotalOrder D) } strictTotalOrder : ∀ {s₁ s₂ s₃ b} {B : Set b} (S : StrictTotalOrder s₁ s₂ s₃) → (B → StrictTotalOrder.Carrier S) → StrictTotalOrder _ _ _ strictTotalOrder S f = record { isStrictTotalOrder = isStrictTotalOrder f (StrictTotalOrder.isStrictTotalOrder S) } lib-0.7/src/Relation/Binary/InducedPreorders.agda0000644000000000000000000000276212101774706020177 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Induced preorders ------------------------------------------------------------------------ open import Relation.Binary module Relation.Binary.InducedPreorders {s₁ s₂} (S : Setoid s₁ s₂) -- The underlying equality. where open import Function open import Data.Product open Setoid S -- Every respectful unary relation induces a preorder. (No claim is -- made that this preorder is unique.) InducedPreorder₁ : ∀ {p} (P : Carrier → Set p) → P Respects _≈_ → Preorder _ _ _ InducedPreorder₁ P resp = record { _≈_ = _≈_ ; _∼_ = λ c₁ c₂ → P c₁ → P c₂ ; isPreorder = record { isEquivalence = isEquivalence ; reflexive = resp ; trans = flip _∘′_ } } -- Every respectful binary relation induces a preorder. (No claim is -- made that this preorder is unique.) InducedPreorder₂ : ∀ {a r} {A : Set a} → (_R_ : A → Carrier → Set r) → (∀ {x} → _R_ x Respects _≈_) → Preorder _ _ _ InducedPreorder₂ _R_ resp = record { _≈_ = _≈_ ; _∼_ = λ c₁ c₂ → ∀ {a} → a R c₁ → a R c₂ ; isPreorder = record { isEquivalence = isEquivalence ; reflexive = λ c₁≈c₂ → resp c₁≈c₂ ; trans = λ c₁∼c₂ c₂∼c₃ → c₂∼c₃ ∘ c₁∼c₂ } } lib-0.7/src/Relation/Binary/PropositionalEquality.agda0000644000000000000000000001454612101774706021321 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Propositional (intensional) equality ------------------------------------------------------------------------ module Relation.Binary.PropositionalEquality where open import Function open import Function.Equality using (Π; _⟶_; ≡-setoid) open import Data.Product open import Data.Unit.Core open import Level open import Relation.Binary import Relation.Binary.Indexed as I open import Relation.Binary.Consequences open import Relation.Binary.HeterogeneousEquality.Core as H using (_≅_) -- Some of the definitions can be found in the following modules: open import Relation.Binary.Core public using (_≡_; refl; _≢_) open import Relation.Binary.PropositionalEquality.Core public ------------------------------------------------------------------------ -- Some properties subst₂ : ∀ {a b p} {A : Set a} {B : Set b} (P : A → B → Set p) {x₁ x₂ y₁ y₂} → x₁ ≡ x₂ → y₁ ≡ y₂ → P x₁ y₁ → P x₂ y₂ subst₂ P refl refl p = p 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₂ : ∀ {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 proof-irrelevance : ∀ {a} {A : Set a} {x y : A} (p q : x ≡ y) → p ≡ q proof-irrelevance refl refl = refl setoid : ∀ {a} → Set a → Setoid _ _ setoid A = record { Carrier = A ; _≈_ = _≡_ ; isEquivalence = isEquivalence } decSetoid : ∀ {a} {A : Set a} → Decidable (_≡_ {A = A}) → DecSetoid _ _ decSetoid dec = record { _≈_ = _≡_ ; isDecEquivalence = record { isEquivalence = isEquivalence ; _≟_ = dec } } isPreorder : ∀ {a} {A : Set a} → IsPreorder {A = A} _≡_ _≡_ isPreorder = record { isEquivalence = isEquivalence ; reflexive = id ; trans = trans } preorder : ∀ {a} → Set a → Preorder _ _ _ preorder A = record { Carrier = A ; _≈_ = _≡_ ; _∼_ = _≡_ ; isPreorder = isPreorder } ------------------------------------------------------------------------ -- Pointwise equality infix 4 _≗_ _→-setoid_ : ∀ {a b} (A : Set a) (B : Set b) → Setoid _ _ A →-setoid B = ≡-setoid A (Setoid.indexedSetoid (setoid B)) _≗_ : ∀ {a b} {A : Set a} {B : Set b} (f g : A → B) → Set _ _≗_ {A = A} {B} = Setoid._≈_ (A →-setoid B) :→-to-Π : ∀ {a b₁ b₂} {A : Set a} {B : I.Setoid _ b₁ b₂} → ((x : A) → I.Setoid.Carrier B x) → Π (setoid A) B :→-to-Π {B = B} f = record { _⟨$⟩_ = f; cong = cong′ } where open I.Setoid B using (_≈_) cong′ : ∀ {x y} → x ≡ y → f x ≈ f y cong′ refl = I.Setoid.refl B →-to-⟶ : ∀ {a b₁ b₂} {A : Set a} {B : Setoid b₁ b₂} → (A → Setoid.Carrier B) → setoid A ⟶ B →-to-⟶ = :→-to-Π ------------------------------------------------------------------------ -- The old inspect idiom -- The old inspect idiom has been deprecated, and may be removed in -- the future. Use inspect on steroids instead. module Deprecated-inspect where -- The inspect idiom can be used when you want to pattern match on -- the result r of some expression e, and you also need to -- "remember" that r ≡ e. -- The inspect idiom has a problem: sometimes you can only pattern -- match on the p part of p with-≡ eq if you also pattern match on -- the eq part, and then you no longer have access to the equality. -- Inspect on steroids solves this problem. data Inspect {a} {A : Set a} (x : A) : Set a where _with-≡_ : (y : A) (eq : x ≡ y) → Inspect x inspect : ∀ {a} {A : Set a} (x : A) → Inspect x inspect x = x with-≡ refl -- Example usage: -- f x y with inspect (g x) -- f x y | c z with-≡ eq = ... ------------------------------------------------------------------------ -- Inspect on steroids -- 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 ] -- Example usage: -- f x y with g x | inspect g x -- f x y | c z | [ eq ] = ... ------------------------------------------------------------------------ -- Convenient syntax for equational reasoning import Relation.Binary.EqReasoning as EqR -- Relation.Binary.EqReasoning is more convenient to use with _≡_ if -- the combinators take the type argument (a) as a hidden argument, -- instead of being locked to a fixed type at module instantiation -- time. module ≡-Reasoning where module _ {a} {A : Set a} where open EqR (setoid A) public hiding (_≡⟨_⟩_) renaming (_≈⟨_⟩_ to _≡⟨_⟩_) infixr 2 _≅⟨_⟩_ _≅⟨_⟩_ : ∀ {a} {A : Set a} (x : A) {y z : A} → x ≅ y → y IsRelatedTo z → x IsRelatedTo z _ ≅⟨ x≅y ⟩ y≡z = _ ≡⟨ H.≅-to-≡ x≅y ⟩ y≡z ------------------------------------------------------------------------ -- Functional extensionality -- If _≡_ were extensional, then the following statement could be -- proved. Extensionality : (a b : Level) → Set _ Extensionality a b = {A : Set a} {B : A → Set b} {f g : (x : A) → B x} → (∀ x → f x ≡ g x) → f ≡ g -- If extensionality holds for a given universe level, then it also -- holds for lower ones. extensionality-for-lower-levels : ∀ {a₁ b₁} a₂ b₂ → Extensionality (a₁ ⊔ a₂) (b₁ ⊔ b₂) → Extensionality a₁ b₁ extensionality-for-lower-levels a₂ b₂ ext f≡g = cong (λ h → lower ∘ h ∘ lift) $ ext (cong (lift {ℓ = b₂}) ∘ f≡g ∘ lower {ℓ = a₂}) -- Functional extensionality implies a form of extensionality for -- Π-types. ∀-extensionality : ∀ {a b} → Extensionality a (suc b) → {A : Set a} (B₁ B₂ : A → Set b) → (∀ x → B₁ x ≡ B₂ x) → (∀ x → B₁ x) ≡ (∀ x → B₂ x) ∀-extensionality ext B₁ B₂ B₁≡B₂ with ext B₁≡B₂ ∀-extensionality ext B .B B₁≡B₂ | refl = refl lib-0.7/src/Relation/Binary/Vec/0000755000000000000000000000000012101774706014626 5ustar0000000000000000lib-0.7/src/Relation/Binary/Vec/Pointwise.agda0000644000000000000000000002056112101774706017431 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Pointwise lifting of relations to vectors ------------------------------------------------------------------------ module Relation.Binary.Vec.Pointwise where open import Category.Applicative.Indexed open import Data.Fin open import Data.Nat open import Data.Plus as Plus hiding (equivalent; map) open import Data.Vec as Vec hiding ([_]; head; tail; map) import Data.Vec.Properties as VecProp open import Function open import Function.Equality using (_⟨$⟩_) open import Function.Equivalence as Equiv using (_⇔_; ⇔-setoid; module Equivalence) import Level open import Relation.Binary open import Relation.Binary.PropositionalEquality as P using (_≡_) open import Relation.Nullary import Relation.Nullary.Decidable as Dec -- Functional definition. record Pointwise {ℓ} {A B : Set ℓ} (_∼_ : REL A B ℓ) {n} (xs : Vec A n) (ys : Vec B n) : Set ℓ where constructor ext field app : ∀ i → lookup i xs ∼ lookup i ys -- Inductive definition. infixr 5 _∷_ data Pointwise′ {ℓ} {A B : Set ℓ} (_∼_ : REL A B ℓ) : ∀ {n} (xs : Vec A n) (ys : Vec B n) → Set ℓ where [] : Pointwise′ _∼_ [] [] _∷_ : ∀ {n x y} {xs : Vec A n} {ys : Vec B n} (x∼y : x ∼ y) (xs∼ys : Pointwise′ _∼_ xs ys) → Pointwise′ _∼_ (x ∷ xs) (y ∷ ys) -- The two definitions are equivalent. equivalent : ∀ {ℓ} {A B : Set ℓ} {_∼_ : REL A B ℓ} {n} {xs : Vec A n} {ys : Vec B n} → Pointwise _∼_ xs ys ⇔ Pointwise′ _∼_ xs ys equivalent {A = A} {B} {_∼_} = Equiv.equivalence (to _ _) from where to : ∀ {n} (xs : Vec A n) (ys : Vec B n) → Pointwise _∼_ xs ys → Pointwise′ _∼_ xs ys to [] [] xs∼ys = [] to (x ∷ xs) (y ∷ ys) xs∼ys = Pointwise.app xs∼ys zero ∷ to xs ys (ext (Pointwise.app xs∼ys ∘ suc)) nil : Pointwise _∼_ [] [] nil = ext λ() cons : ∀ {n x y} {xs : Vec A n} {ys : Vec B n} → x ∼ y → Pointwise _∼_ xs ys → Pointwise _∼_ (x ∷ xs) (y ∷ ys) cons {x = x} {y} {xs} {ys} x∼y xs∼ys = ext helper where helper : ∀ i → lookup i (x ∷ xs) ∼ lookup i (y ∷ ys) helper zero = x∼y helper (suc i) = Pointwise.app xs∼ys i from : ∀ {n} {xs : Vec A n} {ys : Vec B n} → Pointwise′ _∼_ xs ys → Pointwise _∼_ xs ys from [] = nil from (x∼y ∷ xs∼ys) = cons x∼y (from xs∼ys) -- Some destructors. head : ∀ {ℓ} {A B : Set ℓ} {_∼_ : REL A B ℓ} {n x y xs} {ys : Vec B n} → Pointwise′ _∼_ (x ∷ xs) (y ∷ ys) → x ∼ y head (x∼y ∷ xs∼ys) = x∼y tail : ∀ {ℓ} {A B : Set ℓ} {_∼_ : REL A B ℓ} {n x y xs} {ys : Vec B n} → Pointwise′ _∼_ (x ∷ xs) (y ∷ ys) → Pointwise′ _∼_ xs ys tail (x∼y ∷ xs∼ys) = xs∼ys -- Pointwise preserves reflexivity. refl : ∀ {ℓ} {A : Set ℓ} {_∼_ : Rel A ℓ} {n} → Reflexive _∼_ → Reflexive (Pointwise _∼_ {n = n}) refl rfl = ext (λ _ → rfl) -- Pointwise preserves symmetry. sym : ∀ {ℓ} {A B : Set ℓ} {P : REL A B ℓ} {Q : REL B A ℓ} {n} → Sym P Q → Sym (Pointwise P) (Pointwise Q {n = n}) sym sm xs∼ys = ext λ i → sm (Pointwise.app xs∼ys i) -- Pointwise preserves transitivity. trans : ∀ {ℓ} {A B C : Set ℓ} {P : REL A B ℓ} {Q : REL B C ℓ} {R : REL A C ℓ} {n} → Trans P Q R → Trans (Pointwise P) (Pointwise Q) (Pointwise R {n = n}) trans trns xs∼ys ys∼zs = ext λ i → trns (Pointwise.app xs∼ys i) (Pointwise.app ys∼zs i) -- Pointwise preserves equivalences. isEquivalence : ∀ {ℓ} {A : Set ℓ} {_∼_ : Rel A ℓ} {n} → IsEquivalence _∼_ → IsEquivalence (Pointwise _∼_ {n = n}) isEquivalence equiv = record { refl = refl (IsEquivalence.refl equiv) ; sym = sym (IsEquivalence.sym equiv) ; trans = trans (IsEquivalence.trans equiv) } -- Pointwise preserves decidability. decidable : ∀ {ℓ} {A B : Set ℓ} {_∼_ : REL A B ℓ} → Decidable _∼_ → ∀ {n} → Decidable (Pointwise _∼_ {n = n}) decidable {_∼_ = _∼_} dec xs ys = Dec.map (Setoid.sym (⇔-setoid _) equivalent) (decidable′ xs ys) where decidable′ : ∀ {n} → Decidable (Pointwise′ _∼_ {n = n}) decidable′ [] [] = yes [] decidable′ (x ∷ xs) (y ∷ ys) with dec x y ... | no ¬x∼y = no (¬x∼y ∘ head) ... | yes x∼y with decidable′ xs ys ... | no ¬xs∼ys = no (¬xs∼ys ∘ tail) ... | yes xs∼ys = yes (x∼y ∷ xs∼ys) -- Pointwise _≡_ is equivalent to _≡_. Pointwise-≡ : ∀ {ℓ} {A : Set ℓ} {n} {xs ys : Vec A n} → Pointwise _≡_ xs ys ⇔ xs ≡ ys Pointwise-≡ {ℓ} {A} = Equiv.equivalence (to ∘ _⟨$⟩_ {f₂ = ℓ} (Equivalence.to equivalent)) (λ xs≡ys → P.subst (Pointwise _≡_ _) xs≡ys (refl P.refl)) where to : ∀ {n} {xs ys : Vec A n} → Pointwise′ _≡_ xs ys → xs ≡ ys to [] = P.refl to (P.refl ∷ xs∼ys) = P.cong (_∷_ _) $ to xs∼ys -- Pointwise and Plus commute when the underlying relation is -- reflexive. ⁺∙⇒∙⁺ : ∀ {ℓ} {A : Set ℓ} {_∼_ : Rel A ℓ} {n} {xs ys : Vec A n} → Plus (Pointwise _∼_) xs ys → Pointwise (Plus _∼_) xs ys ⁺∙⇒∙⁺ [ ρ≈ρ′ ] = ext (λ x → [ Pointwise.app ρ≈ρ′ x ]) ⁺∙⇒∙⁺ (ρ ∼⁺⟨ ρ≈ρ′ ⟩ ρ′≈ρ″) = ext (λ x → _ ∼⁺⟨ Pointwise.app (⁺∙⇒∙⁺ ρ≈ρ′ ) x ⟩ Pointwise.app (⁺∙⇒∙⁺ ρ′≈ρ″) x) ∙⁺⇒⁺∙ : ∀ {ℓ} {A : Set ℓ} {_∼_ : Rel A ℓ} {n} {xs ys : Vec A n} → Reflexive _∼_ → Pointwise (Plus _∼_) xs ys → Plus (Pointwise _∼_) xs ys ∙⁺⇒⁺∙ {ℓ} {A} {_∼_} x∼x = Plus.map (_⟨$⟩_ {f₂ = ℓ} (Equivalence.from equivalent)) ∘ helper ∘ _⟨$⟩_ {f₂ = ℓ} (Equivalence.to equivalent) where helper : ∀ {n} {xs ys : Vec A n} → Pointwise′ (Plus _∼_) xs ys → Plus (Pointwise′ _∼_) xs ys helper [] = [ [] ] helper (_∷_ {x = x} {y = y} {xs = xs} {ys = ys} x∼y xs∼ys) = x ∷ xs ∼⁺⟨ Plus.map (λ x∼y → x∼y ∷ xs∼xs) x∼y ⟩ y ∷ xs ∼⁺⟨ Plus.map (λ xs∼ys → x∼x ∷ xs∼ys) (helper xs∼ys) ⟩∎ y ∷ ys ∎ where xs∼xs : Pointwise′ _∼_ xs xs xs∼xs = _⟨$⟩_ {f₂ = ℓ} (Equivalence.to equivalent) (refl x∼x) -- Note that ∙⁺⇒⁺∙ cannot be defined if the requirement of reflexivity -- is dropped. private module Counterexample where data D : Set where i j x y z : D data _R_ : Rel D Level.zero where iRj : i R j xRy : x R y yRz : y R z xR⁺z : x [ _R_ ]⁺ z xR⁺z = x ∼⁺⟨ [ xRy ] ⟩ y ∼⁺⟨ [ yRz ] ⟩∎ z ∎ ix : Vec D 2 ix = i ∷ x ∷ [] jz : Vec D 2 jz = j ∷ z ∷ [] ix∙⁺jz : Pointwise′ (Plus _R_) ix jz ix∙⁺jz = [ iRj ] ∷ xR⁺z ∷ [] ¬ix⁺∙jz : ¬ Plus′ (Pointwise′ _R_) ix jz ¬ix⁺∙jz [ iRj ∷ () ∷ [] ] ¬ix⁺∙jz ((iRj ∷ xRy ∷ []) ∷ [ () ∷ yRz ∷ [] ]) ¬ix⁺∙jz ((iRj ∷ xRy ∷ []) ∷ (() ∷ yRz ∷ []) ∷ _) counterexample : ¬ (∀ {n} {xs ys : Vec D n} → Pointwise (Plus _R_) xs ys → Plus (Pointwise _R_) xs ys) counterexample ∙⁺⇒⁺∙ = ¬ix⁺∙jz (Equivalence.to Plus.equivalent ⟨$⟩ Plus.map (_⟨$⟩_ (Equivalence.to equivalent)) (∙⁺⇒⁺∙ (Equivalence.from equivalent ⟨$⟩ ix∙⁺jz))) -- Map. map : ∀ {ℓ} {A : Set ℓ} {_R_ _R′_ : Rel A ℓ} {n} → _R_ ⇒ _R′_ → Pointwise _R_ ⇒ Pointwise _R′_ {n} map R⇒R′ xsRys = ext λ i → R⇒R′ (Pointwise.app xsRys i) -- A variant. gmap : ∀ {ℓ} {A A′ : Set ℓ} {_R_ : Rel A ℓ} {_R′_ : Rel A′ ℓ} {f : A → A′} {n} → _R_ =[ f ]⇒ _R′_ → Pointwise _R_ =[ Vec.map {n = n} f ]⇒ Pointwise _R′_ gmap {_R′_ = _R′_} {f} R⇒R′ {i = xs} {j = ys} xsRys = ext λ i → let module M = Morphism (VecProp.lookup-morphism i) in P.subst₂ _R′_ (P.sym $ M.op-<$> f xs) (P.sym $ M.op-<$> f ys) (R⇒R′ (Pointwise.app xsRys i)) lib-0.7/src/Relation/Binary/Indexed/0000755000000000000000000000000012101774706015471 5ustar0000000000000000lib-0.7/src/Relation/Binary/Indexed/Core.agda0000644000000000000000000000421712101774706017203 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Indexed binary relations ------------------------------------------------------------------------ -- This file contains some core definitions which are reexported by -- Relation.Binary.Indexed. module Relation.Binary.Indexed.Core where open import Function open import Level import Relation.Binary.Core as B import Relation.Binary.Core as P ------------------------------------------------------------------------ -- Indexed binary relations -- Heterogeneous. REL : ∀ {i₁ i₂ a₁ a₂} {I₁ : Set i₁} {I₂ : Set i₂} → (I₁ → Set a₁) → (I₂ → Set a₂) → (ℓ : Level) → Set _ REL A₁ A₂ ℓ = ∀ {i₁ i₂} → A₁ i₁ → A₂ i₂ → Set ℓ -- Homogeneous. Rel : ∀ {i a} {I : Set i} → (I → Set a) → (ℓ : Level) → Set _ Rel A ℓ = REL A A ℓ ------------------------------------------------------------------------ -- Simple properties of indexed binary relations -- Reflexivity. Reflexive : ∀ {i a ℓ} {I : Set i} (A : I → Set a) → Rel A ℓ → Set _ Reflexive _ _∼_ = ∀ {i} → B.Reflexive (_∼_ {i}) -- Symmetry. Symmetric : ∀ {i a ℓ} {I : Set i} (A : I → Set a) → Rel A ℓ → Set _ Symmetric _ _∼_ = ∀ {i j} → B.Sym (_∼_ {i} {j}) _∼_ -- Transitivity. Transitive : ∀ {i a ℓ} {I : Set i} (A : I → Set a) → Rel A ℓ → Set _ Transitive _ _∼_ = ∀ {i j k} → B.Trans _∼_ (_∼_ {j}) (_∼_ {i} {k}) ------------------------------------------------------------------------ -- Setoids record IsEquivalence {i a ℓ} {I : Set i} (A : I → Set a) (_≈_ : Rel A ℓ) : Set (i ⊔ a ⊔ ℓ) where field refl : Reflexive A _≈_ sym : Symmetric A _≈_ trans : Transitive A _≈_ reflexive : ∀ {i} → P._≡_ ⟨ B._⇒_ ⟩ _≈_ {i} reflexive P.refl = refl record Setoid {i} (I : Set i) c ℓ : Set (suc (i ⊔ c ⊔ ℓ)) where infix 4 _≈_ field Carrier : I → Set c _≈_ : Rel Carrier ℓ isEquivalence : IsEquivalence Carrier _≈_ open IsEquivalence isEquivalence public lib-0.7/src/Relation/Binary/Props/0000755000000000000000000000000012101774706015214 5ustar0000000000000000lib-0.7/src/Relation/Binary/Props/DecTotalOrder.agda0000644000000000000000000000154112101774706020526 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Properties satisfied by decidable total orders ------------------------------------------------------------------------ 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 lib-0.7/src/Relation/Binary/Props/Poset.agda0000644000000000000000000000161612101774706017130 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Properties satisfied by posets ------------------------------------------------------------------------ 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 lib-0.7/src/Relation/Binary/Props/StrictTotalOrder.agda0000644000000000000000000000204512101774706021303 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Properties satisfied by strict partial orders ------------------------------------------------------------------------ open import Relation.Binary 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 lib-0.7/src/Relation/Binary/Props/Preorder.agda0000644000000000000000000000144012101774706017613 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Properties satisfied by preorders ------------------------------------------------------------------------ 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) } } lib-0.7/src/Relation/Binary/Props/StrictPartialOrder.agda0000644000000000000000000000172012101774706021613 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Properties satisfied by strict partial orders ------------------------------------------------------------------------ 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 lib-0.7/src/Relation/Binary/Props/TotalOrder.agda0000644000000000000000000000125712101774706020116 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Properties satisfied by total orders ------------------------------------------------------------------------ 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 ≟ } } lib-0.7/src/Relation/Binary/Product/0000755000000000000000000000000012101774706015531 5ustar0000000000000000lib-0.7/src/Relation/Binary/Product/StrictLex.agda0000644000000000000000000003276512101774706020305 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Lexicographic products of binary relations ------------------------------------------------------------------------ -- The definition of lexicographic product used here is suitable if -- the left-hand relation is a strict partial order. module Relation.Binary.Product.StrictLex where open import Function open import Data.Product open import Data.Sum open import Data.Empty open import Level open import Relation.Nullary.Product open import Relation.Nullary.Sum open import Relation.Binary open import Relation.Binary.Consequences open import Relation.Binary.Product.Pointwise as Pointwise using (_×-Rel_) open import Relation.Nullary module _ {a₁ a₂ ℓ₁ ℓ₂} {A₁ : Set a₁} {A₂ : Set a₂} where ×-Lex : (_≈₁_ _<₁_ : Rel A₁ ℓ₁) → (_≤₂_ : Rel A₂ ℓ₂) → Rel (A₁ × A₂) _ ×-Lex _≈₁_ _<₁_ _≤₂_ = (_<₁_ on proj₁) -⊎- (_≈₁_ on proj₁) -×- (_≤₂_ on proj₂) -- Some properties which are preserved by ×-Lex (under certain -- assumptions). ×-reflexive : ∀ _≈₁_ _∼₁_ {_≈₂_ : Rel A₂ ℓ₂} _≤₂_ → _≈₂_ ⇒ _≤₂_ → (_≈₁_ ×-Rel _≈₂_) ⇒ (×-Lex _≈₁_ _∼₁_ _≤₂_) ×-reflexive _ _ _ refl₂ = λ x≈y → inj₂ (proj₁ x≈y , refl₂ (proj₂ x≈y)) _×-irreflexive_ : ∀ {_≈₁_ _<₁_} → Irreflexive _≈₁_ _<₁_ → ∀ {_≈₂_ _<₂_ : Rel A₂ ℓ₂} → Irreflexive _≈₂_ _<₂_ → Irreflexive (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_) (ir₁ ×-irreflexive ir₂) x≈y (inj₁ x₁y₁ = inj₂ (inj₁ x₁>y₁) ×-compare : {_≈₁_ _<₁_ : Rel A₁ ℓ₁} → Symmetric _≈₁_ → Trichotomous _≈₁_ _<₁_ → {_≈₂_ _<₂_ : Rel A₂ ℓ₂} → Trichotomous _≈₂_ _<₂_ → Trichotomous (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_) ×-compare {_≈₁_} {_<₁_} sym₁ compare₁ {_≈₂_} {_<₂_} compare₂ = cmp where cmp″ : ∀ {x₁ y₁ x₂ y₂} → ¬ (x₁ <₁ y₁) → x₁ ≈₁ y₁ → ¬ (y₁ <₁ x₁) → Tri (x₂ <₂ y₂) (x₂ ≈₂ y₂) (y₂ <₂ x₂) → Tri (×-Lex _≈₁_ _<₁_ _<₂_ (x₁ , x₂) (y₁ , y₂)) ((_≈₁_ ×-Rel _≈₂_) (x₁ , x₂) (y₁ , y₂)) (×-Lex _≈₁_ _<₁_ _<₂_ (y₁ , y₂) (x₁ , x₂)) cmp″ x₁≮y₁ x₁≈y₁ x₁≯y₁ (tri< x₂ x₂≮y₂ x₂≉y₂ x₂>y₂) = tri> [ x₁≮y₁ , x₂≮y₂ ∘ proj₂ ] (x₂≉y₂ ∘ proj₂) (inj₂ (sym₁ x₁≈y₁ , x₂>y₂)) cmp″ x₁≮y₁ x₁≈y₁ x₁≯y₁ (tri≈ x₂≮y₂ x₂≈y₂ x₂≯y₂) = tri≈ [ x₁≮y₁ , x₂≮y₂ ∘ proj₂ ] (x₁≈y₁ , x₂≈y₂) [ x₁≯y₁ , x₂≯y₂ ∘ proj₂ ] cmp′ : ∀ {x₁ y₁} → Tri (x₁ <₁ y₁) (x₁ ≈₁ y₁) (y₁ <₁ x₁) → ∀ x₂ y₂ → Tri (×-Lex _≈₁_ _<₁_ _<₂_ (x₁ , x₂) (y₁ , y₂)) ((_≈₁_ ×-Rel _≈₂_) (x₁ , x₂) (y₁ , y₂)) (×-Lex _≈₁_ _<₁_ _<₂_ (y₁ , y₂) (x₁ , x₂)) cmp′ (tri< x₁ x₁≮y₁ x₁≉y₁ x₁>y₁) x₂ y₂ = tri> [ x₁≮y₁ , x₁≉y₁ ∘ proj₁ ] (x₁≉y₁ ∘ proj₁) (inj₁ x₁>y₁) cmp′ (tri≈ x₁≮y₁ x₁≈y₁ x₁≯y₁) x₂ y₂ = cmp″ x₁≮y₁ x₁≈y₁ x₁≯y₁ (compare₂ x₂ y₂) cmp : Trichotomous (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_) cmp (x₁ , x₂) (y₁ , y₂) = cmp′ (compare₁ x₁ y₁) x₂ y₂ -- Some collections of properties which are preserved by ×-Lex. _×-isPreorder_ : ∀ {_≈₁_ _∼₁_} → IsPreorder _≈₁_ _∼₁_ → ∀ {_≈₂_ _∼₂_} → IsPreorder _≈₂_ _∼₂_ → IsPreorder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _∼₁_ _∼₂_) _×-isPreorder_ {_≈₁_ = _≈₁_} {_∼₁_ = _∼₁_} pre₁ {_∼₂_ = _∼₂_} pre₂ = record { isEquivalence = Pointwise._×-isEquivalence_ (isEquivalence pre₁) (isEquivalence pre₂) ; reflexive = λ {x y} → ×-reflexive _≈₁_ _∼₁_ _∼₂_ (reflexive pre₂) {x} {y} ; trans = λ {x y z} → ×-transitive (isEquivalence pre₁) (∼-resp-≈ pre₁) (trans pre₁) {_≤₂_ = _∼₂_} (trans pre₂) {x} {y} {z} } where open IsPreorder _×-isStrictPartialOrder_ : ∀ {_≈₁_ _<₁_} → IsStrictPartialOrder _≈₁_ _<₁_ → ∀ {_≈₂_ _<₂_} → IsStrictPartialOrder _≈₂_ _<₂_ → IsStrictPartialOrder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_) _×-isStrictPartialOrder_ {_<₁_ = _<₁_} spo₁ {_<₂_ = _<₂_} spo₂ = record { isEquivalence = Pointwise._×-isEquivalence_ (isEquivalence spo₁) (isEquivalence spo₂) ; irrefl = λ {x y} → _×-irreflexive_ {_<₁_ = _<₁_} (irrefl spo₁) {_<₂_ = _<₂_} (irrefl spo₂) {x} {y} ; trans = λ {x y z} → ×-transitive {_<₁_ = _<₁_} (isEquivalence spo₁) (<-resp-≈ spo₁) (trans spo₁) {_≤₂_ = _<₂_} (trans spo₂) {x} {y} {z} ; <-resp-≈ = ×-≈-respects₂ (isEquivalence spo₁) (<-resp-≈ spo₁) (<-resp-≈ spo₂) } where open IsStrictPartialOrder _×-isStrictTotalOrder_ : ∀ {_≈₁_ _<₁_} → IsStrictTotalOrder _≈₁_ _<₁_ → ∀ {_≈₂_ _<₂_} → IsStrictTotalOrder _≈₂_ _<₂_ → IsStrictTotalOrder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_) _×-isStrictTotalOrder_ {_<₁_ = _<₁_} spo₁ {_<₂_ = _<₂_} spo₂ = record { isEquivalence = Pointwise._×-isEquivalence_ (isEquivalence spo₁) (isEquivalence spo₂) ; trans = λ {x y z} → ×-transitive {_<₁_ = _<₁_} (isEquivalence spo₁) (<-resp-≈ spo₁) (trans spo₁) {_≤₂_ = _<₂_} (trans spo₂) {x} {y} {z} ; compare = ×-compare (Eq.sym spo₁) (compare spo₁) (compare spo₂) ; <-resp-≈ = ×-≈-respects₂ (isEquivalence spo₁) (<-resp-≈ spo₁) (<-resp-≈ spo₂) } where open IsStrictTotalOrder -- "Packages" (e.g. preorders) can also be combined. _×-preorder_ : ∀ {p₁ p₂ p₃ p₄} → Preorder p₁ p₂ _ → Preorder p₃ p₄ _ → Preorder _ _ _ p₁ ×-preorder p₂ = record { isPreorder = isPreorder p₁ ×-isPreorder isPreorder p₂ } where open Preorder _×-strictPartialOrder_ : ∀ {s₁ s₂ s₃ s₄} → StrictPartialOrder s₁ s₂ _ → StrictPartialOrder s₃ s₄ _ → StrictPartialOrder _ _ _ s₁ ×-strictPartialOrder s₂ = record { isStrictPartialOrder = isStrictPartialOrder s₁ ×-isStrictPartialOrder isStrictPartialOrder s₂ } where open StrictPartialOrder _×-strictTotalOrder_ : ∀ {s₁ s₂ s₃ s₄} → StrictTotalOrder s₁ s₂ _ → StrictTotalOrder s₃ s₄ _ → StrictTotalOrder _ _ _ s₁ ×-strictTotalOrder s₂ = record { isStrictTotalOrder = isStrictTotalOrder s₁ ×-isStrictTotalOrder isStrictTotalOrder s₂ } where open StrictTotalOrder lib-0.7/src/Relation/Binary/Product/NonStrictLex.agda0000644000000000000000000002026312101774706020746 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Lexicographic products of binary relations ------------------------------------------------------------------------ -- The definition of lexicographic product used here is suitable if -- the left-hand relation is a (non-strict) partial order. module Relation.Binary.Product.NonStrictLex where open import Data.Product open import Data.Sum open import Level open import Relation.Binary open import Relation.Binary.Consequences import Relation.Binary.NonStrictToStrict as Conv open import Relation.Binary.Product.Pointwise as Pointwise using (_×-Rel_) import Relation.Binary.Product.StrictLex as Strict module _ {a₁ a₂ ℓ₁ ℓ₂} {A₁ : Set a₁} {A₂ : Set a₂} where ×-Lex : (_≈₁_ _≤₁_ : Rel A₁ ℓ₁) → (_≤₂_ : Rel A₂ ℓ₂) → Rel (A₁ × A₂) _ ×-Lex _≈₁_ _≤₁_ _≤₂_ = Strict.×-Lex _≈₁_ (Conv._<_ _≈₁_ _≤₁_) _≤₂_ -- Some properties which are preserved by ×-Lex (under certain -- assumptions). ×-reflexive : ∀ _≈₁_ _≤₁_ {_≈₂_} _≤₂_ → _≈₂_ ⇒ _≤₂_ → (_≈₁_ ×-Rel _≈₂_) ⇒ (×-Lex _≈₁_ _≤₁_ _≤₂_) ×-reflexive _≈₁_ _≤₁_ _≤₂_ refl₂ {x} {y} = Strict.×-reflexive _≈₁_ (Conv._<_ _≈₁_ _≤₁_) _≤₂_ refl₂ {x} {y} ×-transitive : ∀ {_≈₁_ _≤₁_} → IsPartialOrder _≈₁_ _≤₁_ → ∀ {_≤₂_} → Transitive _≤₂_ → Transitive (×-Lex _≈₁_ _≤₁_ _≤₂_) ×-transitive {_≈₁_ = _≈₁_} {_≤₁_ = _≤₁_} po₁ {_≤₂_ = _≤₂_} trans₂ {x} {y} {z} = Strict.×-transitive {_<₁_ = Conv._<_ _≈₁_ _≤₁_} isEquivalence (Conv.<-resp-≈ _ _ isEquivalence ≤-resp-≈) (Conv.trans _ _ po₁) {_≤₂_ = _≤₂_} trans₂ {x} {y} {z} where open IsPartialOrder po₁ ×-antisymmetric : ∀ {_≈₁_ _≤₁_} → IsPartialOrder _≈₁_ _≤₁_ → ∀ {_≈₂_ _≤₂_} → Antisymmetric _≈₂_ _≤₂_ → Antisymmetric (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _≤₁_ _≤₂_) ×-antisymmetric {_≈₁_ = _≈₁_} {_≤₁_ = _≤₁_} po₁ {_≤₂_ = _≤₂_} antisym₂ {x} {y} = Strict.×-antisymmetric {_<₁_ = Conv._<_ _≈₁_ _≤₁_} ≈-sym₁ irrefl₁ asym₁ {_≤₂_ = _≤₂_} antisym₂ {x} {y} where open IsPartialOrder po₁ open Eq renaming (refl to ≈-refl₁; sym to ≈-sym₁) irrefl₁ : Irreflexive _≈₁_ (Conv._<_ _≈₁_ _≤₁_) irrefl₁ = Conv.irrefl _≈₁_ _≤₁_ asym₁ : Asymmetric (Conv._<_ _≈₁_ _≤₁_) asym₁ = trans∧irr⟶asym {_≈_ = _≈₁_} ≈-refl₁ (Conv.trans _ _ po₁) irrefl₁ ×-≈-respects₂ : ∀ {_≈₁_ _≤₁_} → IsEquivalence _≈₁_ → _≤₁_ Respects₂ _≈₁_ → ∀ {_≈₂_ _≤₂_ : Rel A₂ ℓ₂} → _≤₂_ Respects₂ _≈₂_ → (×-Lex _≈₁_ _≤₁_ _≤₂_) Respects₂ (_≈₁_ ×-Rel _≈₂_) ×-≈-respects₂ eq₁ resp₁ resp₂ = Strict.×-≈-respects₂ eq₁ (Conv.<-resp-≈ _ _ eq₁ resp₁) resp₂ ×-decidable : ∀ {_≈₁_ _≤₁_} → Decidable _≈₁_ → Decidable _≤₁_ → ∀ {_≤₂_} → Decidable _≤₂_ → Decidable (×-Lex _≈₁_ _≤₁_ _≤₂_) ×-decidable dec-≈₁ dec-≤₁ dec-≤₂ = Strict.×-decidable dec-≈₁ (Conv.decidable _ _ dec-≈₁ dec-≤₁) dec-≤₂ ×-total : ∀ {_≈₁_ _≤₁_} → Symmetric _≈₁_ → Decidable _≈₁_ → Antisymmetric _≈₁_ _≤₁_ → Total _≤₁_ → ∀ {_≤₂_} → Total _≤₂_ → Total (×-Lex _≈₁_ _≤₁_ _≤₂_) ×-total {_≈₁_ = _≈₁_} {_≤₁_ = _≤₁_} sym₁ dec₁ antisym₁ total₁ {_≤₂_ = _≤₂_} total₂ = total where tri₁ : Trichotomous _≈₁_ (Conv._<_ _≈₁_ _≤₁_) tri₁ = Conv.trichotomous _ _ sym₁ dec₁ antisym₁ total₁ total : Total (×-Lex _≈₁_ _≤₁_ _≤₂_) total x y with tri₁ (proj₁ x) (proj₁ y) ... | tri< x₁ x₁≮y₁ x₁≉y₁ x₁>y₁ = inj₂ (inj₁ x₁>y₁) ... | tri≈ x₁≮y₁ x₁≈y₁ x₁≯y₁ with total₂ (proj₂ x) (proj₂ y) ... | inj₁ x₂≤y₂ = inj₁ (inj₂ (x₁≈y₁ , x₂≤y₂)) ... | inj₂ x₂≥y₂ = inj₂ (inj₂ (sym₁ x₁≈y₁ , x₂≥y₂)) -- Some collections of properties which are preserved by ×-Lex -- (under certain assumptions). _×-isPartialOrder_ : ∀ {_≈₁_ _≤₁_} → IsPartialOrder _≈₁_ _≤₁_ → ∀ {_≈₂_ _≤₂_} → IsPartialOrder _≈₂_ _≤₂_ → IsPartialOrder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _≤₁_ _≤₂_) _×-isPartialOrder_ {_≈₁_ = _≈₁_} {_≤₁_ = _≤₁_} po₁ {_≤₂_ = _≤₂_} po₂ = record { isPreorder = record { isEquivalence = Pointwise._×-isEquivalence_ (isEquivalence po₁) (isEquivalence po₂) ; reflexive = λ {x y} → ×-reflexive _≈₁_ _≤₁_ _≤₂_ (reflexive po₂) {x} {y} ; trans = λ {x y z} → ×-transitive po₁ {_≤₂_ = _≤₂_} (trans po₂) {x} {y} {z} } ; antisym = λ {x y} → ×-antisymmetric {_≤₁_ = _≤₁_} po₁ {_≤₂_ = _≤₂_} (antisym po₂) {x} {y} } where open IsPartialOrder ×-isTotalOrder : ∀ {_≈₁_ _≤₁_} → Decidable _≈₁_ → IsTotalOrder _≈₁_ _≤₁_ → ∀ {_≈₂_ _≤₂_} → IsTotalOrder _≈₂_ _≤₂_ → IsTotalOrder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _≤₁_ _≤₂_) ×-isTotalOrder {_≤₁_ = _≤₁_} ≈₁-dec to₁ {_≤₂_ = _≤₂_} to₂ = record { isPartialOrder = isPartialOrder to₁ ×-isPartialOrder isPartialOrder to₂ ; total = ×-total {_≤₁_ = _≤₁_} (Eq.sym to₁) ≈₁-dec (antisym to₁) (total to₁) {_≤₂_ = _≤₂_} (total to₂) } where open IsTotalOrder _×-isDecTotalOrder_ : ∀ {_≈₁_ _≤₁_} → IsDecTotalOrder _≈₁_ _≤₁_ → ∀ {_≈₂_ _≤₂_} → IsDecTotalOrder _≈₂_ _≤₂_ → IsDecTotalOrder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _≤₁_ _≤₂_) _×-isDecTotalOrder_ {_≤₁_ = _≤₁_} to₁ {_≤₂_ = _≤₂_} to₂ = record { isTotalOrder = ×-isTotalOrder (_≟_ to₁) (isTotalOrder to₁) (isTotalOrder to₂) ; _≟_ = Pointwise._×-decidable_ (_≟_ to₁) (_≟_ to₂) ; _≤?_ = ×-decidable (_≟_ to₁) (_≤?_ to₁) (_≤?_ to₂) } where open IsDecTotalOrder -- "Packages" (e.g. posets) can also be combined. _×-poset_ : ∀ {p₁ p₂ p₃ p₄} → Poset p₁ p₂ _ → Poset p₃ p₄ _ → Poset _ _ _ p₁ ×-poset p₂ = record { isPartialOrder = isPartialOrder p₁ ×-isPartialOrder isPartialOrder p₂ } where open Poset _×-totalOrder_ : ∀ {d₁ d₂ t₃ t₄} → DecTotalOrder d₁ d₂ _ → TotalOrder t₃ t₄ _ → TotalOrder _ _ _ t₁ ×-totalOrder t₂ = record { isTotalOrder = ×-isTotalOrder T₁._≟_ T₁.isTotalOrder T₂.isTotalOrder } where module T₁ = DecTotalOrder t₁ module T₂ = TotalOrder t₂ _×-decTotalOrder_ : ∀ {d₁ d₂ d₃ d₄} → DecTotalOrder d₁ d₂ _ → DecTotalOrder d₃ d₄ _ → DecTotalOrder _ _ _ t₁ ×-decTotalOrder t₂ = record { isDecTotalOrder = isDecTotalOrder t₁ ×-isDecTotalOrder isDecTotalOrder t₂ } where open DecTotalOrder lib-0.7/src/Relation/Binary/Product/Pointwise.agda0000644000000000000000000004244712101774706020343 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Pointwise products of binary relations ------------------------------------------------------------------------ module Relation.Binary.Product.Pointwise where open import Data.Product as Prod open import Data.Sum open import Data.Unit using (⊤) open import Function open import Function.Equality as F using (_⟶_; _⟨$⟩_) open import Function.Equivalence as Eq using (Equivalence; _⇔_; module Equivalence) open import Function.Injection as Inj using (Injection; _↣_; module Injection) open import Function.Inverse as Inv using (Inverse; _↔_; module Inverse) open import Function.LeftInverse as LeftInv using (LeftInverse; _↞_; _LeftInverseOf_; module LeftInverse) open import Function.Related open import Function.Surjection as Surj using (Surjection; _↠_; module Surjection) open import Level open import Relation.Nullary.Product open import Relation.Binary import Relation.Binary.PropositionalEquality as P module _ {a₁ a₂ ℓ₁ ℓ₂} {A₁ : Set a₁} {A₂ : Set a₂} where infixr 2 _×-Rel_ _×-Rel_ : Rel A₁ ℓ₁ → Rel A₂ ℓ₂ → Rel (A₁ × A₂) _ _∼₁_ ×-Rel _∼₂_ = (_∼₁_ on proj₁) -×- (_∼₂_ on proj₂) -- Some properties which are preserved by ×-Rel (under certain -- assumptions). _×-reflexive_ : ∀ {_≈₁_ _∼₁_ _≈₂_ _∼₂_} → _≈₁_ ⇒ _∼₁_ → _≈₂_ ⇒ _∼₂_ → (_≈₁_ ×-Rel _≈₂_) ⇒ (_∼₁_ ×-Rel _∼₂_) refl₁ ×-reflexive refl₂ = λ x≈y → (refl₁ (proj₁ x≈y) , refl₂ (proj₂ x≈y)) _×-refl_ : ∀ {_∼₁_ _∼₂_} → Reflexive _∼₁_ → Reflexive _∼₂_ → Reflexive (_∼₁_ ×-Rel _∼₂_) refl₁ ×-refl refl₂ = (refl₁ , refl₂) ×-irreflexive₁ : ∀ {_≈₁_ _<₁_ _≈₂_ _<₂_} → Irreflexive _≈₁_ _<₁_ → Irreflexive (_≈₁_ ×-Rel _≈₂_) (_<₁_ ×-Rel _<₂_) ×-irreflexive₁ ir = λ x≈y x _ _ y (λ ()) (λ ()) halt cmp (x ∷ xs) (y ∷ ys) with tri x y ... | tri< x ¬x (¬≤-this ¬x≈y ¬x ¬xs (¬≤-next ¬x>=_ = λ x f → negated-stable (¬¬-map f x) } ¬¬-push : ∀ {p q} {P : Set p} {Q : P → Set q} → ¬ ¬ ((x : P) → Q x) → (x : P) → ¬ ¬ Q x ¬¬-push ¬¬P⟶Q P ¬Q = ¬¬P⟶Q (λ P⟶Q → ¬Q (P⟶Q P)) -- A double-negation-translated variant of excluded middle (or: every -- nullary relation is decidable in the double-negation monad). excluded-middle : ∀ {p} {P : Set p} → ¬ ¬ Dec P excluded-middle ¬h = ¬h (no (λ p → ¬h (yes p))) -- If Whatever is instantiated with ¬ ¬ something, then this function -- is call with current continuation in the double-negation monad, or, -- if you will, a double-negation translation of Peirce's law. -- -- In order to prove ¬ ¬ P one can assume ¬ P and prove ⊥. However, -- sometimes it is nice to avoid leaving the double-negation monad; in -- that case this function can be used (with Whatever instantiated to -- ⊥). call/cc : ∀ {w p} {Whatever : Set w} {P : Set p} → ((P → Whatever) → ¬ ¬ P) → ¬ ¬ P call/cc hyp ¬p = hyp (λ p → ⊥-elim (¬p p)) ¬p -- The "independence of premise" rule, in the double-negation monad. -- It is assumed that the index set (Q) is inhabited. independence-of-premise : ∀ {p q r} {P : Set p} {Q : Set q} {R : Q → Set r} → Q → (P → Σ Q R) → ¬ ¬ (Σ[ x ∈ Q ] (P → R x)) independence-of-premise {P = P} q f = ¬¬-map helper excluded-middle where helper : Dec P → _ helper (yes p) = Prod.map id const (f p) helper (no ¬p) = (q , ⊥-elim ∘′ ¬p) -- The independence of premise rule for binary sums. independence-of-premise-⊎ : ∀ {p q r} {P : Set p} {Q : Set q} {R : Set r} → (P → Q ⊎ R) → ¬ ¬ ((P → Q) ⊎ (P → R)) independence-of-premise-⊎ {P = P} f = ¬¬-map helper excluded-middle where helper : Dec P → _ helper (yes p) = Sum.map const const (f p) helper (no ¬p) = inj₁ (⊥-elim ∘′ ¬p) private -- Note that independence-of-premise-⊎ is a consequence of -- independence-of-premise (for simplicity it is assumed that Q and -- R have the same type here): corollary : ∀ {p ℓ} {P : Set p} {Q R : Set ℓ} → (P → Q ⊎ R) → ¬ ¬ ((P → Q) ⊎ (P → R)) corollary {P = P} {Q} {R} f = ¬¬-map helper (independence-of-premise true ([ _,_ true , _,_ false ] ∘′ f)) where helper : ∃ (λ b → P → if b then Q else R) → (P → Q) ⊎ (P → R) helper (true , f) = inj₁ f helper (false , f) = inj₂ f -- The classical statements of excluded middle and double-negation -- elimination. Excluded-Middle : (ℓ : Level) → Set (suc ℓ) Excluded-Middle p = {P : Set p} → Dec P Double-Negation-Elimination : (ℓ : Level) → Set (suc ℓ) Double-Negation-Elimination p = {P : Set p} → Stable P private -- The two statements above are equivalent. The proofs are so -- simple, given the definitions above, that they are not exported. em⇒dne : ∀ {ℓ} → Excluded-Middle ℓ → Double-Negation-Elimination ℓ em⇒dne em = decidable-stable em dne⇒em : ∀ {ℓ} → Double-Negation-Elimination ℓ → Excluded-Middle ℓ dne⇒em dne = dne excluded-middle lib-0.7/src/Relation/Nullary/Implication.agda0000644000000000000000000000115212101774706017400 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Implications of nullary relations ------------------------------------------------------------------------ module Relation.Nullary.Implication where open import Relation.Nullary open import Data.Empty -- Some properties which are preserved by _→_. infixr 2 _→-dec_ _→-dec_ : ∀ {p q} {P : Set p} {Q : Set q} → Dec P → Dec Q → Dec (P → Q) yes p →-dec no ¬q = no (λ f → ¬q (f p)) yes p →-dec yes q = yes (λ _ → q) no ¬p →-dec _ = yes (λ p → ⊥-elim (¬p p)) lib-0.7/src/Relation/Nullary/Decidable.agda0000644000000000000000000000466612101774706017001 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Operations on and properties of decidable relations ------------------------------------------------------------------------ module Relation.Nullary.Decidable where open import Data.Bool open import Data.Empty open import Data.Product hiding (map) open import Data.Unit open import Function open import Function.Equality using (_⟨$⟩_) open import Function.Equivalence using (_⇔_; equivalence; module Equivalence) open import Level using (Lift) open import Relation.Binary.PropositionalEquality open import Relation.Nullary ⌊_⌋ : ∀ {p} {P : Set p} → Dec P → Bool ⌊ yes _ ⌋ = true ⌊ no _ ⌋ = false True : ∀ {p} {P : Set p} → Dec P → Set True Q = T ⌊ Q ⌋ False : ∀ {p} {P : Set p} → Dec P → Set False Q = T (not ⌊ Q ⌋) -- Gives a witness to the "truth". toWitness : ∀ {p} {P : Set p} {Q : Dec P} → True Q → P toWitness {Q = yes p} _ = p toWitness {Q = no _} () -- Establishes a "truth", given a witness. fromWitness : ∀ {p} {P : Set p} {Q : Dec P} → P → True Q fromWitness {Q = yes p} = const _ fromWitness {Q = no ¬p} = ¬p -- Variants for False. toWitnessFalse : ∀ {p} {P : Set p} {Q : Dec P} → False Q → ¬ P toWitnessFalse {Q = yes _} () toWitnessFalse {Q = no ¬p} _ = ¬p fromWitnessFalse : ∀ {p} {P : Set p} {Q : Dec P} → ¬ P → False Q fromWitnessFalse {Q = yes p} = flip _$_ p fromWitnessFalse {Q = no ¬p} = const _ map : ∀ {p q} {P : Set p} {Q : Set q} → P ⇔ Q → Dec P → Dec Q map P⇔Q (yes p) = yes (Equivalence.to P⇔Q ⟨$⟩ p) map P⇔Q (no ¬p) = no (¬p ∘ _⟨$⟩_ (Equivalence.from P⇔Q)) map′ : ∀ {p q} {P : Set p} {Q : Set q} → (P → Q) → (Q → P) → Dec P → Dec Q map′ P→Q Q→P = map (equivalence P→Q Q→P) -- If a decision procedure returns "yes", then we can extract the -- proof using from-yes. From-yes : ∀ {p} {P : Set p} → Dec P → Set p From-yes {P = P} (yes _) = P From-yes (no _) = Lift ⊤ from-yes : ∀ {p} {P : Set p} (p : Dec P) → From-yes p from-yes (yes p) = p from-yes (no _) = _ -- If a decision procedure returns "no", then we can extract the proof -- using from-no. From-no : ∀ {p} {P : Set p} → Dec P → Set p From-no {P = P} (no _) = ¬ P From-no (yes _) = Lift ⊤ from-no : ∀ {p} {P : Set p} (p : Dec P) → From-no p from-no (no ¬p) = ¬p from-no (yes _) = _ lib-0.7/src/Relation/Nullary/Universe.agda0000644000000000000000000001117712101774706016740 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- A universe of proposition functors, along with some properties ------------------------------------------------------------------------ module Relation.Nullary.Universe where open import Relation.Nullary open import Relation.Nullary.Negation open import Relation.Binary hiding (_⇒_) open import Relation.Binary.Simple open import Relation.Binary.PropositionalEquality as PropEq using (_≡_; refl) open import Relation.Binary.Sum open import Relation.Binary.Product.Pointwise open import Data.Sum as Sum hiding (map) open import Data.Product as Prod hiding (map) open import Function import Function.Equality as FunS open import Data.Empty open import Category.Applicative open import Category.Monad open import Level infix 5 ¬¬_ infixr 4 _⇒_ infixr 3 _∧_ infixr 2 _∨_ infix 1 ⟨_⟩_≈_ -- The universe. data PropF p : Set (suc p) where Id : PropF p K : (P : Set p) → PropF p _∨_ : (F₁ F₂ : PropF p) → PropF p _∧_ : (F₁ F₂ : PropF p) → PropF p _⇒_ : (P₁ : Set p) (F₂ : PropF p) → PropF p ¬¬_ : (F : PropF p) → PropF p -- Equalities for universe inhabitants. mutual setoid : ∀ {p} → PropF p → Set p → Setoid p p setoid Id P = PropEq.setoid P setoid (K P) _ = PropEq.setoid P setoid (F₁ ∨ F₂) P = (setoid F₁ P) ⊎-setoid (setoid F₂ P) setoid (F₁ ∧ F₂) P = (setoid F₁ P) ×-setoid (setoid F₂ P) setoid (P₁ ⇒ F₂) P = FunS.≡-setoid P₁ (Setoid.indexedSetoid (setoid F₂ P)) setoid (¬¬ F) P = Always-setoid (¬ ¬ ⟦ F ⟧ P) ⟦_⟧ : ∀ {p} → PropF p → (Set p → Set p) ⟦ F ⟧ P = Setoid.Carrier (setoid F P) ⟨_⟩_≈_ : ∀ {p} (F : PropF p) {P : Set p} → Rel (⟦ F ⟧ P) p ⟨_⟩_≈_ F = Setoid._≈_ (setoid F _) -- ⟦ F ⟧ is functorial. map : ∀ {p} (F : PropF p) {P Q} → (P → Q) → ⟦ F ⟧ P → ⟦ F ⟧ Q map Id f p = f p map (K P) f p = p map (F₁ ∨ F₂) f FP = Sum.map (map F₁ f) (map F₂ f) FP map (F₁ ∧ F₂) f FP = Prod.map (map F₁ f) (map F₂ f) FP map (P₁ ⇒ F₂) f FP = map F₂ f ∘ FP map (¬¬ F) f FP = ¬¬-map (map F f) FP map-id : ∀ {p} (F : PropF p) {P} → ⟨ ⟦ F ⟧ P ⇒ F ⟩ map F id ≈ id map-id Id x = refl map-id (K P) x = refl map-id (F₁ ∨ F₂) (inj₁ x) = ₁∼₁ (map-id F₁ x) map-id (F₁ ∨ F₂) (inj₂ y) = ₂∼₂ (map-id F₂ y) map-id (F₁ ∧ F₂) (x , y) = (map-id F₁ x , map-id F₂ y) map-id (P₁ ⇒ F₂) f = λ x → map-id F₂ (f x) map-id (¬¬ F) ¬¬x = _ map-∘ : ∀ {p} (F : PropF p) {P Q R} (f : Q → R) (g : P → Q) → ⟨ ⟦ F ⟧ P ⇒ F ⟩ map F f ∘ map F g ≈ map F (f ∘ g) map-∘ Id f g x = refl map-∘ (K P) f g x = refl map-∘ (F₁ ∨ F₂) f g (inj₁ x) = ₁∼₁ (map-∘ F₁ f g x) map-∘ (F₁ ∨ F₂) f g (inj₂ y) = ₂∼₂ (map-∘ F₂ f g y) map-∘ (F₁ ∧ F₂) f g x = (map-∘ F₁ f g (proj₁ x) , map-∘ F₂ f g (proj₂ x)) map-∘ (P₁ ⇒ F₂) f g h = λ x → map-∘ F₂ f g (h x) map-∘ (¬¬ F) f g x = _ -- A variant of sequence can be implemented for ⟦ F ⟧. sequence : ∀ {p AF} → RawApplicative AF → (AF (Lift ⊥) → ⊥) → ({A B : Set p} → (A → AF B) → AF (A → B)) → ∀ F {P} → ⟦ F ⟧ (AF P) → AF (⟦ F ⟧ P) sequence {AF = AF} A extract-⊥ sequence-⇒ = helper where open RawApplicative A helper : ∀ F {P} → ⟦ F ⟧ (AF P) → AF (⟦ F ⟧ P) helper Id x = x helper (K P) x = pure x helper (F₁ ∨ F₂) (inj₁ x) = inj₁ <$> helper F₁ x helper (F₁ ∨ F₂) (inj₂ y) = inj₂ <$> helper F₂ y helper (F₁ ∧ F₂) (x , y) = _,_ <$> helper F₁ x ⊛ helper F₂ y helper (P₁ ⇒ F₂) f = sequence-⇒ (helper F₂ ∘ f) helper (¬¬ F) x = pure (λ ¬FP → x (λ fp → extract-⊥ (lift ∘ ¬FP <$> helper F fp))) -- Some lemmas about double negation. private open module M {p} = RawMonad (¬¬-Monad {p = p}) ¬¬-pull : ∀ {p} (F : PropF p) {P} → ⟦ F ⟧ (¬ ¬ P) → ¬ ¬ ⟦ F ⟧ P ¬¬-pull = sequence rawIApplicative (λ f → f lower) (λ f g → g (λ x → ⊥-elim (f x (λ y → g (λ _ → y))))) ¬¬-remove : ∀ {p} (F : PropF p) {P} → ¬ ¬ ⟦ F ⟧ (¬ ¬ P) → ¬ ¬ ⟦ F ⟧ P ¬¬-remove F = negated-stable ∘ ¬¬-pull (¬¬ F) lib-0.7/src/Level/0000755000000000000000000000000012101774706012157 5ustar0000000000000000lib-0.7/src/Induction/0000755000000000000000000000000012101774706013044 5ustar0000000000000000lib-0.7/src/Induction/WellFounded.agda0000644000000000000000000001113012101774706016066 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Well-founded induction ------------------------------------------------------------------------ open import Relation.Binary module Induction.WellFounded where open import Data.Product open import Function open import Induction open import Relation.Unary -- When using well-founded recursion you can recurse arbitrarily, as -- long as the arguments become smaller, and "smaller" is -- well-founded. WfRec : ∀ {a} {A : Set a} → Rel A a → RecStruct A WfRec _<_ P x = ∀ y → y < x → P y -- The accessibility predicate: x is accessible if everything which is -- smaller than x is also accessible (inductively). data Acc {a} {A : Set a} (_<_ : Rel A a) (x : A) : Set a where acc : (rs : WfRec _<_ (Acc _<_) x) → Acc _<_ x -- The accessibility predicate encodes what it means to be -- well-founded; if all elements are accessible, then _<_ is -- well-founded. Well-founded : ∀ {a} {A : Set a} → Rel A a → Set a Well-founded _<_ = ∀ x → Acc _<_ x -- Well-founded induction for the subset of accessible elements: module Some {a} {A : Set a} {_<_ : Rel A a} where wfRec-builder : SubsetRecursorBuilder (Acc _<_) (WfRec _<_) wfRec-builder P f x (acc rs) = λ y y>=_ postulate return : ∀ {a} {A : Set a} → A → IO A _>>=_ : ∀ {a b} {A : Set a} {B : Set b} → IO A → (A → IO B) → IO B {-# COMPILED return (\_ _ -> return) #-} {-# COMPILED _>>=_ (\_ _ _ _ -> (>>=)) #-} ------------------------------------------------------------------------ -- Simple lazy IO -- Note that the functions below produce commands which, when -- executed, may raise exceptions. -- Note also that the semantics of these functions depends on the -- version of the Haskell base library. If the version is 4.2.0.0 (or -- later?), then the functions use the character encoding specified by -- the locale. For older versions of the library (going back to at -- least version 3) the functions use ISO-8859-1. postulate getContents : IO Costring readFile : String → IO Costring writeFile : String → Costring → IO Unit appendFile : String → Costring → IO Unit putStr : Costring → IO Unit putStrLn : Costring → IO Unit -- Reads a finite file. Raises an exception if the file path refers -- to a non-physical file (like "/dev/zero"). readFiniteFile : String → IO String {-# COMPILED getContents getContents #-} {-# COMPILED readFile readFile #-} {-# COMPILED writeFile writeFile #-} {-# COMPILED appendFile appendFile #-} {-# COMPILED putStr putStr #-} {-# COMPILED putStrLn putStrLn #-} {-# COMPILED readFiniteFile IO.FFI.readFiniteFile #-} lib-0.7/src/Foreign/0000755000000000000000000000000012101774706012501 5ustar0000000000000000lib-0.7/src/Foreign/Haskell.agda0000644000000000000000000000054712101774706014710 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Type(s) used (only) when calling out to Haskell via the FFI ------------------------------------------------------------------------ module Foreign.Haskell where -- A unit type. data Unit : Set where unit : Unit {-# COMPILED_DATA Unit () () #-} lib-0.7/src/Algebra/0000755000000000000000000000000012101774706012445 5ustar0000000000000000lib-0.7/src/Algebra/Morphism.agda0000644000000000000000000000475412101774706015073 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Morphisms between algebraic structures ------------------------------------------------------------------------ module Algebra.Morphism where open import Relation.Binary open import Algebra open import Algebra.FunctionProperties import Algebra.Props.Group as GroupP open import Function open import Data.Product open import Level import Relation.Binary.EqReasoning as EqR ------------------------------------------------------------------------ -- Basic definitions module Definitions {f t ℓ} (From : Set f) (To : Set t) (_≈_ : Rel To ℓ) where Morphism : Set _ Morphism = From → To Homomorphic₀ : Morphism → From → To → Set _ Homomorphic₀ ⟦_⟧ ∙ ∘ = ⟦ ∙ ⟧ ≈ ∘ Homomorphic₁ : Morphism → Fun₁ From → Op₁ To → Set _ Homomorphic₁ ⟦_⟧ ∙_ ∘_ = ∀ x → ⟦ ∙ x ⟧ ≈ ∘ ⟦ x ⟧ Homomorphic₂ : Morphism → Fun₂ From → Op₂ To → Set _ Homomorphic₂ ⟦_⟧ _∙_ _∘_ = ∀ x y → ⟦ x ∙ y ⟧ ≈ (⟦ x ⟧ ∘ ⟦ y ⟧) ------------------------------------------------------------------------ -- An example showing how a morphism type can be defined -- Ring homomorphisms. record _-Ring⟶_ {r₁ r₂ r₃ r₄} (From : Ring r₁ r₂) (To : Ring r₃ r₄) : Set (r₁ ⊔ r₂ ⊔ r₃ ⊔ r₄) where private module F = Ring From module T = Ring To open Definitions F.Carrier T.Carrier T._≈_ field ⟦_⟧ : Morphism ⟦⟧-cong : ⟦_⟧ Preserves F._≈_ ⟶ T._≈_ +-homo : Homomorphic₂ ⟦_⟧ F._+_ T._+_ *-homo : Homomorphic₂ ⟦_⟧ F._*_ T._*_ 1-homo : Homomorphic₀ ⟦_⟧ F.1# T.1# open EqR T.setoid 0-homo : Homomorphic₀ ⟦_⟧ F.0# T.0# 0-homo = GroupP.left-identity-unique T.+-group ⟦ F.0# ⟧ ⟦ F.0# ⟧ (begin T._+_ ⟦ F.0# ⟧ ⟦ F.0# ⟧ ≈⟨ T.sym (+-homo F.0# F.0#) ⟩ ⟦ F._+_ F.0# F.0# ⟧ ≈⟨ ⟦⟧-cong (proj₁ F.+-identity F.0#) ⟩ ⟦ F.0# ⟧ ∎) -‿homo : Homomorphic₁ ⟦_⟧ F.-_ T.-_ -‿homo x = GroupP.left-inverse-unique T.+-group ⟦ F.-_ x ⟧ ⟦ x ⟧ (begin T._+_ ⟦ F.-_ x ⟧ ⟦ x ⟧ ≈⟨ T.sym (+-homo (F.-_ x) x) ⟩ ⟦ F._+_ (F.-_ x) x ⟧ ≈⟨ ⟦⟧-cong (proj₁ F.-‿inverse x) ⟩ ⟦ F.0# ⟧ ≈⟨ 0-homo ⟩ T.0# ∎) lib-0.7/src/Algebra/FunctionProperties.agda0000644000000000000000000000536212101774706017133 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Properties of functions, such as associativity and commutativity ------------------------------------------------------------------------ -- These properties can (for instance) be used to define algebraic -- structures. open import Level open import Relation.Binary -- The properties are specified using the following relation as -- "equality". module Algebra.FunctionProperties {a ℓ} {A : Set a} (_≈_ : Rel A ℓ) where open import Data.Product ------------------------------------------------------------------------ -- Unary and binary operations open import Algebra.FunctionProperties.Core public ------------------------------------------------------------------------ -- Properties of operations Associative : Op₂ A → Set _ Associative _∙_ = ∀ x y z → ((x ∙ y) ∙ z) ≈ (x ∙ (y ∙ z)) Commutative : Op₂ A → Set _ Commutative _∙_ = ∀ x y → (x ∙ y) ≈ (y ∙ x) LeftIdentity : A → Op₂ A → Set _ LeftIdentity e _∙_ = ∀ x → (e ∙ x) ≈ x RightIdentity : A → Op₂ A → Set _ RightIdentity e _∙_ = ∀ x → (x ∙ e) ≈ x Identity : A → Op₂ A → Set _ Identity e ∙ = LeftIdentity e ∙ × RightIdentity e ∙ LeftZero : A → Op₂ A → Set _ LeftZero z _∙_ = ∀ x → (z ∙ x) ≈ z RightZero : A → Op₂ A → Set _ RightZero z _∙_ = ∀ x → (x ∙ z) ≈ z Zero : A → Op₂ A → Set _ Zero z ∙ = LeftZero z ∙ × RightZero z ∙ LeftInverse : A → Op₁ A → Op₂ A → Set _ LeftInverse e _⁻¹ _∙_ = ∀ x → (x ⁻¹ ∙ x) ≈ e RightInverse : A → Op₁ A → Op₂ A → Set _ RightInverse e _⁻¹ _∙_ = ∀ x → (x ∙ (x ⁻¹)) ≈ e Inverse : A → Op₁ A → Op₂ A → Set _ Inverse e ⁻¹ ∙ = LeftInverse e ⁻¹ ∙ × RightInverse e ⁻¹ ∙ _DistributesOverˡ_ : Op₂ A → Op₂ A → Set _ _*_ DistributesOverˡ _+_ = ∀ x y z → (x * (y + z)) ≈ ((x * y) + (x * z)) _DistributesOverʳ_ : Op₂ A → Op₂ A → Set _ _*_ DistributesOverʳ _+_ = ∀ x y z → ((y + z) * x) ≈ ((y * x) + (z * x)) _DistributesOver_ : Op₂ A → Op₂ A → Set _ * DistributesOver + = (* DistributesOverˡ +) × (* DistributesOverʳ +) _IdempotentOn_ : Op₂ A → A → Set _ _∙_ IdempotentOn x = (x ∙ x) ≈ x Idempotent : Op₂ A → Set _ Idempotent ∙ = ∀ x → ∙ IdempotentOn x IdempotentFun : Op₁ A → Set _ IdempotentFun f = ∀ x → f (f x) ≈ f x _Absorbs_ : Op₂ A → Op₂ A → Set _ _∙_ Absorbs _∘_ = ∀ x y → (x ∙ (x ∘ y)) ≈ x Absorptive : Op₂ A → Op₂ A → Set _ Absorptive ∙ ∘ = (∙ Absorbs ∘) × (∘ Absorbs ∙) Involutive : Op₁ A → Set _ Involutive f = ∀ x → f (f x) ≈ x lib-0.7/src/Algebra/Structures.agda0000644000000000000000000003402612101774706015453 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Some algebraic structures (not packed up with sets, operations, -- etc.) ------------------------------------------------------------------------ open import Relation.Binary module Algebra.Structures where import Algebra.FunctionProperties as FunctionProperties open import Data.Product open import Function open import Level using (_⊔_) import Relation.Binary.EqReasoning as EqR open FunctionProperties using (Op₁; Op₂) ------------------------------------------------------------------------ -- One binary operation record IsSemigroup {a ℓ} {A : Set a} (≈ : Rel A ℓ) (∙ : Op₂ A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field isEquivalence : IsEquivalence ≈ assoc : Associative ∙ ∙-cong : ∙ Preserves₂ ≈ ⟶ ≈ ⟶ ≈ open IsEquivalence isEquivalence public record IsMonoid {a ℓ} {A : Set a} (≈ : Rel A ℓ) (∙ : Op₂ A) (ε : A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field isSemigroup : IsSemigroup ≈ ∙ identity : Identity ε ∙ open IsSemigroup isSemigroup public record IsCommutativeMonoid {a ℓ} {A : Set a} (≈ : Rel A ℓ) (_∙_ : Op₂ A) (ε : A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field isSemigroup : IsSemigroup ≈ _∙_ identityˡ : LeftIdentity ε _∙_ comm : Commutative _∙_ open IsSemigroup isSemigroup public identity : Identity ε _∙_ identity = (identityˡ , identityʳ) where open EqR (record { isEquivalence = isEquivalence }) identityʳ : RightIdentity ε _∙_ identityʳ = λ x → begin (x ∙ ε) ≈⟨ comm x ε ⟩ (ε ∙ x) ≈⟨ identityˡ x ⟩ x ∎ isMonoid : IsMonoid ≈ _∙_ ε isMonoid = record { isSemigroup = isSemigroup ; identity = identity } record IsGroup {a ℓ} {A : Set a} (≈ : Rel A ℓ) (_∙_ : Op₂ A) (ε : A) (_⁻¹ : Op₁ A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ infixl 7 _-_ field isMonoid : IsMonoid ≈ _∙_ ε inverse : Inverse ε _⁻¹ _∙_ ⁻¹-cong : _⁻¹ Preserves ≈ ⟶ ≈ open IsMonoid isMonoid public _-_ : Op₂ A x - y = x ∙ (y ⁻¹) record IsAbelianGroup {a ℓ} {A : Set a} (≈ : Rel A ℓ) (∙ : Op₂ A) (ε : A) (⁻¹ : Op₁ A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field isGroup : IsGroup ≈ ∙ ε ⁻¹ comm : Commutative ∙ open IsGroup isGroup public isCommutativeMonoid : IsCommutativeMonoid ≈ ∙ ε isCommutativeMonoid = record { isSemigroup = isSemigroup ; identityˡ = proj₁ identity ; comm = comm } ------------------------------------------------------------------------ -- Two binary operations record IsNearSemiring {a ℓ} {A : Set a} (≈ : Rel A ℓ) (+ * : Op₂ A) (0# : A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field +-isMonoid : IsMonoid ≈ + 0# *-isSemigroup : IsSemigroup ≈ * distribʳ : * DistributesOverʳ + zeroˡ : LeftZero 0# * open IsMonoid +-isMonoid public renaming ( assoc to +-assoc ; ∙-cong to +-cong ; isSemigroup to +-isSemigroup ; identity to +-identity ) open IsSemigroup *-isSemigroup public using () renaming ( assoc to *-assoc ; ∙-cong to *-cong ) record IsSemiringWithoutOne {a ℓ} {A : Set a} (≈ : Rel A ℓ) (+ * : Op₂ A) (0# : A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field +-isCommutativeMonoid : IsCommutativeMonoid ≈ + 0# *-isSemigroup : IsSemigroup ≈ * distrib : * DistributesOver + zero : Zero 0# * open IsCommutativeMonoid +-isCommutativeMonoid public hiding (identityˡ) renaming ( assoc to +-assoc ; ∙-cong to +-cong ; isSemigroup to +-isSemigroup ; identity to +-identity ; isMonoid to +-isMonoid ; comm to +-comm ) open IsSemigroup *-isSemigroup public using () renaming ( assoc to *-assoc ; ∙-cong to *-cong ) isNearSemiring : IsNearSemiring ≈ + * 0# isNearSemiring = record { +-isMonoid = +-isMonoid ; *-isSemigroup = *-isSemigroup ; distribʳ = proj₂ distrib ; zeroˡ = proj₁ zero } record IsSemiringWithoutAnnihilatingZero {a ℓ} {A : Set a} (≈ : Rel A ℓ) (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field -- Note that these structures do have an additive unit, but this -- unit does not necessarily annihilate multiplication. +-isCommutativeMonoid : IsCommutativeMonoid ≈ + 0# *-isMonoid : IsMonoid ≈ * 1# distrib : * DistributesOver + open IsCommutativeMonoid +-isCommutativeMonoid public hiding (identityˡ) renaming ( assoc to +-assoc ; ∙-cong to +-cong ; isSemigroup to +-isSemigroup ; identity to +-identity ; isMonoid to +-isMonoid ; comm to +-comm ) open IsMonoid *-isMonoid public using () renaming ( assoc to *-assoc ; ∙-cong to *-cong ; isSemigroup to *-isSemigroup ; identity to *-identity ) record IsSemiring {a ℓ} {A : Set a} (≈ : Rel A ℓ) (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field isSemiringWithoutAnnihilatingZero : IsSemiringWithoutAnnihilatingZero ≈ + * 0# 1# zero : Zero 0# * open IsSemiringWithoutAnnihilatingZero isSemiringWithoutAnnihilatingZero public isSemiringWithoutOne : IsSemiringWithoutOne ≈ + * 0# isSemiringWithoutOne = record { +-isCommutativeMonoid = +-isCommutativeMonoid ; *-isSemigroup = *-isSemigroup ; distrib = distrib ; zero = zero } open IsSemiringWithoutOne isSemiringWithoutOne public using (isNearSemiring) record IsCommutativeSemiringWithoutOne {a ℓ} {A : Set a} (≈ : Rel A ℓ) (+ * : Op₂ A) (0# : A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field isSemiringWithoutOne : IsSemiringWithoutOne ≈ + * 0# *-comm : Commutative * open IsSemiringWithoutOne isSemiringWithoutOne public record IsCommutativeSemiring {a ℓ} {A : Set a} (≈ : Rel A ℓ) (_+_ _*_ : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field +-isCommutativeMonoid : IsCommutativeMonoid ≈ _+_ 0# *-isCommutativeMonoid : IsCommutativeMonoid ≈ _*_ 1# distribʳ : _*_ DistributesOverʳ _+_ zeroˡ : LeftZero 0# _*_ private module +-CM = IsCommutativeMonoid +-isCommutativeMonoid open module *-CM = IsCommutativeMonoid *-isCommutativeMonoid public using () renaming (comm to *-comm) open EqR (record { isEquivalence = +-CM.isEquivalence }) distrib : _*_ DistributesOver _+_ distrib = (distribˡ , distribʳ) where distribˡ : _*_ DistributesOverˡ _+_ distribˡ x y z = begin (x * (y + z)) ≈⟨ *-comm x (y + z) ⟩ ((y + z) * x) ≈⟨ distribʳ x y z ⟩ ((y * x) + (z * x)) ≈⟨ *-comm y x ⟨ +-CM.∙-cong ⟩ *-comm z x ⟩ ((x * y) + (x * z)) ∎ zero : Zero 0# _*_ zero = (zeroˡ , zeroʳ) where zeroʳ : RightZero 0# _*_ zeroʳ x = begin (x * 0#) ≈⟨ *-comm x 0# ⟩ (0# * x) ≈⟨ zeroˡ x ⟩ 0# ∎ isSemiring : IsSemiring ≈ _+_ _*_ 0# 1# isSemiring = record { isSemiringWithoutAnnihilatingZero = record { +-isCommutativeMonoid = +-isCommutativeMonoid ; *-isMonoid = *-CM.isMonoid ; distrib = distrib } ; zero = zero } open IsSemiring isSemiring public hiding (distrib; zero; +-isCommutativeMonoid) isCommutativeSemiringWithoutOne : IsCommutativeSemiringWithoutOne ≈ _+_ _*_ 0# isCommutativeSemiringWithoutOne = record { isSemiringWithoutOne = isSemiringWithoutOne ; *-comm = *-CM.comm } record IsRing {a ℓ} {A : Set a} (≈ : Rel A ℓ) (_+_ _*_ : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field +-isAbelianGroup : IsAbelianGroup ≈ _+_ 0# -_ *-isMonoid : IsMonoid ≈ _*_ 1# distrib : _*_ DistributesOver _+_ open IsAbelianGroup +-isAbelianGroup public renaming ( assoc to +-assoc ; ∙-cong to +-cong ; isSemigroup to +-isSemigroup ; identity to +-identity ; isMonoid to +-isMonoid ; inverse to -‿inverse ; ⁻¹-cong to -‿cong ; isGroup to +-isGroup ; comm to +-comm ; isCommutativeMonoid to +-isCommutativeMonoid ) open IsMonoid *-isMonoid public using () renaming ( assoc to *-assoc ; ∙-cong to *-cong ; isSemigroup to *-isSemigroup ; identity to *-identity ) zero : Zero 0# _*_ zero = (zeroˡ , zeroʳ) where open EqR (record { isEquivalence = isEquivalence }) zeroˡ : LeftZero 0# _*_ zeroˡ x = begin 0# * x ≈⟨ sym $ proj₂ +-identity _ ⟩ (0# * x) + 0# ≈⟨ refl ⟨ +-cong ⟩ sym (proj₂ -‿inverse _) ⟩ (0# * x) + ((0# * x) + - (0# * x)) ≈⟨ sym $ +-assoc _ _ _ ⟩ ((0# * x) + (0# * x)) + - (0# * x) ≈⟨ sym (proj₂ distrib _ _ _) ⟨ +-cong ⟩ refl ⟩ ((0# + 0#) * x) + - (0# * x) ≈⟨ (proj₂ +-identity _ ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ refl ⟩ (0# * x) + - (0# * x) ≈⟨ proj₂ -‿inverse _ ⟩ 0# ∎ zeroʳ : RightZero 0# _*_ zeroʳ x = begin x * 0# ≈⟨ sym $ proj₂ +-identity _ ⟩ (x * 0#) + 0# ≈⟨ refl ⟨ +-cong ⟩ sym (proj₂ -‿inverse _) ⟩ (x * 0#) + ((x * 0#) + - (x * 0#)) ≈⟨ sym $ +-assoc _ _ _ ⟩ ((x * 0#) + (x * 0#)) + - (x * 0#) ≈⟨ sym (proj₁ distrib _ _ _) ⟨ +-cong ⟩ refl ⟩ (x * (0# + 0#)) + - (x * 0#) ≈⟨ (refl ⟨ *-cong ⟩ proj₂ +-identity _) ⟨ +-cong ⟩ refl ⟩ (x * 0#) + - (x * 0#) ≈⟨ proj₂ -‿inverse _ ⟩ 0# ∎ isSemiringWithoutAnnihilatingZero : IsSemiringWithoutAnnihilatingZero ≈ _+_ _*_ 0# 1# isSemiringWithoutAnnihilatingZero = record { +-isCommutativeMonoid = +-isCommutativeMonoid ; *-isMonoid = *-isMonoid ; distrib = distrib } isSemiring : IsSemiring ≈ _+_ _*_ 0# 1# isSemiring = record { isSemiringWithoutAnnihilatingZero = isSemiringWithoutAnnihilatingZero ; zero = zero } open IsSemiring isSemiring public using (isNearSemiring; isSemiringWithoutOne) record IsCommutativeRing {a ℓ} {A : Set a} (≈ : Rel A ℓ) (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field isRing : IsRing ≈ + * - 0# 1# *-comm : Commutative * open IsRing isRing public isCommutativeSemiring : IsCommutativeSemiring ≈ + * 0# 1# isCommutativeSemiring = record { +-isCommutativeMonoid = +-isCommutativeMonoid ; *-isCommutativeMonoid = record { isSemigroup = *-isSemigroup ; identityˡ = proj₁ *-identity ; comm = *-comm } ; distribʳ = proj₂ distrib ; zeroˡ = proj₁ zero } open IsCommutativeSemiring isCommutativeSemiring public using ( *-isCommutativeMonoid ; isCommutativeSemiringWithoutOne ) record IsLattice {a ℓ} {A : Set a} (≈ : Rel A ℓ) (∨ ∧ : Op₂ A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field isEquivalence : IsEquivalence ≈ ∨-comm : Commutative ∨ ∨-assoc : Associative ∨ ∨-cong : ∨ Preserves₂ ≈ ⟶ ≈ ⟶ ≈ ∧-comm : Commutative ∧ ∧-assoc : Associative ∧ ∧-cong : ∧ Preserves₂ ≈ ⟶ ≈ ⟶ ≈ absorptive : Absorptive ∨ ∧ open IsEquivalence isEquivalence public record IsDistributiveLattice {a ℓ} {A : Set a} (≈ : Rel A ℓ) (∨ ∧ : Op₂ A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field isLattice : IsLattice ≈ ∨ ∧ ∨-∧-distribʳ : ∨ DistributesOverʳ ∧ open IsLattice isLattice public record IsBooleanAlgebra {a ℓ} {A : Set a} (≈ : Rel A ℓ) (∨ ∧ : Op₂ A) (¬ : Op₁ A) (⊤ ⊥ : A) : Set (a ⊔ ℓ) where open FunctionProperties ≈ field isDistributiveLattice : IsDistributiveLattice ≈ ∨ ∧ ∨-complementʳ : RightInverse ⊤ ¬ ∨ ∧-complementʳ : RightInverse ⊥ ¬ ∧ ¬-cong : ¬ Preserves ≈ ⟶ ≈ open IsDistributiveLattice isDistributiveLattice public lib-0.7/src/Algebra/Operations.agda0000644000000000000000000001156412101774706015415 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Some defined operations (multiplication by natural number and -- exponentiation) ------------------------------------------------------------------------ open import Algebra module Algebra.Operations {s₁ s₂} (S : Semiring s₁ s₂) where open Semiring S renaming (zero to *-zero) open import Data.Nat using (zero; suc; ℕ) renaming (_+_ to _ℕ+_; _*_ to _ℕ*_) open import Data.Product using (module Σ) open import Function open import Relation.Binary open import Relation.Binary.PropositionalEquality as PropEq using (_≡_) import Relation.Binary.EqReasoning as EqR open EqR setoid ------------------------------------------------------------------------ -- Operations -- Multiplication by natural number. infixr 7 _×_ _×_ : ℕ → Carrier → Carrier 0 × x = 0# suc n × x = x + n × x -- A variant that includes a "redundant" case which ensures that 1 × y -- is definitionally equal to y. _×′_ : ℕ → Carrier → Carrier 0 ×′ x = 0# 1 ×′ x = x suc n ×′ x = x + n ×′ x -- Exponentiation. infixr 8 _^_ _^_ : Carrier → ℕ → Carrier x ^ zero = 1# x ^ suc n = x * x ^ n ------------------------------------------------------------------------ -- Some properties -- Unfolding lemma for _×′_. 1+×′ : ∀ n x → suc n ×′ x ≈ x + n ×′ x 1+×′ 0 x = begin x ≈⟨ sym $ Σ.proj₂ +-identity x ⟩ x + 0# ∎ 1+×′ (suc n) x = begin x + suc n ×′ x ≡⟨⟩ x + suc n ×′ x ∎ -- _×_ and _×′_ are extensionally equal (up to the setoid -- equivalence). ×≈×′ : ∀ n x → n × x ≈ n ×′ x ×≈×′ 0 x = begin 0# ∎ ×≈×′ (suc n) x = begin x + n × x ≈⟨ +-cong refl (×≈×′ n x) ⟩ x + n ×′ x ≈⟨ sym $ 1+×′ n x ⟩ suc n ×′ x ∎ -- _×_ is homomorphic with respect to _ℕ+_/_+_. ×-homo-+ : ∀ c m n → (m ℕ+ n) × c ≈ m × c + n × c ×-homo-+ c 0 n = begin n × c ≈⟨ sym $ Σ.proj₁ +-identity (n × c) ⟩ 0# + n × c ∎ ×-homo-+ c (suc m) n = begin c + (m ℕ+ n) × c ≈⟨ +-cong refl (×-homo-+ c m n) ⟩ c + (m × c + n × c) ≈⟨ sym $ +-assoc c (m × c) (n × c) ⟩ c + m × c + n × c ∎ -- _×′_ is homomorphic with respect to _ℕ+_/_+_. ×′-homo-+ : ∀ c m n → (m ℕ+ n) ×′ c ≈ m ×′ c + n ×′ c ×′-homo-+ c m n = begin (m ℕ+ n) ×′ c ≈⟨ sym $ ×≈×′ (m ℕ+ n) c ⟩ (m ℕ+ n) × c ≈⟨ ×-homo-+ c m n ⟩ m × c + n × c ≈⟨ +-cong (×≈×′ m c) (×≈×′ n c) ⟩ m ×′ c + n ×′ c ∎ -- _× 1# is homomorphic with respect to _ℕ*_/_*_. ×1-homo-* : ∀ m n → (m ℕ* n) × 1# ≈ (m × 1#) * (n × 1#) ×1-homo-* 0 n = begin 0# ≈⟨ sym $ Σ.proj₁ *-zero (n × 1#) ⟩ 0# * (n × 1#) ∎ ×1-homo-* (suc m) n = begin (n ℕ+ m ℕ* n) × 1# ≈⟨ ×-homo-+ 1# n (m ℕ* n) ⟩ n × 1# + (m ℕ* n) × 1# ≈⟨ +-cong refl (×1-homo-* m n) ⟩ n × 1# + (m × 1#) * (n × 1#) ≈⟨ sym $ +-cong (Σ.proj₁ *-identity (n × 1#)) refl ⟩ 1# * (n × 1#) + (m × 1#) * (n × 1#) ≈⟨ sym $ Σ.proj₂ distrib (n × 1#) 1# (m × 1#) ⟩ (1# + m × 1#) * (n × 1#) ∎ -- _×′ 1# is homomorphic with respect to _ℕ*_/_*_. ×′1-homo-* : ∀ m n → (m ℕ* n) ×′ 1# ≈ (m ×′ 1#) * (n ×′ 1#) ×′1-homo-* m n = begin (m ℕ* n) ×′ 1# ≈⟨ sym $ ×≈×′ (m ℕ* n) 1# ⟩ (m ℕ* n) × 1# ≈⟨ ×1-homo-* m n ⟩ (m × 1#) * (n × 1#) ≈⟨ *-cong (×≈×′ m 1#) (×≈×′ n 1#) ⟩ (m ×′ 1#) * (n ×′ 1#) ∎ -- _×_ preserves equality. ×-cong : _×_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_ ×-cong {n} {n′} {x} {x′} n≡n′ x≈x′ = begin n × x ≈⟨ reflexive $ PropEq.cong (λ n → n × x) n≡n′ ⟩ n′ × x ≈⟨ ×-congʳ n′ x≈x′ ⟩ n′ × x′ ∎ where ×-congʳ : ∀ n → (_×_ n) Preserves _≈_ ⟶ _≈_ ×-congʳ 0 x≈x′ = refl ×-congʳ (suc n) x≈x′ = x≈x′ ⟨ +-cong ⟩ ×-congʳ n x≈x′ -- _×′_ preserves equality. ×′-cong : _×′_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_ ×′-cong {n} {n′} {x} {x′} n≡n′ x≈x′ = begin n ×′ x ≈⟨ sym $ ×≈×′ n x ⟩ n × x ≈⟨ ×-cong n≡n′ x≈x′ ⟩ n′ × x′ ≈⟨ ×≈×′ n′ x′ ⟩ n′ ×′ x′ ∎ -- _^_ preserves equality. ^-cong : _^_ Preserves₂ _≈_ ⟶ _≡_ ⟶ _≈_ ^-cong {x} {x'} {n} {n'} x≈x' n≡n' = begin x ^ n ≈⟨ reflexive $ PropEq.cong (_^_ x) n≡n' ⟩ x ^ n' ≈⟨ ^-congˡ n' x≈x' ⟩ x' ^ n' ∎ where ^-congˡ : ∀ n → (λ x → x ^ n) Preserves _≈_ ⟶ _≈_ ^-congˡ zero x≈x' = refl ^-congˡ (suc n) x≈x' = x≈x' ⟨ *-cong ⟩ ^-congˡ n x≈x' lib-0.7/src/Algebra/RingSolver.agda0000644000000000000000000005442712101774706015371 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Solver for commutative ring or semiring equalities ------------------------------------------------------------------------ -- Uses ideas from the Coq ring tactic. See "Proving Equalities in a -- Commutative Ring Done Right in Coq" by Grégoire and Mahboubi. The -- code below is not optimised like theirs, though (in particular, our -- Horner normal forms are not sparse). open import Algebra open import Algebra.RingSolver.AlmostCommutativeRing open import Relation.Binary module Algebra.RingSolver {r₁ r₂ r₃} (Coeff : RawRing r₁) -- Coefficient "ring". (R : AlmostCommutativeRing r₂ r₃) -- Main "ring". (morphism : Coeff -Raw-AlmostCommutative⟶ R) (_coeff≟_ : Decidable (Induced-equivalence morphism)) where import Algebra.RingSolver.Lemmas as L; open L Coeff R morphism private module C = RawRing Coeff open AlmostCommutativeRing R renaming (zero to zero*) import Algebra.FunctionProperties as P; open P _≈_ open import Algebra.Morphism open _-Raw-AlmostCommutative⟶_ morphism renaming (⟦_⟧ to ⟦_⟧′) import Algebra.Operations as Ops; open Ops semiring open import Relation.Binary open import Relation.Nullary.Core import Relation.Binary.EqReasoning as EqR; open EqR setoid import Relation.Binary.PropositionalEquality as PropEq import Relation.Binary.Reflection as Reflection open import Data.Empty open import Data.Product open import Data.Nat as Nat using (ℕ; suc; zero) open import Data.Fin as Fin using (Fin; zero; suc) open import Data.Vec open import Function open import Level using (_⊔_) infix 9 :-_ -H_ -N_ infixr 9 _:^_ _^N_ infix 8 _*x+_ _*x+HN_ _*x+H_ infixl 8 _:*_ _*N_ _*H_ _*NH_ _*HN_ infixl 7 _:+_ _:-_ _+H_ _+N_ infix 4 _≈H_ _≈N_ ------------------------------------------------------------------------ -- Polynomials data Op : Set where [+] : Op [*] : Op -- The polynomials are indexed by the number of variables. data Polynomial (m : ℕ) : Set r₁ where op : (o : Op) (p₁ : Polynomial m) (p₂ : Polynomial m) → Polynomial m con : (c : C.Carrier) → Polynomial m var : (x : Fin m) → Polynomial m _:^_ : (p : Polynomial m) (n : ℕ) → Polynomial m :-_ : (p : Polynomial m) → Polynomial m -- Short-hand notation. _:+_ : ∀ {n} → Polynomial n → Polynomial n → Polynomial n _:+_ = op [+] _:*_ : ∀ {n} → Polynomial n → Polynomial n → Polynomial n _:*_ = op [*] _:-_ : ∀ {n} → Polynomial n → Polynomial n → Polynomial n x :- y = x :+ :- y -- Semantics. sem : Op → Op₂ Carrier sem [+] = _+_ sem [*] = _*_ ⟦_⟧ : ∀ {n} → Polynomial n → Vec Carrier n → Carrier ⟦ op o p₁ p₂ ⟧ ρ = ⟦ p₁ ⟧ ρ ⟨ sem o ⟩ ⟦ p₂ ⟧ ρ ⟦ con c ⟧ ρ = ⟦ c ⟧′ ⟦ var x ⟧ ρ = lookup x ρ ⟦ p :^ n ⟧ ρ = ⟦ p ⟧ ρ ^ n ⟦ :- p ⟧ ρ = - ⟦ p ⟧ ρ ------------------------------------------------------------------------ -- Normal forms of polynomials -- A univariate polynomial of degree d, -- -- p = a_d x^d + a_{d-1}x^{d-1} + … + a_0, -- -- is represented in Horner normal form by -- -- p = ((a_d x + a_{d-1})x + …)x + a_0. -- -- Note that Horner normal forms can be represented as lists, with the -- empty list standing for the zero polynomial of degree "-1". -- -- Given this representation of univariate polynomials over an -- arbitrary ring, polynomials in any number of variables over the -- ring C can be represented via the isomorphisms -- -- C[] ≅ C -- -- and -- -- C[X_0,...X_{n+1}] ≅ C[X_0,...,X_n][X_{n+1}]. mutual -- The polynomial representations are indexed by the polynomial's -- degree. data HNF : ℕ → Set r₁ where ∅ : ∀ {n} → HNF (suc n) _*x+_ : ∀ {n} → HNF (suc n) → Normal n → HNF (suc n) data Normal : ℕ → Set r₁ where con : C.Carrier → Normal zero poly : ∀ {n} → HNF (suc n) → Normal (suc n) -- Note that the data types above do /not/ ensure uniqueness of -- normal forms: the zero polynomial of degree one can be -- represented using both ∅ and ∅ *x+ con C.0#. mutual -- Semantics. ⟦_⟧H : ∀ {n} → HNF (suc n) → Vec Carrier (suc n) → Carrier ⟦ ∅ ⟧H _ = 0# ⟦ p *x+ c ⟧H (x ∷ ρ) = ⟦ p ⟧H (x ∷ ρ) * x + ⟦ c ⟧N ρ ⟦_⟧N : ∀ {n} → Normal n → Vec Carrier n → Carrier ⟦ con c ⟧N _ = ⟦ c ⟧′ ⟦ poly p ⟧N ρ = ⟦ p ⟧H ρ ------------------------------------------------------------------------ -- Equality and decidability mutual -- Equality. data _≈H_ : ∀ {n} → HNF n → HNF n → Set (r₁ ⊔ r₃) where ∅ : ∀ {n} → _≈H_ {suc n} ∅ ∅ _*x+_ : ∀ {n} {p₁ p₂ : HNF (suc n)} {c₁ c₂ : Normal n} → p₁ ≈H p₂ → c₁ ≈N c₂ → (p₁ *x+ c₁) ≈H (p₂ *x+ c₂) data _≈N_ : ∀ {n} → Normal n → Normal n → Set (r₁ ⊔ r₃) where con : ∀ {c₁ c₂} → ⟦ c₁ ⟧′ ≈ ⟦ c₂ ⟧′ → con c₁ ≈N con c₂ poly : ∀ {n} {p₁ p₂ : HNF (suc n)} → p₁ ≈H p₂ → poly p₁ ≈N poly p₂ mutual -- Equality is decidable. _≟H_ : ∀ {n} → Decidable (_≈H_ {n = n}) ∅ ≟H ∅ = yes ∅ ∅ ≟H (_ *x+ _) = no λ() (_ *x+ _) ≟H ∅ = no λ() (p₁ *x+ c₁) ≟H (p₂ *x+ c₂) with p₁ ≟H p₂ | c₁ ≟N c₂ ... | yes p₁≈p₂ | yes c₁≈c₂ = yes (p₁≈p₂ *x+ c₁≈c₂) ... | _ | no c₁≉c₂ = no λ { (_ *x+ c₁≈c₂) → c₁≉c₂ c₁≈c₂ } ... | no p₁≉p₂ | _ = no λ { (p₁≈p₂ *x+ _) → p₁≉p₂ p₁≈p₂ } _≟N_ : ∀ {n} → Decidable (_≈N_ {n = n}) con c₁ ≟N con c₂ with c₁ coeff≟ c₂ ... | yes c₁≈c₂ = yes (con c₁≈c₂) ... | no c₁≉c₂ = no λ { (con c₁≈c₂) → c₁≉c₂ c₁≈c₂} poly p₁ ≟N poly p₂ with p₁ ≟H p₂ ... | yes p₁≈p₂ = yes (poly p₁≈p₂) ... | no p₁≉p₂ = no λ { (poly p₁≈p₂) → p₁≉p₂ p₁≈p₂ } mutual -- The semantics respect the equality relations defined above. ⟦_⟧H-cong : ∀ {n} {p₁ p₂ : HNF (suc n)} → p₁ ≈H p₂ → ∀ ρ → ⟦ p₁ ⟧H ρ ≈ ⟦ p₂ ⟧H ρ ⟦ ∅ ⟧H-cong _ = refl ⟦ p₁≈p₂ *x+ c₁≈c₂ ⟧H-cong (x ∷ ρ) = (⟦ p₁≈p₂ ⟧H-cong (x ∷ ρ) ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ ⟦ c₁≈c₂ ⟧N-cong ρ ⟦_⟧N-cong : ∀ {n} {p₁ p₂ : Normal n} → p₁ ≈N p₂ → ∀ ρ → ⟦ p₁ ⟧N ρ ≈ ⟦ p₂ ⟧N ρ ⟦ con c₁≈c₂ ⟧N-cong _ = c₁≈c₂ ⟦ poly p₁≈p₂ ⟧N-cong ρ = ⟦ p₁≈p₂ ⟧H-cong ρ ------------------------------------------------------------------------ -- Ring operations on Horner normal forms -- Zero. 0H : ∀ {n} → HNF (suc n) 0H = ∅ 0N : ∀ {n} → Normal n 0N {zero} = con C.0# 0N {suc n} = poly 0H mutual -- One. 1H : ∀ {n} → HNF (suc n) 1H {n} = ∅ *x+ 1N {n} 1N : ∀ {n} → Normal n 1N {zero} = con C.1# 1N {suc n} = poly 1H -- A simplifying variant of _*x+_. _*x+HN_ : ∀ {n} → HNF (suc n) → Normal n → HNF (suc n) (p *x+ c′) *x+HN c = (p *x+ c′) *x+ c ∅ *x+HN c with c ≟N 0N ... | yes c≈0 = ∅ ... | no c≉0 = ∅ *x+ c mutual -- Addition. _+H_ : ∀ {n} → HNF (suc n) → HNF (suc n) → HNF (suc n) ∅ +H p = p p +H ∅ = p (p₁ *x+ c₁) +H (p₂ *x+ c₂) = (p₁ +H p₂) *x+HN (c₁ +N c₂) _+N_ : ∀ {n} → Normal n → Normal n → Normal n con c₁ +N con c₂ = con (c₁ C.+ c₂) poly p₁ +N poly p₂ = poly (p₁ +H p₂) -- Multiplication. _*x+H_ : ∀ {n} → HNF (suc n) → HNF (suc n) → HNF (suc n) p₁ *x+H (p₂ *x+ c) = (p₁ +H p₂) *x+HN c ∅ *x+H ∅ = ∅ (p₁ *x+ c) *x+H ∅ = (p₁ *x+ c) *x+ 0N mutual _*NH_ : ∀ {n} → Normal n → HNF (suc n) → HNF (suc n) c *NH ∅ = ∅ c *NH (p *x+ c′) with c ≟N 0N ... | yes c≈0 = ∅ ... | no c≉0 = (c *NH p) *x+ (c *N c′) _*HN_ : ∀ {n} → HNF (suc n) → Normal n → HNF (suc n) ∅ *HN c = ∅ (p *x+ c′) *HN c with c ≟N 0N ... | yes c≈0 = ∅ ... | no c≉0 = (p *HN c) *x+ (c′ *N c) _*H_ : ∀ {n} → HNF (suc n) → HNF (suc n) → HNF (suc n) ∅ *H _ = ∅ (_ *x+ _) *H ∅ = ∅ (p₁ *x+ c₁) *H (p₂ *x+ c₂) = ((p₁ *H p₂) *x+H (p₁ *HN c₂ +H c₁ *NH p₂)) *x+HN (c₁ *N c₂) _*N_ : ∀ {n} → Normal n → Normal n → Normal n con c₁ *N con c₂ = con (c₁ C.* c₂) poly p₁ *N poly p₂ = poly (p₁ *H p₂) -- Exponentiation. _^N_ : ∀ {n} → Normal n → ℕ → Normal n p ^N zero = 1N p ^N suc n = p *N (p ^N n) mutual -- Negation. -H_ : ∀ {n} → HNF (suc n) → HNF (suc n) -H p = (-N 1N) *NH p -N_ : ∀ {n} → Normal n → Normal n -N con c = con (C.- c) -N poly p = poly (-H p) ------------------------------------------------------------------------ -- Normalisation normalise-con : ∀ {n} → C.Carrier → Normal n normalise-con {zero} c = con c normalise-con {suc n} c = poly (∅ *x+HN normalise-con c) normalise-var : ∀ {n} → Fin n → Normal n normalise-var zero = poly ((∅ *x+ 1N) *x+ 0N) normalise-var (suc i) = poly (∅ *x+HN normalise-var i) normalise : ∀ {n} → Polynomial n → Normal n normalise (op [+] t₁ t₂) = normalise t₁ +N normalise t₂ normalise (op [*] t₁ t₂) = normalise t₁ *N normalise t₂ normalise (con c) = normalise-con c normalise (var i) = normalise-var i normalise (t :^ k) = normalise t ^N k normalise (:- t) = -N normalise t -- Evaluation after normalisation. ⟦_⟧↓ : ∀ {n} → Polynomial n → Vec Carrier n → Carrier ⟦ p ⟧↓ ρ = ⟦ normalise p ⟧N ρ ------------------------------------------------------------------------ -- Homomorphism lemmas 0N-homo : ∀ {n} ρ → ⟦ 0N {n} ⟧N ρ ≈ 0# 0N-homo [] = 0-homo 0N-homo (x ∷ ρ) = refl -- If c is equal to 0N, then c is semantically equal to 0#. 0≈⟦0⟧ : ∀ {n} {c : Normal n} → c ≈N 0N → ∀ ρ → 0# ≈ ⟦ c ⟧N ρ 0≈⟦0⟧ {c = c} c≈0 ρ = sym (begin ⟦ c ⟧N ρ ≈⟨ ⟦ c≈0 ⟧N-cong ρ ⟩ ⟦ 0N ⟧N ρ ≈⟨ 0N-homo ρ ⟩ 0# ∎) 1N-homo : ∀ {n} ρ → ⟦ 1N {n} ⟧N ρ ≈ 1# 1N-homo [] = 1-homo 1N-homo (x ∷ ρ) = begin 0# * x + ⟦ 1N ⟧N ρ ≈⟨ refl ⟨ +-cong ⟩ 1N-homo ρ ⟩ 0# * x + 1# ≈⟨ lemma₆ _ _ ⟩ 1# ∎ -- _*x+HN_ is equal to _*x+_. *x+HN≈*x+ : ∀ {n} (p : HNF (suc n)) (c : Normal n) → ∀ ρ → ⟦ p *x+HN c ⟧H ρ ≈ ⟦ p *x+ c ⟧H ρ *x+HN≈*x+ (p *x+ c′) c ρ = refl *x+HN≈*x+ ∅ c (x ∷ ρ) with c ≟N 0N ... | yes c≈0 = begin 0# ≈⟨ 0≈⟦0⟧ c≈0 ρ ⟩ ⟦ c ⟧N ρ ≈⟨ sym $ lemma₆ _ _ ⟩ 0# * x + ⟦ c ⟧N ρ ∎ ... | no c≉0 = refl ∅*x+HN-homo : ∀ {n} (c : Normal n) x ρ → ⟦ ∅ *x+HN c ⟧H (x ∷ ρ) ≈ ⟦ c ⟧N ρ ∅*x+HN-homo c x ρ with c ≟N 0N ... | yes c≈0 = 0≈⟦0⟧ c≈0 ρ ... | no c≉0 = lemma₆ _ _ mutual +H-homo : ∀ {n} (p₁ p₂ : HNF (suc n)) → ∀ ρ → ⟦ p₁ +H p₂ ⟧H ρ ≈ ⟦ p₁ ⟧H ρ + ⟦ p₂ ⟧H ρ +H-homo ∅ p₂ ρ = sym (proj₁ +-identity _) +H-homo (p₁ *x+ x₁) ∅ ρ = sym (proj₂ +-identity _) +H-homo (p₁ *x+ c₁) (p₂ *x+ c₂) (x ∷ ρ) = begin ⟦ (p₁ +H p₂) *x+HN (c₁ +N c₂) ⟧H (x ∷ ρ) ≈⟨ *x+HN≈*x+ (p₁ +H p₂) (c₁ +N c₂) (x ∷ ρ) ⟩ ⟦ p₁ +H p₂ ⟧H (x ∷ ρ) * x + ⟦ c₁ +N c₂ ⟧N ρ ≈⟨ (+H-homo p₁ p₂ (x ∷ ρ) ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ +N-homo c₁ c₂ ρ ⟩ (⟦ p₁ ⟧H (x ∷ ρ) + ⟦ p₂ ⟧H (x ∷ ρ)) * x + (⟦ c₁ ⟧N ρ + ⟦ c₂ ⟧N ρ) ≈⟨ lemma₁ _ _ _ _ _ ⟩ (⟦ p₁ ⟧H (x ∷ ρ) * x + ⟦ c₁ ⟧N ρ) + (⟦ p₂ ⟧H (x ∷ ρ) * x + ⟦ c₂ ⟧N ρ) ∎ +N-homo : ∀ {n} (p₁ p₂ : Normal n) → ∀ ρ → ⟦ p₁ +N p₂ ⟧N ρ ≈ ⟦ p₁ ⟧N ρ + ⟦ p₂ ⟧N ρ +N-homo (con c₁) (con c₂) _ = +-homo _ _ +N-homo (poly p₁) (poly p₂) ρ = +H-homo p₁ p₂ ρ *x+H-homo : ∀ {n} (p₁ p₂ : HNF (suc n)) x ρ → ⟦ p₁ *x+H p₂ ⟧H (x ∷ ρ) ≈ ⟦ p₁ ⟧H (x ∷ ρ) * x + ⟦ p₂ ⟧H (x ∷ ρ) *x+H-homo ∅ ∅ _ _ = sym $ lemma₆ _ _ *x+H-homo (p *x+ c) ∅ x ρ = begin ⟦ p *x+ c ⟧H (x ∷ ρ) * x + ⟦ 0N ⟧N ρ ≈⟨ refl ⟨ +-cong ⟩ 0N-homo ρ ⟩ ⟦ p *x+ c ⟧H (x ∷ ρ) * x + 0# ∎ *x+H-homo p₁ (p₂ *x+ c₂) x ρ = begin ⟦ (p₁ +H p₂) *x+HN c₂ ⟧H (x ∷ ρ) ≈⟨ *x+HN≈*x+ (p₁ +H p₂) c₂ (x ∷ ρ) ⟩ ⟦ p₁ +H p₂ ⟧H (x ∷ ρ) * x + ⟦ c₂ ⟧N ρ ≈⟨ (+H-homo p₁ p₂ (x ∷ ρ) ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ refl ⟩ (⟦ p₁ ⟧H (x ∷ ρ) + ⟦ p₂ ⟧H (x ∷ ρ)) * x + ⟦ c₂ ⟧N ρ ≈⟨ lemma₀ _ _ _ _ ⟩ ⟦ p₁ ⟧H (x ∷ ρ) * x + (⟦ p₂ ⟧H (x ∷ ρ) * x + ⟦ c₂ ⟧N ρ) ∎ mutual *NH-homo : ∀ {n} (c : Normal n) (p : HNF (suc n)) x ρ → ⟦ c *NH p ⟧H (x ∷ ρ) ≈ ⟦ c ⟧N ρ * ⟦ p ⟧H (x ∷ ρ) *NH-homo c ∅ x ρ = sym (proj₂ zero* _) *NH-homo c (p *x+ c′) x ρ with c ≟N 0N ... | yes c≈0 = begin 0# ≈⟨ sym (proj₁ zero* _) ⟩ 0# * (⟦ p ⟧H (x ∷ ρ) * x + ⟦ c′ ⟧N ρ) ≈⟨ 0≈⟦0⟧ c≈0 ρ ⟨ *-cong ⟩ refl ⟩ ⟦ c ⟧N ρ * (⟦ p ⟧H (x ∷ ρ) * x + ⟦ c′ ⟧N ρ) ∎ ... | no c≉0 = begin ⟦ c *NH p ⟧H (x ∷ ρ) * x + ⟦ c *N c′ ⟧N ρ ≈⟨ (*NH-homo c p x ρ ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ *N-homo c c′ ρ ⟩ (⟦ c ⟧N ρ * ⟦ p ⟧H (x ∷ ρ)) * x + (⟦ c ⟧N ρ * ⟦ c′ ⟧N ρ) ≈⟨ lemma₃ _ _ _ _ ⟩ ⟦ c ⟧N ρ * (⟦ p ⟧H (x ∷ ρ) * x + ⟦ c′ ⟧N ρ) ∎ *HN-homo : ∀ {n} (p : HNF (suc n)) (c : Normal n) x ρ → ⟦ p *HN c ⟧H (x ∷ ρ) ≈ ⟦ p ⟧H (x ∷ ρ) * ⟦ c ⟧N ρ *HN-homo ∅ c x ρ = sym (proj₁ zero* _) *HN-homo (p *x+ c′) c x ρ with c ≟N 0N ... | yes c≈0 = begin 0# ≈⟨ sym (proj₂ zero* _) ⟩ (⟦ p ⟧H (x ∷ ρ) * x + ⟦ c′ ⟧N ρ) * 0# ≈⟨ refl ⟨ *-cong ⟩ 0≈⟦0⟧ c≈0 ρ ⟩ (⟦ p ⟧H (x ∷ ρ) * x + ⟦ c′ ⟧N ρ) * ⟦ c ⟧N ρ ∎ ... | no c≉0 = begin ⟦ p *HN c ⟧H (x ∷ ρ) * x + ⟦ c′ *N c ⟧N ρ ≈⟨ (*HN-homo p c x ρ ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ *N-homo c′ c ρ ⟩ (⟦ p ⟧H (x ∷ ρ) * ⟦ c ⟧N ρ) * x + (⟦ c′ ⟧N ρ * ⟦ c ⟧N ρ) ≈⟨ lemma₂ _ _ _ _ ⟩ (⟦ p ⟧H (x ∷ ρ) * x + ⟦ c′ ⟧N ρ) * ⟦ c ⟧N ρ ∎ *H-homo : ∀ {n} (p₁ p₂ : HNF (suc n)) → ∀ ρ → ⟦ p₁ *H p₂ ⟧H ρ ≈ ⟦ p₁ ⟧H ρ * ⟦ p₂ ⟧H ρ *H-homo ∅ p₂ ρ = sym $ proj₁ zero* _ *H-homo (p₁ *x+ c₁) ∅ ρ = sym $ proj₂ zero* _ *H-homo (p₁ *x+ c₁) (p₂ *x+ c₂) (x ∷ ρ) = begin ⟦ ((p₁ *H p₂) *x+H ((p₁ *HN c₂) +H (c₁ *NH p₂))) *x+HN (c₁ *N c₂) ⟧H (x ∷ ρ) ≈⟨ *x+HN≈*x+ ((p₁ *H p₂) *x+H ((p₁ *HN c₂) +H (c₁ *NH p₂))) (c₁ *N c₂) (x ∷ ρ) ⟩ ⟦ (p₁ *H p₂) *x+H ((p₁ *HN c₂) +H (c₁ *NH p₂)) ⟧H (x ∷ ρ) * x + ⟦ c₁ *N c₂ ⟧N ρ ≈⟨ (*x+H-homo (p₁ *H p₂) ((p₁ *HN c₂) +H (c₁ *NH p₂)) x ρ ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ *N-homo c₁ c₂ ρ ⟩ (⟦ p₁ *H p₂ ⟧H (x ∷ ρ) * x + ⟦ (p₁ *HN c₂) +H (c₁ *NH p₂) ⟧H (x ∷ ρ)) * x + ⟦ c₁ ⟧N ρ * ⟦ c₂ ⟧N ρ ≈⟨ (((*H-homo p₁ p₂ (x ∷ ρ) ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ (+H-homo (p₁ *HN c₂) (c₁ *NH p₂) (x ∷ ρ))) ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ refl ⟩ (⟦ p₁ ⟧H (x ∷ ρ) * ⟦ p₂ ⟧H (x ∷ ρ) * x + (⟦ p₁ *HN c₂ ⟧H (x ∷ ρ) + ⟦ c₁ *NH p₂ ⟧H (x ∷ ρ))) * x + ⟦ c₁ ⟧N ρ * ⟦ c₂ ⟧N ρ ≈⟨ ((refl ⟨ +-cong ⟩ (*HN-homo p₁ c₂ x ρ ⟨ +-cong ⟩ *NH-homo c₁ p₂ x ρ)) ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ refl ⟩ (⟦ p₁ ⟧H (x ∷ ρ) * ⟦ p₂ ⟧H (x ∷ ρ) * x + (⟦ p₁ ⟧H (x ∷ ρ) * ⟦ c₂ ⟧N ρ + ⟦ c₁ ⟧N ρ * ⟦ p₂ ⟧H (x ∷ ρ))) * x + (⟦ c₁ ⟧N ρ * ⟦ c₂ ⟧N ρ) ≈⟨ lemma₄ _ _ _ _ _ ⟩ (⟦ p₁ ⟧H (x ∷ ρ) * x + ⟦ c₁ ⟧N ρ) * (⟦ p₂ ⟧H (x ∷ ρ) * x + ⟦ c₂ ⟧N ρ) ∎ *N-homo : ∀ {n} (p₁ p₂ : Normal n) → ∀ ρ → ⟦ p₁ *N p₂ ⟧N ρ ≈ ⟦ p₁ ⟧N ρ * ⟦ p₂ ⟧N ρ *N-homo (con c₁) (con c₂) _ = *-homo _ _ *N-homo (poly p₁) (poly p₂) ρ = *H-homo p₁ p₂ ρ ^N-homo : ∀ {n} (p : Normal n) (k : ℕ) → ∀ ρ → ⟦ p ^N k ⟧N ρ ≈ ⟦ p ⟧N ρ ^ k ^N-homo p zero ρ = 1N-homo ρ ^N-homo p (suc k) ρ = begin ⟦ p *N (p ^N k) ⟧N ρ ≈⟨ *N-homo p (p ^N k) ρ ⟩ ⟦ p ⟧N ρ * ⟦ p ^N k ⟧N ρ ≈⟨ refl ⟨ *-cong ⟩ ^N-homo p k ρ ⟩ ⟦ p ⟧N ρ * (⟦ p ⟧N ρ ^ k) ∎ mutual -H‿-homo : ∀ {n} (p : HNF (suc n)) → ∀ ρ → ⟦ -H p ⟧H ρ ≈ - ⟦ p ⟧H ρ -H‿-homo p (x ∷ ρ) = begin ⟦ (-N 1N) *NH p ⟧H (x ∷ ρ) ≈⟨ *NH-homo (-N 1N) p x ρ ⟩ ⟦ -N 1N ⟧N ρ * ⟦ p ⟧H (x ∷ ρ) ≈⟨ trans (-N‿-homo 1N ρ) (-‿cong (1N-homo ρ)) ⟨ *-cong ⟩ refl ⟩ - 1# * ⟦ p ⟧H (x ∷ ρ) ≈⟨ lemma₇ _ ⟩ - ⟦ p ⟧H (x ∷ ρ) ∎ -N‿-homo : ∀ {n} (p : Normal n) → ∀ ρ → ⟦ -N p ⟧N ρ ≈ - ⟦ p ⟧N ρ -N‿-homo (con c) _ = -‿homo _ -N‿-homo (poly p) ρ = -H‿-homo p ρ ------------------------------------------------------------------------ -- Correctness correct-con : ∀ {n} (c : C.Carrier) (ρ : Vec Carrier n) → ⟦ normalise-con c ⟧N ρ ≈ ⟦ c ⟧′ correct-con c [] = refl correct-con c (x ∷ ρ) = begin ⟦ ∅ *x+HN normalise-con c ⟧H (x ∷ ρ) ≈⟨ ∅*x+HN-homo (normalise-con c) x ρ ⟩ ⟦ normalise-con c ⟧N ρ ≈⟨ correct-con c ρ ⟩ ⟦ c ⟧′ ∎ correct-var : ∀ {n} (i : Fin n) → ∀ ρ → ⟦ normalise-var i ⟧N ρ ≈ lookup i ρ correct-var () [] correct-var (suc i) (x ∷ ρ) = begin ⟦ ∅ *x+HN normalise-var i ⟧H (x ∷ ρ) ≈⟨ ∅*x+HN-homo (normalise-var i) x ρ ⟩ ⟦ normalise-var i ⟧N ρ ≈⟨ correct-var i ρ ⟩ lookup i ρ ∎ correct-var zero (x ∷ ρ) = begin (0# * x + ⟦ 1N ⟧N ρ) * x + ⟦ 0N ⟧N ρ ≈⟨ ((refl ⟨ +-cong ⟩ 1N-homo ρ) ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ 0N-homo ρ ⟩ (0# * x + 1#) * x + 0# ≈⟨ lemma₅ _ ⟩ x ∎ correct : ∀ {n} (p : Polynomial n) → ∀ ρ → ⟦ p ⟧↓ ρ ≈ ⟦ p ⟧ ρ correct (op [+] p₁ p₂) ρ = begin ⟦ normalise p₁ +N normalise p₂ ⟧N ρ ≈⟨ +N-homo (normalise p₁) (normalise p₂) ρ ⟩ ⟦ p₁ ⟧↓ ρ + ⟦ p₂ ⟧↓ ρ ≈⟨ correct p₁ ρ ⟨ +-cong ⟩ correct p₂ ρ ⟩ ⟦ p₁ ⟧ ρ + ⟦ p₂ ⟧ ρ ∎ correct (op [*] p₁ p₂) ρ = begin ⟦ normalise p₁ *N normalise p₂ ⟧N ρ ≈⟨ *N-homo (normalise p₁) (normalise p₂) ρ ⟩ ⟦ p₁ ⟧↓ ρ * ⟦ p₂ ⟧↓ ρ ≈⟨ correct p₁ ρ ⟨ *-cong ⟩ correct p₂ ρ ⟩ ⟦ p₁ ⟧ ρ * ⟦ p₂ ⟧ ρ ∎ correct (con c) ρ = correct-con c ρ correct (var i) ρ = correct-var i ρ correct (p :^ k) ρ = begin ⟦ normalise p ^N k ⟧N ρ ≈⟨ ^N-homo (normalise p) k ρ ⟩ ⟦ p ⟧↓ ρ ^ k ≈⟨ correct p ρ ⟨ ^-cong ⟩ PropEq.refl {x = k} ⟩ ⟦ p ⟧ ρ ^ k ∎ correct (:- p) ρ = begin ⟦ -N normalise p ⟧N ρ ≈⟨ -N‿-homo (normalise p) ρ ⟩ - ⟦ p ⟧↓ ρ ≈⟨ -‿cong (correct p ρ) ⟩ - ⟦ p ⟧ ρ ∎ ------------------------------------------------------------------------ -- "Tactics" open Reflection setoid var ⟦_⟧ ⟦_⟧↓ correct public using (prove; solve) renaming (_⊜_ to _:=_) -- For examples of how solve and _:=_ can be used to -- semi-automatically prove ring equalities, see, for instance, -- Data.Digit or Data.Nat.DivMod. lib-0.7/src/Algebra/Props/0000755000000000000000000000000012101774706013550 5ustar0000000000000000lib-0.7/src/Algebra/Props/Group.agda0000644000000000000000000000533312101774706015466 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Some derivable properties ------------------------------------------------------------------------ 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 ⁻¹ ∎ lib-0.7/src/Algebra/Props/BooleanAlgebra.agda0000644000000000000000000006117412101774706017234 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Some derivable properties ------------------------------------------------------------------------ open import Algebra 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) lib-0.7/src/Algebra/Props/Lattice.agda0000644000000000000000000000761212101774706015761 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Some derivable properties ------------------------------------------------------------------------ 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}) lib-0.7/src/Algebra/Props/DistributiveLattice.agda0000644000000000000000000000660012101774706020353 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Some derivable properties ------------------------------------------------------------------------ open import Algebra 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}) lib-0.7/src/Algebra/Props/Ring.agda0000644000000000000000000000442712101774706015274 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Some derivable properties ------------------------------------------------------------------------ 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) ∎ lib-0.7/src/Algebra/Props/AbelianGroup.agda0000644000000000000000000000324612101774706016743 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Some derivable properties ------------------------------------------------------------------------ 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 ⁻¹ ∎ lib-0.7/src/Algebra/Props/BooleanAlgebra/0000755000000000000000000000000012101774706016405 5ustar0000000000000000lib-0.7/src/Algebra/Props/BooleanAlgebra/Expression.agda0000644000000000000000000002064312101774706021367 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Boolean algebra expressions ------------------------------------------------------------------------ open import Algebra module Algebra.Props.BooleanAlgebra.Expression {b} (B : BooleanAlgebra b b) where open BooleanAlgebra B open import Category.Applicative import Category.Applicative.Indexed as Applicative open import Category.Monad open import Category.Monad.Identity open import Data.Fin using (Fin) open import Data.Nat open import Data.Vec as Vec using (Vec) open import Data.Product import Data.Vec.Properties as VecProp open import Function open import Relation.Binary.PropositionalEquality as P using (_≗_) import Relation.Binary.Reflection as Reflection open import Relation.Binary.Vec.Pointwise as PW using (Pointwise; module Pointwise; ext) -- Expressions made up of variables and the operations of a boolean -- algebra. infixr 7 _and_ infixr 6 _or_ data Expr n : Set b where var : (x : Fin n) → Expr n _or_ _and_ : (e₁ e₂ : Expr n) → Expr n not : (e : Expr n) → Expr n top bot : Expr n -- The semantics of an expression, parametrised by an applicative -- functor. module Semantics {F : Set b → Set b} (A : RawApplicative F) where open RawApplicative A ⟦_⟧ : ∀ {n} → Expr n → Vec (F Carrier) n → F Carrier ⟦ var x ⟧ ρ = Vec.lookup x ρ ⟦ e₁ or e₂ ⟧ ρ = pure _∨_ ⊛ ⟦ e₁ ⟧ ρ ⊛ ⟦ e₂ ⟧ ρ ⟦ e₁ and e₂ ⟧ ρ = pure _∧_ ⊛ ⟦ e₁ ⟧ ρ ⊛ ⟦ e₂ ⟧ ρ ⟦ not e ⟧ ρ = pure ¬_ ⊛ ⟦ e ⟧ ρ ⟦ top ⟧ ρ = pure ⊤ ⟦ bot ⟧ ρ = pure ⊥ -- flip Semantics.⟦_⟧ e is natural. module Naturality {F₁ F₂ : Set b → Set b} {A₁ : RawApplicative F₁} {A₂ : RawApplicative F₂} (f : Applicative.Morphism A₁ A₂) where open P.≡-Reasoning open Applicative.Morphism f open Semantics A₁ renaming (⟦_⟧ to ⟦_⟧₁) open Semantics A₂ renaming (⟦_⟧ to ⟦_⟧₂) open RawApplicative A₁ renaming (pure to pure₁; _⊛_ to _⊛₁_) open RawApplicative A₂ renaming (pure to pure₂; _⊛_ to _⊛₂_) natural : ∀ {n} (e : Expr n) → op ∘ ⟦ e ⟧₁ ≗ ⟦ e ⟧₂ ∘ Vec.map op natural (var x) ρ = begin op (Vec.lookup x ρ) ≡⟨ P.sym $ Applicative.Morphism.op-<$> (VecProp.lookup-morphism x) op ρ ⟩ Vec.lookup x (Vec.map op ρ) ∎ natural (e₁ or e₂) ρ = begin op (pure₁ _∨_ ⊛₁ ⟦ e₁ ⟧₁ ρ ⊛₁ ⟦ e₂ ⟧₁ ρ) ≡⟨ op-⊛ _ _ ⟩ op (pure₁ _∨_ ⊛₁ ⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (op-⊛ _ _) P.refl ⟩ op (pure₁ _∨_) ⊛₂ op (⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (P.cong₂ _⊛₂_ (op-pure _) (natural e₁ ρ)) (natural e₂ ρ) ⟩ pure₂ _∨_ ⊛₂ ⟦ e₁ ⟧₂ (Vec.map op ρ) ⊛₂ ⟦ e₂ ⟧₂ (Vec.map op ρ) ∎ natural (e₁ and e₂) ρ = begin op (pure₁ _∧_ ⊛₁ ⟦ e₁ ⟧₁ ρ ⊛₁ ⟦ e₂ ⟧₁ ρ) ≡⟨ op-⊛ _ _ ⟩ op (pure₁ _∧_ ⊛₁ ⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (op-⊛ _ _) P.refl ⟩ op (pure₁ _∧_) ⊛₂ op (⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (P.cong₂ _⊛₂_ (op-pure _) (natural e₁ ρ)) (natural e₂ ρ) ⟩ pure₂ _∧_ ⊛₂ ⟦ e₁ ⟧₂ (Vec.map op ρ) ⊛₂ ⟦ e₂ ⟧₂ (Vec.map op ρ) ∎ natural (not e) ρ = begin op (pure₁ ¬_ ⊛₁ ⟦ e ⟧₁ ρ) ≡⟨ op-⊛ _ _ ⟩ op (pure₁ ¬_) ⊛₂ op (⟦ e ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (op-pure _) (natural e ρ) ⟩ pure₂ ¬_ ⊛₂ ⟦ e ⟧₂ (Vec.map op ρ) ∎ natural top ρ = begin op (pure₁ ⊤) ≡⟨ op-pure _ ⟩ pure₂ ⊤ ∎ natural bot ρ = begin op (pure₁ ⊥) ≡⟨ op-pure _ ⟩ pure₂ ⊥ ∎ -- An example of how naturality can be used: Any boolean algebra can -- be lifted, in a pointwise manner, to vectors of carrier elements. lift : ℕ → BooleanAlgebra b b lift n = record { Carrier = Vec Carrier n ; _≈_ = Pointwise _≈_ ; _∨_ = Vec.zipWith _∨_ ; _∧_ = Vec.zipWith _∧_ ; ¬_ = Vec.map ¬_ ; ⊤ = Vec.replicate ⊤ ; ⊥ = Vec.replicate ⊥ ; isBooleanAlgebra = record { isDistributiveLattice = record { isLattice = record { isEquivalence = PW.isEquivalence isEquivalence ; ∨-comm = λ _ _ → ext λ i → solve i 2 (λ x y → x or y , y or x) (∨-comm _ _) _ _ ; ∨-assoc = λ _ _ _ → ext λ i → solve i 3 (λ x y z → (x or y) or z , x or (y or z)) (∨-assoc _ _ _) _ _ _ ; ∨-cong = λ xs≈us ys≈vs → ext λ i → solve₁ i 4 (λ x y u v → x or y , u or v) _ _ _ _ (∨-cong (Pointwise.app xs≈us i) (Pointwise.app ys≈vs i)) ; ∧-comm = λ _ _ → ext λ i → solve i 2 (λ x y → x and y , y and x) (∧-comm _ _) _ _ ; ∧-assoc = λ _ _ _ → ext λ i → solve i 3 (λ x y z → (x and y) and z , x and (y and z)) (∧-assoc _ _ _) _ _ _ ; ∧-cong = λ xs≈ys us≈vs → ext λ i → solve₁ i 4 (λ x y u v → x and y , u and v) _ _ _ _ (∧-cong (Pointwise.app xs≈ys i) (Pointwise.app us≈vs i)) ; absorptive = (λ _ _ → ext λ i → solve i 2 (λ x y → x or (x and y) , x) (proj₁ absorptive _ _) _ _) , (λ _ _ → ext λ i → solve i 2 (λ x y → x and (x or y) , x) (proj₂ absorptive _ _) _ _) } ; ∨-∧-distribʳ = λ _ _ _ → ext λ i → solve i 3 (λ x y z → (y and z) or x , (y or x) and (z or x)) (∨-∧-distribʳ _ _ _) _ _ _ } ; ∨-complementʳ = λ _ → ext λ i → solve i 1 (λ x → x or (not x) , top) (∨-complementʳ _) _ ; ∧-complementʳ = λ _ → ext λ i → solve i 1 (λ x → x and (not x) , bot) (∧-complementʳ _) _ ; ¬-cong = λ xs≈ys → ext λ i → solve₁ i 2 (λ x y → not x , not y) _ _ (¬-cong (Pointwise.app xs≈ys i)) } } where ⟦_⟧Id : ∀ {n} → Expr n → Vec Carrier n → Carrier ⟦_⟧Id = Semantics.⟦_⟧ (RawMonad.rawIApplicative IdentityMonad) ⟦_⟧Vec : ∀ {m n} → Expr n → Vec (Vec Carrier m) n → Vec Carrier m ⟦_⟧Vec = Semantics.⟦_⟧ Vec.applicative open module R {n} (i : Fin n) = Reflection setoid var (λ e ρ → Vec.lookup i (⟦ e ⟧Vec ρ)) (λ e ρ → ⟦ e ⟧Id (Vec.map (Vec.lookup i) ρ)) (λ e ρ → sym $ reflexive $ Naturality.natural (VecProp.lookup-morphism i) e ρ) lib-0.7/src/Algebra/FunctionProperties/0000755000000000000000000000000012101774706016307 5ustar0000000000000000lib-0.7/src/Algebra/FunctionProperties/Core.agda0000644000000000000000000000143612101774706020021 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Properties of functions, such as associativity and commutativity ------------------------------------------------------------------------ -- This file contains some core definitions which are reexported by -- Algebra.FunctionProperties. They are placed here because -- Algebra.FunctionProperties is a parameterised module, and some of -- the parameters are irrelevant for these definitions. module Algebra.FunctionProperties.Core where open import Level ------------------------------------------------------------------------ -- Unary and binary operations Op₁ : ∀ {ℓ} → Set ℓ → Set ℓ Op₁ A = A → A Op₂ : ∀ {ℓ} → Set ℓ → Set ℓ Op₂ A = A → A → A lib-0.7/src/Algebra/RingSolver/0000755000000000000000000000000012101774706014537 5ustar0000000000000000lib-0.7/src/Algebra/RingSolver/Natural-coefficients.agda0000644000000000000000000000447312101774706021432 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Instantiates the ring solver, using the natural numbers as the -- coefficient "ring" ------------------------------------------------------------------------ open import Algebra import Algebra.Operations open import Relation.Nullary module Algebra.RingSolver.Natural-coefficients {r₁ r₂} (R : CommutativeSemiring r₁ r₂) (dec : let open CommutativeSemiring R open Algebra.Operations semiring in ∀ m n → Dec (m × 1# ≈ n × 1#)) where import Algebra.RingSolver open import Algebra.RingSolver.AlmostCommutativeRing open import Data.Nat as ℕ open import Data.Product using (module Σ) open import Function import Relation.Binary.EqReasoning import Relation.Nullary.Decidable as Dec open CommutativeSemiring R open Algebra.Operations semiring open Relation.Binary.EqReasoning setoid private -- The coefficient "ring". ℕ-ring : RawRing _ ℕ-ring = record { Carrier = ℕ ; _+_ = ℕ._+_ ; _*_ = ℕ._*_ ; -_ = id ; 0# = 0 ; 1# = 1 } -- There is a homomorphism from ℕ to R. -- -- Note that _×′_ is used rather than _×_. If _×_ were used, then -- Function.Related.TypeIsomorphisms.test would fail to type-check. homomorphism : ℕ-ring -Raw-AlmostCommutative⟶ fromCommutativeSemiring R homomorphism = record { ⟦_⟧ = λ n → n ×′ 1# ; +-homo = ×′-homo-+ 1# ; *-homo = ×′1-homo-* ; -‿homo = λ _ → refl ; 0-homo = refl ; 1-homo = refl } -- Equality of certain expressions can be decided. dec′ : ∀ m n → Dec (m ×′ 1# ≈ n ×′ 1#) dec′ m n = Dec.map′ to from (dec m n) where to : m × 1# ≈ n × 1# → m ×′ 1# ≈ n ×′ 1# to m≈n = begin m ×′ 1# ≈⟨ sym $ ×≈×′ m 1# ⟩ m × 1# ≈⟨ m≈n ⟩ n × 1# ≈⟨ ×≈×′ n 1# ⟩ n ×′ 1# ∎ from : m ×′ 1# ≈ n ×′ 1# → m × 1# ≈ n × 1# from m≈n = begin m × 1# ≈⟨ ×≈×′ m 1# ⟩ m ×′ 1# ≈⟨ m≈n ⟩ n ×′ 1# ≈⟨ sym $ ×≈×′ n 1# ⟩ n × 1# ∎ -- The instantiation. open Algebra.RingSolver _ _ homomorphism dec′ public lib-0.7/src/Algebra/RingSolver/Lemmas.agda0000644000000000000000000001136412101774706016600 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Some boring lemmas used by the ring solver ------------------------------------------------------------------------ -- Note that these proofs use all "almost commutative ring" properties. open import Algebra open import Algebra.RingSolver.AlmostCommutativeRing module Algebra.RingSolver.Lemmas {r₁ r₂ r₃} (coeff : RawRing r₁) (r : AlmostCommutativeRing r₂ r₃) (morphism : coeff -Raw-AlmostCommutative⟶ r) where private module C = RawRing coeff open AlmostCommutativeRing r open import Algebra.Morphism open _-Raw-AlmostCommutative⟶_ morphism import Relation.Binary.EqReasoning as EqR; open EqR setoid open import Function open import Data.Product lemma₀ : ∀ a b c x → (a + b) * x + c ≈ a * x + (b * x + c) lemma₀ a b c x = begin (a + b) * x + c ≈⟨ proj₂ distrib _ _ _ ⟨ +-cong ⟩ refl ⟩ (a * x + b * x) + c ≈⟨ +-assoc _ _ _ ⟩ a * x + (b * x + c) ∎ lemma₁ : ∀ a b c d x → (a + b) * x + (c + d) ≈ (a * x + c) + (b * x + d) lemma₁ a b c d x = begin (a + b) * x + (c + d) ≈⟨ lemma₀ _ _ _ _ ⟩ a * x + (b * x + (c + d)) ≈⟨ refl ⟨ +-cong ⟩ sym (+-assoc _ _ _) ⟩ a * x + ((b * x + c) + d) ≈⟨ refl ⟨ +-cong ⟩ (+-comm _ _ ⟨ +-cong ⟩ refl) ⟩ a * x + ((c + b * x) + d) ≈⟨ refl ⟨ +-cong ⟩ +-assoc _ _ _ ⟩ a * x + (c + (b * x + d)) ≈⟨ sym $ +-assoc _ _ _ ⟩ (a * x + c) + (b * x + d) ∎ lemma₂ : ∀ a b c x → a * c * x + b * c ≈ (a * x + b) * c lemma₂ a b c x = begin a * c * x + b * c ≈⟨ lem ⟨ +-cong ⟩ refl ⟩ a * x * c + b * c ≈⟨ sym $ proj₂ distrib _ _ _ ⟩ (a * x + b) * c ∎ where lem = begin a * c * x ≈⟨ *-assoc _ _ _ ⟩ a * (c * x) ≈⟨ refl ⟨ *-cong ⟩ *-comm _ _ ⟩ a * (x * c) ≈⟨ sym $ *-assoc _ _ _ ⟩ a * x * c ∎ lemma₃ : ∀ a b c x → a * b * x + a * c ≈ a * (b * x + c) lemma₃ a b c x = begin a * b * x + a * c ≈⟨ *-assoc _ _ _ ⟨ +-cong ⟩ refl ⟩ a * (b * x) + a * c ≈⟨ sym $ proj₁ distrib _ _ _ ⟩ a * (b * x + c) ∎ lemma₄ : ∀ a b c d x → (a * c * x + (a * d + b * c)) * x + b * d ≈ (a * x + b) * (c * x + d) lemma₄ a b c d x = begin (a * c * x + (a * d + b * c)) * x + b * d ≈⟨ proj₂ distrib _ _ _ ⟨ +-cong ⟩ refl ⟩ (a * c * x * x + (a * d + b * c) * x) + b * d ≈⟨ refl ⟨ +-cong ⟩ ((refl ⟨ +-cong ⟩ refl) ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ refl ⟩ (a * c * x * x + (a * d + b * c) * x) + b * d ≈⟨ +-assoc _ _ _ ⟩ a * c * x * x + ((a * d + b * c) * x + b * d) ≈⟨ lem₁ ⟨ +-cong ⟩ (lem₂ ⟨ +-cong ⟩ refl) ⟩ a * x * (c * x) + (a * x * d + b * (c * x) + b * d) ≈⟨ refl ⟨ +-cong ⟩ +-assoc _ _ _ ⟩ a * x * (c * x) + (a * x * d + (b * (c * x) + b * d)) ≈⟨ sym $ +-assoc _ _ _ ⟩ a * x * (c * x) + a * x * d + (b * (c * x) + b * d) ≈⟨ sym $ proj₁ distrib _ _ _ ⟨ +-cong ⟩ proj₁ distrib _ _ _ ⟩ a * x * (c * x + d) + b * (c * x + d) ≈⟨ sym $ proj₂ distrib _ _ _ ⟩ (a * x + b) * (c * x + d) ∎ where lem₁′ = begin a * c * x ≈⟨ *-assoc _ _ _ ⟩ a * (c * x) ≈⟨ refl ⟨ *-cong ⟩ *-comm _ _ ⟩ a * (x * c) ≈⟨ sym $ *-assoc _ _ _ ⟩ a * x * c ∎ lem₁ = begin a * c * x * x ≈⟨ lem₁′ ⟨ *-cong ⟩ refl ⟩ a * x * c * x ≈⟨ *-assoc _ _ _ ⟩ a * x * (c * x) ∎ lem₂ = begin (a * d + b * c) * x ≈⟨ proj₂ distrib _ _ _ ⟩ a * d * x + b * c * x ≈⟨ *-assoc _ _ _ ⟨ +-cong ⟩ *-assoc _ _ _ ⟩ a * (d * x) + b * (c * x) ≈⟨ (refl ⟨ *-cong ⟩ *-comm _ _) ⟨ +-cong ⟩ refl ⟩ a * (x * d) + b * (c * x) ≈⟨ sym $ *-assoc _ _ _ ⟨ +-cong ⟩ refl ⟩ a * x * d + b * (c * x) ∎ lemma₅ : ∀ x → (0# * x + 1#) * x + 0# ≈ x lemma₅ x = begin (0# * x + 1#) * x + 0# ≈⟨ ((zeroˡ _ ⟨ +-cong ⟩ refl) ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ refl ⟩ (0# + 1#) * x + 0# ≈⟨ (proj₁ +-identity _ ⟨ *-cong ⟩ refl) ⟨ +-cong ⟩ refl ⟩ 1# * x + 0# ≈⟨ proj₂ +-identity _ ⟩ 1# * x ≈⟨ proj₁ *-identity _ ⟩ x ∎ lemma₆ : ∀ a x → 0# * x + a ≈ a lemma₆ a x = begin 0# * x + a ≈⟨ zeroˡ _ ⟨ +-cong ⟩ refl ⟩ 0# + a ≈⟨ proj₁ +-identity _ ⟩ a ∎ lemma₇ : ∀ x → - 1# * x ≈ - x lemma₇ x = begin - 1# * x ≈⟨ -‿*-distribˡ _ _ ⟩ - (1# * x) ≈⟨ -‿cong (proj₁ *-identity _) ⟩ - x ∎ lib-0.7/src/Algebra/RingSolver/Simple.agda0000644000000000000000000000120412101774706016603 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Instantiates the ring solver with two copies of the same ring with -- decidable equality ------------------------------------------------------------------------ open import Algebra.RingSolver.AlmostCommutativeRing open import Relation.Binary module Algebra.RingSolver.Simple {r₁ r₂} (R : AlmostCommutativeRing r₁ r₂) (_≟_ : Decidable (AlmostCommutativeRing._≈_ R)) where open AlmostCommutativeRing R import Algebra.RingSolver as RS open RS rawRing R (-raw-almostCommutative⟶ R) _≟_ public lib-0.7/src/Algebra/RingSolver/AlmostCommutativeRing.agda0000644000000000000000000001143612101774706021657 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Commutative semirings with some additional structure ("almost" -- commutative rings), used by the ring solver ------------------------------------------------------------------------ module Algebra.RingSolver.AlmostCommutativeRing where open import Relation.Binary open import Algebra open import Algebra.Structures open import Algebra.FunctionProperties import Algebra.Morphism as Morphism open import Function open import Level ------------------------------------------------------------------------ -- Definitions record IsAlmostCommutativeRing {a ℓ} {A : Set a} (_≈_ : Rel A ℓ) (_+_ _*_ : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where field isCommutativeSemiring : IsCommutativeSemiring _≈_ _+_ _*_ 0# 1# -‿cong : -_ Preserves _≈_ ⟶ _≈_ -‿*-distribˡ : ∀ x y → ((- x) * y) ≈ (- (x * y)) -‿+-comm : ∀ x y → ((- x) + (- y)) ≈ (- (x + y)) open IsCommutativeSemiring isCommutativeSemiring public record AlmostCommutativeRing c ℓ : Set (suc (c ⊔ ℓ)) where infix 8 -_ infixl 7 _*_ infixl 6 _+_ infix 4 _≈_ field Carrier : Set c _≈_ : Rel Carrier ℓ _+_ : Op₂ Carrier _*_ : Op₂ Carrier -_ : Op₁ Carrier 0# : Carrier 1# : Carrier isAlmostCommutativeRing : IsAlmostCommutativeRing _≈_ _+_ _*_ -_ 0# 1# open IsAlmostCommutativeRing isAlmostCommutativeRing public commutativeSemiring : CommutativeSemiring _ _ commutativeSemiring = record { isCommutativeSemiring = isCommutativeSemiring } open CommutativeSemiring commutativeSemiring public using ( setoid ; +-semigroup; +-monoid; +-commutativeMonoid ; *-semigroup; *-monoid; *-commutativeMonoid ; semiring ) rawRing : RawRing _ rawRing = record { _+_ = _+_ ; _*_ = _*_ ; -_ = -_ ; 0# = 0# ; 1# = 1# } ------------------------------------------------------------------------ -- Homomorphisms record _-Raw-AlmostCommutative⟶_ {r₁ r₂ r₃} (From : RawRing r₁) (To : AlmostCommutativeRing r₂ r₃) : Set (r₁ ⊔ r₂ ⊔ r₃) where private module F = RawRing From module T = AlmostCommutativeRing To open Morphism.Definitions F.Carrier T.Carrier T._≈_ field ⟦_⟧ : Morphism +-homo : Homomorphic₂ ⟦_⟧ F._+_ T._+_ *-homo : Homomorphic₂ ⟦_⟧ F._*_ T._*_ -‿homo : Homomorphic₁ ⟦_⟧ F.-_ T.-_ 0-homo : Homomorphic₀ ⟦_⟧ F.0# T.0# 1-homo : Homomorphic₀ ⟦_⟧ F.1# T.1# -raw-almostCommutative⟶ : ∀ {r₁ r₂} (R : AlmostCommutativeRing r₁ r₂) → AlmostCommutativeRing.rawRing R -Raw-AlmostCommutative⟶ R -raw-almostCommutative⟶ R = record { ⟦_⟧ = id ; +-homo = λ _ _ → refl ; *-homo = λ _ _ → refl ; -‿homo = λ _ → refl ; 0-homo = refl ; 1-homo = refl } where open AlmostCommutativeRing R -- A homomorphism induces a notion of equivalence on the raw ring. Induced-equivalence : ∀ {c₁ c₂ ℓ} {Coeff : RawRing c₁} {R : AlmostCommutativeRing c₂ ℓ} → Coeff -Raw-AlmostCommutative⟶ R → Rel (RawRing.Carrier Coeff) ℓ Induced-equivalence {R = R} morphism a b = ⟦ a ⟧ ≈ ⟦ b ⟧ where open AlmostCommutativeRing R open _-Raw-AlmostCommutative⟶_ morphism ------------------------------------------------------------------------ -- Conversions -- Commutative rings are almost commutative rings. fromCommutativeRing : ∀ {r₁ r₂} → CommutativeRing r₁ r₂ → AlmostCommutativeRing _ _ fromCommutativeRing CR = record { isAlmostCommutativeRing = record { isCommutativeSemiring = isCommutativeSemiring ; -‿cong = -‿cong ; -‿*-distribˡ = -‿*-distribˡ ; -‿+-comm = ⁻¹-∙-comm } } where open CommutativeRing CR import Algebra.Props.Ring as R; open R ring import Algebra.Props.AbelianGroup as AG; open AG +-abelianGroup -- Commutative semirings can be viewed as almost commutative rings by -- using identity as the "almost negation". fromCommutativeSemiring : ∀ {r₁ r₂} → CommutativeSemiring r₁ r₂ → AlmostCommutativeRing _ _ fromCommutativeSemiring CS = record { -_ = id ; isAlmostCommutativeRing = record { isCommutativeSemiring = isCommutativeSemiring ; -‿cong = id ; -‿*-distribˡ = λ _ _ → refl ; -‿+-comm = λ _ _ → refl } } where open CommutativeSemiring CS lib-0.7/src/Data/0000755000000000000000000000000012101774706011761 5ustar0000000000000000lib-0.7/src/Data/Sum.agda0000644000000000000000000000313312101774706013343 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Sums (disjoint unions) ------------------------------------------------------------------------ module Data.Sum where open import Function open import Data.Maybe.Core open import Level ------------------------------------------------------------------------ -- Definition infixr 1 _⊎_ data _⊎_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where inj₁ : (x : A) → A ⊎ B inj₂ : (y : B) → A ⊎ B {-# IMPORT Data.FFI #-} {-# COMPILED_DATA _⊎_ Data.FFI.AgdaEither Left Right #-} ------------------------------------------------------------------------ -- Functions [_,_] : ∀ {a b c} {A : Set a} {B : Set b} {C : A ⊎ B → Set c} → ((x : A) → C (inj₁ x)) → ((x : B) → C (inj₂ x)) → ((x : A ⊎ B) → C x) [ f , g ] (inj₁ x) = f x [ f , g ] (inj₂ y) = g y [_,_]′ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → (A → C) → (B → C) → (A ⊎ B → C) [_,_]′ = [_,_] map : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → (A → C) → (B → D) → (A ⊎ B → C ⊎ D) map f g = [ inj₁ ∘ f , inj₂ ∘ g ] infixr 1 _-⊎-_ _-⊎-_ : ∀ {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₁ (inj₁ x) = just x isInj₁ (inj₂ y) = nothing isInj₂ : {A B : Set} → A ⊎ B → Maybe B isInj₂ (inj₁ x) = nothing isInj₂ (inj₂ y) = just y lib-0.7/src/Data/ReflexiveClosure.agda0000644000000000000000000000261112101774706016065 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Reflexive closures ------------------------------------------------------------------------ module Data.ReflexiveClosure where open import Data.Unit open import Level open import Relation.Binary open import Relation.Binary.Simple ------------------------------------------------------------------------ -- Reflexive closure data Refl {a ℓ} {A : Set a} (_∼_ : Rel A ℓ) : Rel A (a ⊔ ℓ) where [_] : ∀ {x y} (x∼y : x ∼ y) → Refl _∼_ x y refl : Reflexive (Refl _∼_) -- Map. map : ∀ {a a′ ℓ ℓ′} {A : Set a} {A′ : Set a′} {_R_ : Rel A ℓ} {_R′_ : Rel A′ ℓ′} {f : A → A′} → _R_ =[ f ]⇒ _R′_ → Refl _R_ =[ f ]⇒ Refl _R′_ map R⇒R′ [ xRy ] = [ R⇒R′ xRy ] map R⇒R′ refl = refl -- The reflexive closure has no effect on reflexive relations. drop-refl : ∀ {a ℓ} {A : Set a} {_R_ : Rel A ℓ} → Reflexive _R_ → Refl _R_ ⇒ _R_ drop-refl rfl [ x∼y ] = x∼y drop-refl rfl refl = rfl ------------------------------------------------------------------------ -- Example: Maybe module Maybe where Maybe : ∀ {ℓ} → Set ℓ → Set ℓ Maybe A = Refl (Const A) tt tt nothing : ∀ {a} {A : Set a} → Maybe A nothing = refl just : ∀ {a} {A : Set a} → A → Maybe A just = [_] lib-0.7/src/Data/DifferenceVec.agda0000644000000000000000000000310012101774706015261 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Vectors with fast append ------------------------------------------------------------------------ module Data.DifferenceVec where open import Data.DifferenceNat open import Data.Vec as V using (Vec) open import Function import Data.Nat as N infixr 5 _∷_ _++_ DiffVec : ∀ {ℓ} → Set ℓ → Diffℕ → Set ℓ DiffVec A m = ∀ {n} → Vec A n → Vec A (m n) [] : ∀ {a} {A : Set a} → DiffVec A 0# [] = λ k → k _∷_ : ∀ {a} {A : Set a} {n} → A → DiffVec A n → DiffVec A (suc n) x ∷ xs = λ k → V._∷_ x (xs k) [_] : ∀ {a} {A : Set a} → A → DiffVec A 1# [ x ] = x ∷ [] _++_ : ∀ {a} {A : Set a} {m n} → DiffVec A m → DiffVec A n → DiffVec A (m + n) xs ++ ys = λ k → xs (ys k) toVec : ∀ {a} {A : Set a} {n} → DiffVec A n → Vec A (toℕ n) toVec xs = xs V.[] -- fromVec xs is linear in the length of xs. fromVec : ∀ {a} {A : Set a} {n} → Vec A n → DiffVec A (fromℕ n) fromVec xs = λ k → xs ⟨ V._++_ ⟩ k head : ∀ {a} {A : Set a} {n} → DiffVec A (suc n) → A head xs = V.head (toVec xs) tail : ∀ {a} {A : Set a} {n} → DiffVec A (suc n) → DiffVec A n tail xs = λ k → V.tail (xs k) take : ∀ {a} {A : Set a} m {n} → DiffVec A (fromℕ m + n) → DiffVec A (fromℕ m) take N.zero xs = [] take (N.suc m) xs = head xs ∷ take m (tail xs) drop : ∀ {a} {A : Set a} m {n} → DiffVec A (fromℕ m + n) → DiffVec A n drop N.zero xs = xs drop (N.suc m) xs = drop m (tail xs) lib-0.7/src/Data/Char.agda0000644000000000000000000000264712101774706013465 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Characters ------------------------------------------------------------------------ module Data.Char where 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.Binary import Relation.Binary.On as On open import Relation.Binary.PropositionalEquality as PropEq using (_≡_) open import Relation.Binary.PropositionalEquality.TrustMe ------------------------------------------------------------------------ -- The type postulate Char : Set {-# BUILTIN CHAR Char #-} {-# COMPILED_TYPE Char Char #-} ------------------------------------------------------------------------ -- Operations private primitive primCharToNat : Char → ℕ primCharEquality : Char → Char → Bool toNat : Char → ℕ toNat = primCharToNat infix 4 _==_ _==_ : Char → Char → Bool _==_ = primCharEquality _≟_ : Decidable {A = Char} _≡_ s₁ ≟ s₂ with s₁ == s₂ ... | true = yes trustMe ... | false = no whatever where postulate whatever : _ setoid : Setoid _ _ setoid = PropEq.setoid Char decSetoid : DecSetoid _ _ decSetoid = PropEq.decSetoid _≟_ -- An ordering induced by the toNat function. strictTotalOrder : StrictTotalOrder _ _ _ strictTotalOrder = On.strictTotalOrder NatProp.strictTotalOrder toNat lib-0.7/src/Data/Integer.agda0000644000000000000000000001570012101774706014177 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- Integers ------------------------------------------------------------------------ module Data.Integer where open import Data.Nat as ℕ using (ℕ) renaming (_+_ to _ℕ+_; _*_ to _ℕ*_) import Data.Nat.Show as ℕ open import Data.Sign as Sign using (Sign) renaming (_*_ to _S*_) open import Data.Product as Prod open import Data.String using (String; _++_) open import Function open import Data.Sum as Sum open import Relation.Nullary import Relation.Nullary.Decidable as Dec open import Relation.Binary open import Relation.Binary.PropositionalEquality as PropEq using (_≡_; refl; sym; cong; cong₂) open PropEq.≡-Reasoning infix 8 +_ -_ infixl 7 _*_ _⊓_ infixl 6 _+_ _-_ _⊖_ _⊔_ infix 4 _≤_ _≤?_ ------------------------------------------------------------------------ -- The types -- Integers. data ℤ : Set where -[1+_] : (n : ℕ) → ℤ -- -[1+ n ] stands for - (1 + n). +_ : (n : ℕ) → ℤ -- + n stands for n. ------------------------------------------------------------------------ -- Conversions -- Absolute value. ∣_∣ : ℤ → ℕ ∣ + n ∣ = n ∣ -[1+ n ] ∣ = ℕ.suc n -- Gives the sign. For zero the sign is arbitrarily chosen to be +. sign : ℤ → Sign sign (+ _) = Sign.+ sign -[1+ _ ] = Sign.- -- Decimal string representation. show : ℤ → String show i = showSign (sign i) ++ ℕ.show ∣ i ∣ where showSign : Sign → String showSign Sign.- = "-" showSign Sign.+ = "" ------------------------------------------------------------------------ -- A view of integers as sign + absolute value infix 5 _◂_ _◃_ _◃_ : Sign → ℕ → ℤ _ ◃ ℕ.zero = + ℕ.zero Sign.+ ◃ n = + n Sign.- ◃ ℕ.suc n = -[1+ n ] ◃-left-inverse : ∀ i → sign i ◃ ∣ i ∣ ≡ i ◃-left-inverse -[1+ n ] = refl ◃-left-inverse (+ ℕ.zero) = refl ◃-left-inverse (+ ℕ.suc n) = refl ◃-cong : ∀ {i j} → sign i ≡ sign j → ∣ i ∣ ≡ ∣ j ∣ → i ≡ j ◃-cong {i} {j} sign-≡ abs-≡ = begin i ≡⟨ sym $ ◃-left-inverse i ⟩ sign i ◃ ∣ i ∣ ≡⟨ cong₂ _◃_ sign-≡ abs-≡ ⟩ sign j ◃ ∣ j ∣ ≡⟨ ◃-left-inverse j ⟩ j ∎ data SignAbs : ℤ → Set where _◂_ : (s : Sign) (n : ℕ) → SignAbs (s ◃ n) signAbs : ∀ i → SignAbs i signAbs i = PropEq.subst SignAbs (◃-left-inverse i) $ sign i ◂ ∣ i ∣ ------------------------------------------------------------------------ -- Equality is decidable _≟_ : Decidable {A = ℤ} _≡_ i ≟ j with Sign._≟_ (sign i) (sign j) | ℕ._≟_ ∣ i ∣ ∣ j ∣ i ≟ j | yes sign-≡ | yes abs-≡ = yes (◃-cong sign-≡ abs-≡) i ≟ j | no sign-≢ | _ = no (sign-≢ ∘ cong sign) i ≟ j | _ | no abs-≢ = no (abs-≢ ∘ cong ∣_∣) decSetoid : DecSetoid _ _ decSetoid = PropEq.decSetoid _≟_ ------------------------------------------------------------------------ -- Arithmetic -- Negation. -_ : ℤ → ℤ - (+ ℕ.suc n) = -[1+ n ] - (+ ℕ.zero) = + ℕ.zero - -[1+ n ] = + ℕ.suc n -- Subtraction of natural numbers. _⊖_ : ℕ → ℕ → ℤ m ⊖ ℕ.zero = + m ℕ.zero ⊖ ℕ.suc n = -[1+ n ] ℕ.suc m ⊖ ℕ.suc n = m ⊖ n -- Addition. _+_ : ℤ → ℤ → ℤ -[1+ m ] + -[1+ n ] = -[1+ ℕ.suc (m ℕ+ n) ] -[1+ m ] + + n = n ⊖ ℕ.suc m + m + -[1+ n ] = m ⊖ ℕ.suc n + m + + n = + (m ℕ+ n) -- Subtraction. _-_ : ℤ → ℤ → ℤ i - j = i + - j -- Successor. suc : ℤ → ℤ suc i = + 1 + i private suc-is-lazy⁺ : ∀ n → suc (+ n) ≡ + ℕ.suc n suc-is-lazy⁺ n = refl suc-is-lazy⁻ : ∀ n → suc -[1+ ℕ.suc n ] ≡ -[1+ n ] suc-is-lazy⁻ n = refl -- Predecessor. pred : ℤ → ℤ pred i = - + 1 + i private pred-is-lazy⁺ : ∀ n → pred (+ ℕ.suc n) ≡ + n pred-is-lazy⁺ n = refl pred-is-lazy⁻ : ∀ n → pred -[1+ n ] ≡ -[1+ ℕ.suc n ] pred-is-lazy⁻ n = refl -- Multiplication. _*_ : ℤ → ℤ → ℤ i * j = sign i S* sign j ◃ ∣ i ∣ ℕ* ∣ j ∣ -- Maximum. _⊔_ : ℤ → ℤ → ℤ -[1+ m ] ⊔ -[1+ n ] = -[1+ ℕ._⊓_ m n ] -[1+ m ] ⊔ + n = + n + m ⊔ -[1+ n ] = + m + m ⊔ + n = + (ℕ._⊔_ m n) -- Minimum. _⊓_ : ℤ → ℤ → ℤ -[1+ m ] ⊓ -[1+ n ] = -[1+ ℕ._⊔_ m n ] -[1+ m ] ⊓ + n = -[1+ m ] + m ⊓ -[1+ n ] = -[1+ n ] + m ⊓ + n = + (ℕ._⊓_ m n) ------------------------------------------------------------------------ -- Ordering data _≤_ : ℤ → ℤ → Set where -≤+ : ∀ {m n} → -[1+ m ] ≤ + n -≤- : ∀ {m n} → (n≤m : ℕ._≤_ n m) → -[1+ m ] ≤ -[1+ n ] +≤+ : ∀ {m n} → (m≤n : ℕ._≤_ m n) → + m ≤ + n drop‿+≤+ : ∀ {m n} → + m ≤ + n → ℕ._≤_ m n drop‿+≤+ (+≤+ m≤n) = m≤n drop‿-≤- : ∀ {m n} → -[1+ m ] ≤ -[1+ n ] → ℕ._≤_ n m drop‿-≤- (-≤- n≤m) = n≤m _≤?_ : Decidable _≤_ -[1+ m ] ≤? -[1+ n ] = Dec.map′ -≤- drop‿-≤- (ℕ._≤?_ n m) -[1+ m ] ≤? + n = yes -≤+ + m ≤? -[1+ n ] = no λ () + m ≤? + n = Dec.map′ +≤+ drop‿+≤+ (ℕ._≤?_ m n) decTotalOrder : DecTotalOrder _ _ _ decTotalOrder = record { Carrier = ℤ ; _≈_ = _≡_ ; _≤_ = _≤_ ; isDecTotalOrder = record { isTotalOrder = record { isPartialOrder = record { isPreorder = record { isEquivalence = PropEq.isEquivalence ; reflexive = refl′ ; trans = trans } ; antisym = antisym } ; total = total } ; _≟_ = _≟_ ; _≤?_ = _≤?_ } } where module ℕO = DecTotalOrder ℕ.decTotalOrder refl′ : _≡_ ⇒ _≤_ refl′ { -[1+ n ]} refl = -≤- ℕO.refl refl′ {+ n} refl = +≤+ ℕO.refl trans : Transitive _≤_ trans -≤+ (+≤+ n≤m) = -≤+ trans (-≤- n≤m) -≤+ = -≤+ trans (-≤- n≤m) (-≤- k≤n) = -≤- (ℕO.trans k≤n n≤m) trans (+≤+ m≤n) (+≤+ n≤k) = +≤+ (ℕO.trans m≤n n≤k) antisym : Antisymmetric _≡_ _≤_ antisym -≤+ () antisym (-≤- n≤m) (-≤- m≤n) = cong -[1+_] $ ℕO.antisym m≤n n≤m antisym (+≤+ m≤n) (+≤+ n≤m) = cong +_ $ ℕO.antisym m≤n n≤m total : Total _≤_ total (-[1+ m ]) (-[1+ n ]) = [ inj₂ ∘′ -≤- , inj₁ ∘′ -≤- ]′ $ ℕO.total m n total (-[1+ m ]) (+ n ) = inj₁ -≤+ total (+ m ) (-[1+ n ]) = inj₂ -≤+ total (+ m ) (+ n ) = [ inj₁ ∘′ +≤+ , inj₂ ∘′ +≤+ ]′ $ ℕO.total m n poset : Poset _ _ _ poset = DecTotalOrder.poset decTotalOrder import Relation.Binary.PartialOrderReasoning as POR module ≤-Reasoning = POR poset renaming (_≈⟨_⟩_ to _≡⟨_⟩_) lib-0.7/src/Data/Maybe.agda0000644000000000000000000001226612101774706013643 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- The Maybe type ------------------------------------------------------------------------ module Data.Maybe where open import Level ------------------------------------------------------------------------ -- The type open import Data.Maybe.Core public ------------------------------------------------------------------------ -- Some operations open import Data.Bool using (Bool; true; false) open import Data.Unit using (⊤) boolToMaybe : Bool → Maybe ⊤ boolToMaybe true = just _ boolToMaybe false = nothing maybeToBool : ∀ {a} {A : Set a} → Maybe A → Bool maybeToBool (just _) = true maybeToBool nothing = false -- A dependent eliminator. maybe : ∀ {a b} {A : Set a} {B : Maybe A → Set b} → ((x : A) → B (just x)) → B nothing → (x : Maybe A) → B x maybe j n (just x) = j x maybe j n nothing = n -- A non-dependent eliminator. maybe′ : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → B → Maybe A → B maybe′ = maybe -- A safe variant of "fromJust". If the value is nothing, then the -- return type is the unit type. From-just : ∀ {a} (A : Set a) → Maybe A → Set a From-just A (just _) = A From-just A nothing = Lift ⊤ from-just : ∀ {a} {A : Set a} (x : Maybe A) → From-just A x from-just (just x) = x from-just nothing = _ ------------------------------------------------------------------------ -- Maybe monad open import Function open import Category.Functor open import Category.Monad open import Category.Monad.Identity functor : ∀ {f} → RawFunctor (Maybe {a = f}) functor = record { _<$>_ = λ f → maybe (just ∘ f) nothing } monadT : ∀ {f} {M : Set f → Set f} → RawMonad M → RawMonad (λ A → M (Maybe A)) monadT M = record { return = M.return ∘ just ; _>>=_ = λ m f → M._>>=_ m (maybe f (M.return nothing)) } where module M = RawMonad M monad : ∀ {f} → RawMonad (Maybe {a = f}) monad = monadT IdentityMonad monadZero : ∀ {f} → RawMonadZero (Maybe {a = f}) monadZero = record { monad = monad ; ∅ = nothing } monadPlus : ∀ {f} → RawMonadPlus (Maybe {a = f}) monadPlus {f} = record { monadZero = monadZero ; _∣_ = _∣_ } where _∣_ : {A : Set f} → Maybe A → Maybe A → Maybe A nothing ∣ y = y just x ∣ y = just x ------------------------------------------------------------------------ -- Equality open import Relation.Nullary open import Relation.Binary as B data Eq {a ℓ} {A : Set a} (_≈_ : Rel A ℓ) : Rel (Maybe A) (a ⊔ ℓ) where just : ∀ {x y} (x≈y : x ≈ y) → Eq _≈_ (just x) (just y) nothing : Eq _≈_ nothing nothing drop-just : ∀ {a ℓ} {A : Set a} {_≈_ : Rel A ℓ} {x y : A} → just x ⟨ Eq _≈_ ⟩ just y → x ≈ y drop-just (just x≈y) = x≈y setoid : ∀ {ℓ₁ ℓ₂} → Setoid ℓ₁ ℓ₂ → Setoid _ _ setoid S = record { Carrier = Maybe S.Carrier ; _≈_ = _≈_ ; isEquivalence = record { refl = refl ; sym = sym ; trans = trans } } where module S = Setoid S _≈_ = Eq S._≈_ refl : ∀ {x} → x ≈ x refl {just x} = just S.refl refl {nothing} = nothing sym : ∀ {x y} → x ≈ y → y ≈ x sym (just x≈y) = just (S.sym x≈y) sym nothing = nothing trans : ∀ {x y z} → x ≈ y → y ≈ z → x ≈ z trans (just x≈y) (just y≈z) = just (S.trans x≈y y≈z) trans nothing nothing = nothing decSetoid : ∀ {ℓ₁ ℓ₂} → DecSetoid ℓ₁ ℓ₂ → DecSetoid _ _ decSetoid D = record { isDecEquivalence = record { isEquivalence = Setoid.isEquivalence (setoid (DecSetoid.setoid D)) ; _≟_ = _≟_ } } where _≟_ : B.Decidable (Eq (DecSetoid._≈_ D)) just x ≟ just y with DecSetoid._≟_ D x y just x ≟ just y | yes x≈y = yes (just x≈y) just x ≟ just y | no x≉y = no (x≉y ∘ drop-just) just x ≟ nothing = no λ() nothing ≟ just y = no λ() nothing ≟ nothing = yes nothing ------------------------------------------------------------------------ -- Any and All 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 just : ∀ {x} (px : P x) → Any P (just x) data All {a} {A : Set a} (P : A → Set a) : Maybe A → Set a where just : ∀ {x} (px : P x) → All P (just x) nothing : All P nothing IsJust : ∀ {a} {A : Set a} → Maybe A → Set a IsJust = Any (λ _ → Lift ⊤) IsNothing : ∀ {a} {A : Set a} → Maybe A → Set a IsNothing = All (λ _ → Lift ⊥) private drop-just-Any : ∀ {A} {P : A → Set} {x} → Any P (just x) → P x drop-just-Any (just px) = px drop-just-All : ∀ {A} {P : A → Set} {x} → All P (just x) → P x drop-just-All (just px) = px anyDec : ∀ {A} {P : A → Set} → U.Decidable P → U.Decidable (Any P) anyDec p nothing = no λ() anyDec p (just x) = Dec.map′ just drop-just-Any (p x) allDec : ∀ {A} {P : A → Set} → U.Decidable P → U.Decidable (All P) allDec p nothing = yes nothing allDec p (just x) = Dec.map′ just drop-just-All (p x) lib-0.7/src/Data/AVL.agda0000644000000000000000000004070212101774706013224 0ustar0000000000000000------------------------------------------------------------------------ -- The Agda standard library -- -- AVL trees ------------------------------------------------------------------------ -- AVL trees are balanced binary search trees. -- The search tree invariant is specified using the technique -- described by Conor McBride in his talk "Pivotal pragmatism". open import Relation.Binary open import Relation.Binary.PropositionalEquality as P using (_≡_) module Data.AVL {k v ℓ} {Key : Set k} (Value : Key → Set v) {_<_ : Rel Key ℓ} (isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_) where 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.Nat hiding (_<_; compare; _⊔_) open import Data.Product hiding (map) open import Data.Unit open import Function open import Level using (_⊔_; Lift; lift) open IsStrictTotalOrder isStrictTotalOrder ------------------------------------------------------------------------ -- Extended keys module Extended-key where -- The key type extended with a new minimum and maximum. data Key⁺ : Set k where ⊥⁺ ⊤⁺ : Key⁺ [_] : (k : Key) → Key⁺ -- An extended strict ordering relation. infix 4 _<⁺_ _<⁺_ : Key⁺ → Key⁺ → Set ℓ ⊥⁺ <⁺ [ _ ] = Lift ⊤ ⊥⁺ <⁺ ⊤⁺ = Lift ⊤ [ x ] <⁺ [ y ] = x < y [ _ ] <⁺ ⊤⁺ = Lift ⊤ _ <⁺ _ = Lift ⊥ -- A pair of ordering constraints. infix 4 _<_<_ _<_<_ : Key⁺ → Key → Key⁺ → Set ℓ l < x < u = l <⁺ [ x ] × [ x ] <⁺ u -- _<⁺_ is transitive. trans⁺ : ∀ l {m u} → l <⁺ m → m <⁺ u → l <⁺ u trans⁺ [ l ] {m = [ m ]} {u = [ u ]} l _ _ k′ _ _ _ = joinʳ⁻ _ p lp (delete k pu) bal ... | tri≈ _ _ _ = join lp pu bal -- Looks up a key. Logarithmic in the size of the tree (assuming -- constant-time comparisons). lookup : ∀ {l u h} → (k : Key) → Tree l u h → Maybe (Value k) lookup k (leaf _) = nothing lookup k (node (k′ , v) lk′ k′u _) with compare k k′ ... | tri< _ _ _ = lookup k lk′ ... | tri> _ _ _ = lookup k k′u ... | tri≈ _ eq _ rewrite eq = just v -- Maps a function over all values in the tree. map : (∀ {k} → Value k → Value k) → ∀ {l u h} → Tree l u h → Tree l u h map f (leaf l