summaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorGianfranco Costamagna <costamagnagianfranco@yahoo.it>2017-11-26 11:35:06 +0100
committerGianfranco Costamagna <costamagnagianfranco@yahoo.it>2017-11-26 11:35:06 +0100
commit0b3946998d82e1be6c50e2872e7474db69b5a0dc (patch)
treec935ec365b32a5d11dddea90cc53aa9e7ad9bf3c
parentc7745ec5e54ac1945fe7161ee9d6e47b10cc59bd (diff)
New upstream version 0.14
-rw-r--r--.dropbox_uploader.enc1
-rw-r--r--.travis.yml70
-rw-r--r--CHANGELOG.md773
-rw-r--r--GNUmakefile9
-rw-r--r--HACKING31
-rw-r--r--HACKING.md45
-rw-r--r--LICENCE5
-rw-r--r--README.agda16
-rw-r--r--README/AVL.agda6
-rw-r--r--README/Nat.agda13
-rw-r--r--doc/release-notes/future.txt2
-rw-r--r--lib.cabal18
-rw-r--r--notes/stdlib-releases.txt8
-rw-r--r--src/Algebra/CommutativeMonoidSolver.agda5
-rw-r--r--src/Algebra/FunctionProperties.agda9
-rw-r--r--src/Algebra/FunctionProperties/Consequences.agda113
-rw-r--r--src/Algebra/Properties/BooleanAlgebra.agda675
-rw-r--r--src/Algebra/Properties/BooleanAlgebra/Expression.agda18
-rw-r--r--src/Algebra/Properties/DistributiveLattice.agda52
-rw-r--r--src/Category/Applicative/Indexed.agda6
-rw-r--r--src/Category/Functor.agda14
-rw-r--r--src/Category/Functor/Identity.agda17
-rw-r--r--src/Data/Bin.agda192
-rw-r--r--src/Data/Bin/Properties.agda141
-rw-r--r--src/Data/BoundedVec.agda45
-rw-r--r--src/Data/Char.agda4
-rw-r--r--src/Data/Colist.agda6
-rw-r--r--src/Data/Colist/Infinite-merge.agda4
-rw-r--r--src/Data/Container.agda10
-rw-r--r--src/Data/Container/Combinator.agda2
-rw-r--r--src/Data/Container/Indexed.agda12
-rw-r--r--src/Data/Digit.agda52
-rw-r--r--src/Data/Empty.agda4
-rw-r--r--src/Data/Empty/Irrelevant.agda12
-rw-r--r--src/Data/Fin.agda29
-rw-r--r--src/Data/Fin/Dec.agda8
-rw-r--r--src/Data/Fin/Properties.agda223
-rw-r--r--src/Data/Fin/Subset.agda10
-rw-r--r--src/Data/Fin/Subset/Properties.agda146
-rw-r--r--src/Data/Fin/Substitution/Lemmas.agda9
-rw-r--r--src/Data/Graph/Acyclic.agda6
-rw-r--r--src/Data/Integer.agda67
-rw-r--r--src/Data/Integer/Addition/Properties.agda126
-rw-r--r--src/Data/Integer/Base.agda27
-rw-r--r--src/Data/Integer/Multiplication/Properties.agda96
-rw-r--r--src/Data/Integer/Properties.agda743
-rw-r--r--src/Data/List/All.agda4
-rw-r--r--src/Data/List/All/Properties.agda298
-rw-r--r--src/Data/List/Any.agda192
-rw-r--r--src/Data/List/Any/BagAndSetEquality.agda3
-rw-r--r--src/Data/List/Any/Membership.agda286
-rw-r--r--src/Data/List/Any/Membership/Properties.agda72
-rw-r--r--src/Data/List/Any/Membership/Propositional.agda84
-rw-r--r--src/Data/List/Any/Membership/Propositional/Properties.agda286
-rw-r--r--src/Data/List/Any/Properties.agda718
-rw-r--r--src/Data/List/Base.agda35
-rw-r--r--src/Data/List/Countdown.agda45
-rw-r--r--src/Data/List/Properties.agda358
-rw-r--r--src/Data/List/Reverse.agda6
-rw-r--r--src/Data/Maybe.agda98
-rw-r--r--src/Data/Maybe/Base.agda4
-rw-r--r--src/Data/Nat.agda55
-rw-r--r--src/Data/Nat/Base.agda21
-rw-r--r--src/Data/Nat/Coprimality.agda9
-rw-r--r--src/Data/Nat/DivMod.agda9
-rw-r--r--src/Data/Nat/Divisibility.agda242
-rw-r--r--src/Data/Nat/GCD.agda21
-rw-r--r--src/Data/Nat/GCD/Lemmas.agda2
-rw-r--r--src/Data/Nat/GeneralisedArithmetic.agda88
-rw-r--r--src/Data/Nat/InfinitelyOften.agda12
-rw-r--r--src/Data/Nat/LCM.agda7
-rw-r--r--src/Data/Nat/Properties.agda1240
-rw-r--r--src/Data/Nat/Properties/Simple.agda118
-rw-r--r--src/Data/Product.agda10
-rw-r--r--src/Data/Product/N-ary.agda2
-rw-r--r--src/Data/Rational.agda71
-rw-r--r--src/Data/Rational/Properties.agda97
-rw-r--r--src/Data/Sign/Properties.agda47
-rw-r--r--src/Data/Stream.agda4
-rw-r--r--src/Data/Sum.agda21
-rw-r--r--src/Data/Vec.agda14
-rw-r--r--src/Data/Vec/All.agda36
-rw-r--r--src/Data/Vec/All/Properties.agda214
-rw-r--r--src/Data/Vec/Equality.agda14
-rw-r--r--src/Data/Vec/Properties.agda99
-rw-r--r--src/Foreign/Haskell.agda4
-rw-r--r--src/Function.agda18
-rw-r--r--src/IO/Primitive.agda46
-rw-r--r--src/Induction/Nat.agda62
-rw-r--r--src/Irrelevance.agda9
-rw-r--r--src/Relation/Binary/Consequences.agda5
-rw-r--r--src/Relation/Binary/HeterogeneousEquality.agda10
-rw-r--r--src/Relation/Binary/StrictPartialOrderReasoning.agda8
-rw-r--r--src/Relation/Binary/Vec/Pointwise.agda4
-rw-r--r--src/Relation/Nullary/Negation.agda9
-rw-r--r--src/Strict.agda29
96 files changed, 5688 insertions, 3338 deletions
diff --git a/.dropbox_uploader.enc b/.dropbox_uploader.enc
deleted file mode 100644
index 8dc5135..0000000
--- a/.dropbox_uploader.enc
+++ /dev/null
@@ -1 +0,0 @@
-$%7 "6_TXNݤPp~ *'O^6<W)aQ Ak_9c -)hfFה9C>_[AYo@o%674ugT%";Bv`Rfo \ No newline at end of file
diff --git a/.travis.yml b/.travis.yml
index 98571a2..f617871 100644
--- a/.travis.yml
+++ b/.travis.yml
@@ -1,26 +1,60 @@
-language: haskell
+language: c
branches:
only:
- master
+sudo: false
+
+dist: trusty
+
+cache:
+ directories:
+ - $HOME/.cabsnap
+
+matrix:
+ include:
+ - env: TEST=MAIN GHC_VER=8.0.2 BUILD=CABAL CABAL_VER=1.24
+ addons:
+ apt:
+ packages:
+ - alex-3.1.7
+ - cabal-install-1.24
+ - ghc-8.0.2
+ - happy-1.19.5
+ sources:
+ - hvr-ghc
+
before_install:
- # Decrypting the dropbox credentials
- - openssl aes-256-cbc -K $encrypted_3f6e0f4132ed_key -iv $encrypted_3f6e0f4132ed_iv -in .dropbox_uploader.enc -out .dropbox_uploader -d
- - cp .dropbox_uploader $HOME/
+ - export PATH=/opt/ghc/$GHC_VER/bin:/opt/cabal/$CABAL_VER/bin:/opt/alex/3.1.7/bin:/opt/happy/1.19.5/bin:~/.cabal/bin/:$PATH;
install:
- # Installing the dropbox download script
- - cd $HOME
- - curl "https://raw.githubusercontent.com/andreafabrizi/Dropbox-Uploader/master/dropbox_uploader.sh" -o dropbox_uploader.sh
- - chmod +x dropbox_uploader.sh
- # downloading the last (latest?) tarball and installing agda
- - AGDATAR=$(./dropbox_uploader.sh list | tail -n 1 | sed 's/^.*\(Agda.*\)$/\1/')
- - ./dropbox_uploader.sh download $AGDATAR .
- - tar -xzvf $AGDATAR
- - cd $( basename $AGDATAR .tar.gz )
- - sudo ./deploy.sh
+ - git clone https://github.com/agda/agda --depth=1 --single-branch
+ - cabal update
+ - sed -i 's/^jobs:/-- jobs:/' $HOME/.cabal/config
+ # checking whether .ghc is still valid
+ - cabal install --only-dependencies --dry -v > $HOME/installplan.txt
+ - sed -i -e '1,/^Resolving /d' $HOME/installplan.txt; cat $HOME/installplan.txt
+ - touch $HOME/.cabsnap/intallplan.txt
+ - mkdir -p $HOME/.cabsnap/ghc $HOME/.cabsnap/lib $HOME/.cabsnap/share $HOME/.cabsnap/bin
+ - if diff -u $HOME/.cabsnap/installplan.txt $HOME/installplan.txt;
+ then
+ echo "cabal build-cache HIT";
+ rm -rfv .ghc;
+ cp -a $HOME/.cabsnap/ghc $HOME/.ghc;
+ cp -a $HOME/.cabsnap/lib $HOME/.cabsnap/share $HOME/.cabsnap/bin $HOME/.cabal/;
+ else
+ echo "cabal build-cache MISS";
+ rm -rf $HOME/.cabsnap;
+ mkdir -p $HOME/.ghc $HOME/.cabal/lib $HOME/.cabal/share $HOME/.cabal/bin;
+ fi
+ - cabal install cpphs
+ - cd agda && cabal install --only-dependencies && make CABAL_OPTS=-v2 install-bin
+ # snapshot package-db on cache miss
+ - echo "snapshotting package-db to build-cache";
+ mkdir $HOME/.cabsnap;
+ cp -a $HOME/.ghc $HOME/.cabsnap/ghc;
+ cp -a $HOME/.cabal/lib $HOME/.cabal/share $HOME/.cabal/bin $HOME/installplan.txt $HOME/.cabsnap/;
# generating Everything.agda
- - cabal install filemanip
- cd $HOME/build/agda/agda-stdlib
- runghc GenerateEverything.hs
@@ -41,12 +75,8 @@ after_success:
- git fetch upstream && git reset upstream/gh-pages
- git add -f \*.html
- git commit -m "Automatic HTML update via Travis"
- - git push -q upstream HEAD:gh-pages &>/dev/null
+ - if [ "$TRAVIS_PULL_REQUEST" = "false" ]; then git push -q upstream HEAD:gh-pages &>/dev/null; fi
notifications:
email: false
-# Github token
-env:
- global:
- secure: f0GRSKTlAw6FdrvkI8LjP5ZhwcGBltJ1t5+nxipqfoR3VczMxcNgXD7bdwtUojFYR48jMs2W+LxEoJcuSH4bsdCeoSfpZ0nU424MDajGh0wuhAbSYKQXPWjFOxwMgC23sCySKhDAZiqN/Wd6orwV1p5JhuJkSCHdVeyUp+hLvIw=
diff --git a/CHANGELOG.md b/CHANGELOG.md
index 89d843f..fc25e91 100644
--- a/CHANGELOG.md
+++ b/CHANGELOG.md
@@ -1,3 +1,776 @@
+Version 0.14
+============
+
+The library has been tested using Agda version 2.5.3.
+
+Non-backwards compatible changes
+--------------------------------
+
+#### 1st stage of overhaul of list membership
+
+* The current setup for list membership is difficult to work with as both setoid membership
+ and propositional membership exist as internal modules of `Data.Any`. Furthermore the
+ top-level module `Data.List.Any.Membership` actually contains properties of propositional
+ membership rather than the membership relation itself as its name would suggest.
+ Consequently this leaves no place to reason about the properties of setoid membership.
+
+ Therefore the two internal modules `Membership` and `Membership-≡` have been moved out
+ of `Data.List.Any` into top-level `Data.List.Any.Membership` and
+ `Data.List.Any.Membership.Propositional` respectively. The previous module
+ `Data.List.Any.Membership` has been renamed
+ `Data.List.Any.Membership.Propositional.Properties`.
+
+ Accordingly some lemmas have been moved to more logical locations:
+ - `lift-resp` has been moved from `Data.List.Any.Membership` to `Data.List.Any.Properties`
+ - `∈-resp-≈`, `⊆-preorder` and `⊆-Reasoning` have been moved from `Data.List.Any.Membership`
+ to `Data.List.Any.Membership.Properties`.
+ - `∈-resp-list-≈` has been moved from `Data.List.Any.Membership` to
+ `Data.List.Any.Membership.Properties` and renamed `∈-resp-≋`.
+ - `swap` in `Data.List.Any.Properties` has been renamed `swap↔` and made more generic with
+ respect to levels.
+
+#### Moving `decTotalOrder` and `decSetoid` from `Data.X` to `Data.X.Properties`
+
+* Currently the library does not directly expose proofs of basic properties such as reflexivity,
+ transitivity etc. for `_≤_` in numeric datatypes such as `Nat`, `Integer` etc. In order to use these
+ properties it was necessary to first import the `decTotalOrder` proof from `Data.X` and then
+ separately open it, often having to rename the proofs as well. This adds unneccessary lines of
+ code to the import statements for what are very commonly used properties.
+
+ These basic proofs have now been added in `Data.X.Properties` along with proofs that they form
+ pre-orders, partial orders and total orders. This should make them considerably easier to work with
+ and simplify files' import preambles. However consequently the records `decTotalOrder` and
+ `decSetoid` have been moved from `Data.X` to `≤-decTotalOrder` and `≡-decSetoid` in
+ `Data.X.Properties`.
+
+ The numeric datatypes for which this has been done are `Nat`, `Integer`, `Rational` and `Bin`.
+
+ As a consequence the module `≤-Reasoning` has also had to have been moved from `Data.Nat` to
+ `Data.Nat.Properties`.
+
+#### New well-founded induction proofs for `Data.Nat`
+
+* Currently `Induction.Nat` only proves that the non-standard `_<′_`relation over `ℕ` is
+ well-founded. Unfortunately these existing proofs are named `<-Rec` and `<-well-founded`
+ which clash with the sensible names for new proofs over the standard `_<_` relation.
+
+ Therefore `<-Rec` and `<-well-founded` have been renamed to `<′-Rec` and `<′-well-founded`
+ respectively. The original names `<-Rec` and `<-well-founded` now refer to new
+ corresponding proofs for `_<_`.
+
+#### Other
+
+* Changed the implementation of `map` and `zipWith` in `Data.Vec` to use native
+ (pattern-matching) definitions. Previously they were defined using the
+ `applicative` operations of `Vec`. The new definitions can be converted back
+ to the old using the new proofs `⊛-is-zipWith`, `map-is-⊛` and `zipWith-is-⊛`
+ in `Data.Vec.Properties`. It has been argued that `zipWith` is fundamental than `_⊛_`
+ and this change allows better printing of goals involving `map` or `zipWith`.
+
+* Changed the implementation of `All₂` in `Data.Vec.All` to a native datatype. This
+ improved improves pattern matching on terms and allows the new datatype to be more
+ generic with respect to types and levels.
+
+* Changed the implementation of `downFrom` in `Data.List` to a native
+ (pattern-matching) definition. Previously it was defined using a private
+ internal module which made pattern matching difficult.
+
+* The arguments of `≤pred⇒≤` and `≤⇒pred≤` in `Data.Nat.Properties` are now implicit
+ rather than explicit (was `∀ m n → m ≤ pred n → m ≤ n` and is now
+ `∀ {m n} → m ≤ pred n → m ≤ n`). This makes it consistent with `<⇒≤pred` which
+ already used implicit arguments, and shouldn't introduce any significant problems
+ as both parameters can be inferred by Agda.
+
+* Moved `¬∀⟶∃¬` from `Relation.Nullary.Negation` to `Data.Fin.Dec`. Its old
+ location was causing dependency cyles to form between `Data.Fin.Dec`,
+ `Relation.Nullary.Negation` and `Data.Fin`.
+
+* Moved `fold`, `add` and `mul` from `Data.Nat` to new module `Data.Nat.GeneralisedArithmetic`.
+
+* Changed type of second parameter of `Relation.Binary.StrictPartialOrderReasoning._<⟨_⟩_`
+ from `x < y ⊎ x ≈ y` to `x < y`. `_≈⟨_⟩_` is left unchanged to take a value with type `x ≈ y`.
+ Old code may be fixed by prefixing the contents of `_<⟨_⟩_` with `inj₁`.
+
+Deprecated features
+-------------------
+
+Deprecated features still exist and therefore existing code should still work
+but they may be removed in some future release of the library.
+
+* The module `Data.Nat.Properties.Simple` is now deprecated. All proofs
+ have been moved to `Data.Nat.Properties` where they should be used directly.
+ The `Simple` file still exists for backwards compatability reasons and
+ re-exports the proofs from `Data.Nat.Properties` but will be removed in some
+ future release.
+
+* The modules `Data.Integer.Addition.Properties` and
+ `Data.Integer.Multiplication.Properties` are now deprecated. All proofs
+ have been moved to `Data.Integer.Properties` where they should be used
+ directly. The `Addition.Properties` and `Multiplication.Properties` files
+ still exist for backwards compatability reasons and re-exports the proofs from
+ `Data.Integer.Properties` but will be removed in some future release.
+
+* The following renaming has occured in `Data.Nat.Properties`
+ ```agda
+ _+-mono_ ↦ +-mono-≤
+ _*-mono_ ↦ *-mono-≤
+
+ +-right-identity ↦ +-identityʳ
+ *-right-zero ↦ *-zeroʳ
+ distribʳ-*-+ ↦ *-distribʳ-+
+ *-distrib-∸ʳ ↦ *-distribʳ-∸
+ cancel-+-left ↦ +-cancelˡ-≡
+ cancel-+-left-≤ ↦ +-cancelˡ-≤
+ cancel-*-right ↦ *-cancelʳ-≡
+ cancel-*-right-≤ ↦ *-cancelʳ-≤
+
+ strictTotalOrder ↦ <-strictTotalOrder
+ isCommutativeSemiring ↦ *-+-isCommutativeSemiring
+ commutativeSemiring ↦ *-+-commutativeSemiring
+ isDistributiveLattice ↦ ⊓-⊔-isDistributiveLattice
+ distributiveLattice ↦ ⊓-⊔-distributiveLattice
+ ⊔-⊓-0-isSemiringWithoutOne ↦ ⊔-⊓-isSemiringWithoutOne
+ ⊔-⊓-0-isCommutativeSemiringWithoutOne ↦ ⊔-⊓-isCommutativeSemiringWithoutOne
+ ⊔-⊓-0-commutativeSemiringWithoutOne ↦ ⊔-⊓-commutativeSemiringWithoutOne
+ ```
+
+* The following renaming has occurred in `Data.Nat.Divisibility`:
+ ```agda
+ ∣-* ↦ n|m*n
+ ∣-+ ↦ ∣m∣n⇒∣m+n
+ ∣-∸ ↦ ∣m+n|m⇒|n
+ ```
+
+Backwards compatible changes
+----------------------------
+
+* Added support for GHC 8.0.2 and 8.2.1.
+
+* Removed the empty `Irrelevance` module
+
+* Added `Category.Functor.Morphism` and module `Category.Functor.Identity`.
+
+* `Data.Container` and `Data.Container.Indexed` now allow for different
+ levels in the container and in the data it contains.
+
+* Made `Data.BoundedVec` polymorphic with respect to levels.
+
+* Access to `primForce` and `primForceLemma` has been provided via the new
+ top-level module `Strict`.
+
+* New call-by-value application combinator `_$!_` in `Function`.
+
+* Added properties to `Algebra.FunctionProperties`:
+ ```agda
+ LeftCancellative _•_ = ∀ x {y z} → (x • y) ≈ (x • z) → y ≈ z
+ RightCancellative _•_ = ∀ {x} y z → (y • x) ≈ (z • x) → y ≈ z
+ Cancellative _•_ = LeftCancellative _•_ × RightCancellative _•_
+ ```
+
+* Added new module `Algebra.FunctionProperties.Consequences` for basic causal relationships between
+ properties, containing:
+ ```agda
+ comm+idˡ⇒idʳ : Commutative _•_ → LeftIdentity e _•_ → RightIdentity e _•_
+ comm+idʳ⇒idˡ : Commutative _•_ → RightIdentity e _•_ → LeftIdentity e _•_
+ comm+zeˡ⇒zeʳ : Commutative _•_ → LeftZero e _•_ → RightZero e _•_
+ comm+zeʳ⇒zeˡ : Commutative _•_ → RightZero e _•_ → LeftZero e _•_
+ comm+invˡ⇒invʳ : Commutative _•_ → LeftInverse e _⁻¹ _•_ → RightInverse e _⁻¹ _•_
+ comm+invʳ⇒invˡ : Commutative _•_ → RightInverse e _⁻¹ _•_ → LeftInverse e _⁻¹ _•_
+ comm+distrˡ⇒distrʳ : Commutative _•_ → _•_ DistributesOverˡ _◦_ → _•_ DistributesOverʳ _◦_
+ comm+distrʳ⇒distrˡ : Commutative _•_ → _•_ DistributesOverʳ _◦_ → _•_ DistributesOverˡ _◦_
+ comm+cancelˡ⇒cancelʳ : Commutative _•_ → LeftCancellative _•_ → RightCancellative _•_
+ comm+cancelˡ⇒cancelʳ : Commutative _•_ → LeftCancellative _•_ → RightCancellative _•_
+ sel⇒idem : Selective _•_ → Idempotent _•_
+ ```
+
+* Added proofs to `Algebra.Properties.BooleanAlgebra`:
+ ```agda
+ ∨-complementˡ : LeftInverse ⊤ ¬_ _∨_
+ ∧-complementˡ : LeftInverse ⊥ ¬_ _∧_
+
+ ∧-identityʳ : RightIdentity ⊤ _∧_
+ ∧-identityˡ : LeftIdentity ⊤ _∧_
+ ∧-identity : Identity ⊤ _∧_
+
+ ∨-identityʳ : RightIdentity ⊥ _∨_
+ ∨-identityˡ : LeftIdentity ⊥ _∨_
+ ∨-identity : Identity ⊥ _∨_
+
+ ∧-zeroʳ : RightZero ⊥ _∧_
+ ∧-zeroˡ : LeftZero ⊥ _∧_
+ ∧-zero : Zero ⊥ _∧_
+
+ ∨-zeroʳ : RightZero ⊤ _∨_
+ ∨-zeroˡ : LeftZero ⊤ _∨_
+ ∨-zero : Zero ⊤ _∨_
+
+ ⊕-identityˡ : LeftIdentity ⊥ _⊕_
+ ⊕-identityʳ : RightIdentity ⊥ _⊕_
+ ⊕-identity : Identity ⊥ _⊕_
+
+ ⊕-inverseˡ : LeftInverse ⊥ id _⊕_
+ ⊕-inverseʳ : RightInverse ⊥ id _⊕_
+ ⊕-inverse : Inverse ⊥ id _⊕_
+
+ ⊕-cong : Congruent₂ _⊕_
+ ⊕-comm : Commutative _⊕_
+ ⊕-assoc : Associative _⊕_
+
+ ∧-distribˡ-⊕ : _∧_ DistributesOverˡ _⊕_
+ ∧-distribʳ-⊕ : _∧_ DistributesOverʳ _⊕_
+ ∧-distrib-⊕ : _∧_ DistributesOver _⊕_
+
+ ∨-isSemigroup : IsSemigroup _≈_ _∨_
+ ∧-isSemigroup : IsSemigroup _≈_ _∧_
+ ∨-⊥-isMonoid : IsMonoid _≈_ _∨_ ⊥
+ ∧-⊤-isMonoid : IsMonoid _≈_ _∧_ ⊤
+ ∨-⊥-isCommutativeMonoid : IsCommutativeMonoid _≈_ _∨_ ⊥
+ ∧-⊤-isCommutativeMonoid : IsCommutativeMonoid _≈_ _∧_ ⊤
+
+ ⊕-isSemigroup : IsSemigroup _≈_ _⊕_
+ ⊕-⊥-isMonoid : IsMonoid _≈_ _⊕_ ⊥
+ ⊕-⊥-isGroup : IsGroup _≈_ _⊕_ ⊥ id
+ ⊕-⊥-isAbelianGroup : IsAbelianGroup _≈_ _⊕_ ⊥ id
+ ⊕-∧-isRing : IsRing _≈_ _⊕_ _∧_ id ⊥ ⊤
+ ```
+
+* Added proofs to `Algebra.Properties.DistributiveLattice`:
+ ```agda
+ ∨-∧-distribˡ : _∨_ DistributesOverˡ _∧_
+ ∧-∨-distribˡ : _∧_ DistributesOverˡ _∨_
+ ∧-∨-distribʳ : _∧_ DistributesOverʳ _∨_
+ ```
+
+* Added pattern synonyms to `Data.Bin` to improve readability:
+ ```agda
+ pattern 0b = zero
+ pattern 1b = 1+ zero
+ pattern ⊥b = 1+ 1+ ()
+ ```
+
+* A new module `Data.Bin.Properties` has been added, containing proofs:
+ ```agda
+ 1#-injective : as 1# ≡ bs 1# → as ≡ bs
+ _≟_ : Decidable {A = Bin} _≡_
+ ≡-isDecEquivalence : IsDecEquivalence _≡_
+ ≡-decSetoid : DecSetoid _ _
+
+ <-trans : Transitive _<_
+ <-asym : Asymmetric _<_
+ <-irrefl : Irreflexive _≡_ _<_
+ <-cmp : Trichotomous _≡_ _<_
+ <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
+
+ <⇒≢ : a < b → a ≢ b
+ 1<[23] : [] 1# < (b ∷ []) 1#
+ 1<2+ : [] 1# < (b ∷ bs) 1#
+ 0<1+ : 0# < bs 1#
+ ```
+
+* Added functions to `Data.BoundedVec`:
+ ```agda
+ toInefficient : BoundedVec A n → Ineff.BoundedVec A n
+ fromInefficient : Ineff.BoundedVec A n → BoundedVec A n
+ ```
+
+* Added the following to `Data.Digit`:
+ ```agda
+ Expansion : ℕ → Set
+ Expansion base = List (Fin base)
+ ```
+
+* Added new module `Data.Empty.Irrelevant` containing an irrelevant version of `⊥-elim`.
+
+* Added functions to `Data.Fin`:
+ ```agda
+ punchIn i j ≈ if j≥i then j+1 else j
+ punchOut i j ≈ if j>i then j-1 else j
+ ```
+
+* Added proofs to `Data.Fin.Properties`:
+ ```agda
+ isDecEquivalence : ∀ {n} → IsDecEquivalence (_≡_ {A = Fin n})
+
+ ≤-reflexive : ∀ {n} → _≡_ ⇒ (_≤_ {n})
+ ≤-refl : ∀ {n} → Reflexive (_≤_ {n})
+ ≤-trans : ∀ {n} → Transitive (_≤_ {n})
+ ≤-antisymmetric : ∀ {n} → Antisymmetric _≡_ (_≤_ {n})
+ ≤-total : ∀ {n} → Total (_≤_ {n})
+ ≤-isPreorder : ∀ {n} → IsPreorder _≡_ (_≤_ {n})
+ ≤-isPartialOrder : ∀ {n} → IsPartialOrder _≡_ (_≤_ {n})
+ ≤-isTotalOrder : ∀ {n} → IsTotalOrder _≡_ (_≤_ {n})
+
+ _<?_ : ∀ {n} → Decidable (_<_ {n})
+ <-trans : ∀ {n} → Transitive (_<_ {n})
+ <-isStrictTotalOrder : ∀ {n} → IsStrictTotalOrder _≡_ (_<_ {n})
+
+ punchOut-injective : punchOut i≢j ≡ punchOut i≢k → j ≡ k
+ punchIn-injective : punchIn i j ≡ punchIn i k → j ≡ k
+ punchIn-punchOut : punchIn i (punchOut i≢j) ≡ j
+ punchInᵢ≢i : punchIn i j ≢ i
+ ```
+
+* Added proofs to `Data.Fin.Subset.Properties`:
+ ```agda
+ x∈⁅x⁆ : x ∈ ⁅ x ⁆
+ x∈⁅y⁆⇒x≡y : x ∈ ⁅ y ⁆ → x ≡ y
+
+ ∪-assoc : Associative _≡_ _∪_
+ ∩-assoc : Associative _≡_ _∩_
+ ∪-comm : Commutative _≡_ _∪_
+ ∩-comm : Commutative _≡_ _∩_
+
+ p⊆p∪q : p ⊆ p ∪ q
+ q⊆p∪q : q ⊆ p ∪ q
+ x∈p∪q⁻ : x ∈ p ∪ q → x ∈ p ⊎ x ∈ q
+ x∈p∪q⁺ : x ∈ p ⊎ x ∈ q → x ∈ p ∪ q
+
+ p∩q⊆p : p ∩ q ⊆ p
+ p∩q⊆q : p ∩ q ⊆ q
+ x∈p∩q⁺ : x ∈ p × x ∈ q → x ∈ p ∩ q
+ x∈p∩q⁻ : x ∈ p ∩ q → x ∈ p × x ∈ q
+ ∩⇔× : x ∈ p ∩ q ⇔ (x ∈ p × x ∈ q)
+ ```
+
+* Added relations to `Data.Integer`
+ ```agda
+ _≥_ : Rel ℤ _
+ _<_ : Rel ℤ _
+ _>_ : Rel ℤ _
+ _≰_ : Rel ℤ _
+ _≱_ : Rel ℤ _
+ _≮_ : Rel ℤ _
+ _≯_ : Rel ℤ _
+ ```
+
+* Added proofs to `Data.Integer.Properties`
+ ```agda
+ +-injective : + m ≡ + n → m ≡ n
+ -[1+-injective : -[1+ m ] ≡ -[1+ n ] → m ≡ n
+
+ doubleNeg : - - n ≡ n
+ neg-injective : - m ≡ - n → m ≡ n
+
+ ∣n∣≡0⇒n≡0 : ∣ n ∣ ≡ 0 → n ≡ + 0
+ ∣-n∣≡∣n∣ : ∣ - n ∣ ≡ ∣ n ∣
+
+ +◃n≡+n : Sign.+ ◃ n ≡ + n
+ -◃n≡-n : Sign.- ◃ n ≡ - + n
+ signₙ◃∣n∣≡n : sign n ◃ ∣ n ∣ ≡ n
+ ∣s◃m∣*∣t◃n∣≡m*n : ∣ s ◃ m ∣ ℕ* ∣ t ◃ n ∣ ≡ m ℕ* n
+
+ ⊖-≰ : n ≰ m → m ⊖ n ≡ - + (n ∸ m)
+ ∣⊖∣-≰ : n ≰ m → ∣ m ⊖ n ∣ ≡ n ∸ m
+ sign-⊖-≰ : n ≰ m → sign (m ⊖ n) ≡ Sign.-
+ -[n⊖m]≡-m+n : - (m ⊖ n) ≡ (- (+ m)) + (+ n)
+
+ +-identity : Identity (+ 0) _+_
+ +-inverse : Inverse (+ 0) -_ _+_
+ +-0-isMonoid : IsMonoid _≡_ _+_ (+ 0)
+ +-0-isGroup : IsGroup _≡_ _+_ (+ 0) (-_)
+ +-0-abelianGroup : AbelianGroup _ _
+
+ n≢1+n : n ≢ suc n
+ 1-[1+n]≡-n : suc -[1+ n ] ≡ - (+ n)
+ neg-distrib-+ : - (m + n) ≡ (- m) + (- n)
+ ◃-distrib-+ : s ◃ (m + n) ≡ (s ◃ m) + (s ◃ n)
+
+ *-identityʳ : RightIdentity (+ 1) _*_
+ *-identity : Identity (+ 1) _*_
+ *-zeroˡ : LeftZero (+ 0) _*_
+ *-zeroʳ : RightZero (+ 0) _*_
+ *-zero : Zero (+ 0) _*_
+ *-1-isMonoid : IsMonoid _≡_ _*_ (+ 1)
+ -1*n≡-n : -[1+ 0 ] * n ≡ - n
+ ◃-distrib-* : (s 𝕊* t) ◃ (m ℕ* n) ≡ (s ◃ m) * (t ◃ n)
+
+ +-*-isRing : IsRing _≡_ _+_ _*_ -_ (+ 0) (+ 1)
+ +-*-isCommutativeRing : IsCommutativeRing _≡_ _+_ _*_ -_ (+ 0) (+ 1)
+
+ ≤-reflexive : _≡_ ⇒ _≤_
+ ≤-refl : Reflexive _≤_
+ ≤-trans : Transitive _≤_
+ ≤-antisym : Antisymmetric _≡_ _≤_
+ ≤-total : Total _≤_
+
+ ≤-isPreorder : IsPreorder _≡_ _≤_
+ ≤-isPartialOrder : IsPartialOrder _≡_ _≤_
+ ≤-isTotalOrder : IsTotalOrder _≡_ _≤_
+ ≤-isDecTotalOrder : IsDecTotalOrder _≡_ _≤_
+
+ ≤-step : n ≤ m → n ≤ suc m
+ n≤1+n : n ≤ + 1 + n
+
+ <-irrefl : Irreflexive _≡_ _<_
+ <-asym : Asymmetric _<_
+ <-trans : Transitive _<_
+ <-cmp : Trichotomous _≡_ _<_
+ <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
+
+ n≮n : n ≮ n
+ -<+ : -[1+ m ] < + n
+ <⇒≤ : m < n → m ≤ n
+ ≰→> : x ≰ y → x > y
+ ```
+
+* Added functions to `Data.List`
+ ```agda
+ applyUpTo f n ≈ f[0] ∷ f[1] ∷ ... ∷ f[n-1] ∷ []
+ upTo n ≈ 0 ∷ 1 ∷ ... ∷ n-1 ∷ []
+ applyDownFrom f n ≈ f[n-1] ∷ f[n-2] ∷ ... ∷ f[0] ∷ []
+ tabulate f ≈ f[0] ∷ f[1] ∷ ... ∷ f[n-1] ∷ []
+ allFin n ≈ 0f ∷ 1f ∷ ... ∷ n-1f ∷ []
+ ```
+
+* Added proofs to `Data.List.Properties`
+ ```agda
+ map-id₂ : All (λ x → f x ≡ x) xs → map f xs ≡ xs
+ map-cong₂ : All (λ x → f x ≡ g x) xs → map f xs ≡ map g xs
+ foldr-++ : foldr f x (ys ++ zs) ≡ foldr f (foldr f x zs) ys
+ foldl-++ : foldl f x (ys ++ zs) ≡ foldl f (foldl f x ys) zs
+ foldr-∷ʳ : foldr f x (ys ∷ʳ y) ≡ foldr f (f y x) ys
+ foldl-∷ʳ : foldl f x (ys ∷ʳ y) ≡ f (foldl f x ys) y
+ reverse-foldr : foldr f x (reverse ys) ≡ foldl (flip f) x ys
+ reverse-foldr : foldl f x (reverse ys) ≡ foldr (flip f) x ys
+ length-reverse : length (reverse xs) ≡ length xs
+ ```
+
+* Added proofs to `Data.List.All.Properties`
+ ```agda
+ All-universal : Universal P → All P xs
+
+ ¬Any⇒All¬ : ¬ Any P xs → All (¬_ ∘ P) xs
+ All¬⇒¬Any : All (¬_ ∘ P) xs → ¬ Any P xs
+ ¬All⇒Any¬ : Decidable P → ¬ All P xs → Any (¬_ ∘ P) xs
+
+ ++⁺ : All P xs → All P ys → All P (xs ++ ys)
+ ++⁻ˡ : All P (xs ++ ys) → All P xs
+ ++⁻ʳ : All P (xs ++ ys) → All P ys
+ ++⁻ : All P (xs ++ ys) → All P xs × All P ys
+
+ concat⁺ : All (All P) xss → All P (concat xss)
+ concat⁻ : All P (concat xss) → All (All P) xss
+
+ drop⁺ : All P xs → All P (drop n xs)
+ take⁺ : All P xs → All P (take n xs)
+
+ tabulate⁺ : (∀ i → P (f i)) → All P (tabulate f)
+ tabulate⁻ : All P (tabulate f) → (∀ i → P (f i))
+
+ applyUpTo⁺₁ : (∀ {i} → i < n → P (f i)) → All P (applyUpTo f n)
+ applyUpTo⁺₂ : (∀ i → P (f i)) → All P (applyUpTo f n)
+ applyUpTo⁻ : All P (applyUpTo f n) → ∀ {i} → i < n → P (f i)
+ ```
+
+* Added proofs to `Data.List.Any.Properties`
+ ```agda
+ lose∘find : uncurry′ lose (proj₂ (find p)) ≡ p
+ find∘lose : find (lose x∈xs pp) ≡ (x , x∈xs , pp)
+
+ swap : Any (λ x → Any (P x) ys) xs → Any (λ y → Any (flip P y) xs) ys
+ swap-invol : swap (swap any) ≡ any
+
+ ∃∈-Any : (∃ λ x → x ∈ xs × P x) → Any P xs
+
+ Any-⊎⁺ : Any P xs ⊎ Any Q xs → Any (λ x → P x ⊎ Q x) xs
+ Any-⊎⁻ : Any (λ x → P x ⊎ Q x) xs → Any P xs ⊎ Any Q xs
+ Any-×⁺ : Any P xs × Any Q ys → Any (λ x → Any (λ y → P x × Q y) ys) xs
+ Any-×⁻ : Any (λ x → Any (λ y → P x × Q y) ys) xs → Any P xs × Any Q ys
+
+ map⁺ : Any (P ∘ f) xs → Any P (map f xs)
+ map⁻ : Any P (map f xs) → Any (P ∘ f) xs
+
+ ++⁺ˡ : Any P xs → Any P (xs ++ ys)
+ ++⁺ʳ : Any P ys → Any P (xs ++ ys)
+ ++⁻ : Any P (xs ++ ys) → Any P xs ⊎ Any P ys
+
+ concat⁺ : Any (Any P) xss → Any P (concat xss)
+ concat⁻ : Any P (concat xss) → Any (Any P) xss
+
+ applyUpTo⁺ : P (f i) → i < n → Any P (applyUpTo f n)
+ applyUpTo⁻ : Any P (applyUpTo f n) → ∃ λ i → i < n × P (f i)
+
+ tabulate⁺ : P (f i) → Any P (tabulate f)
+ tabulate⁻ : Any P (tabulate f) → ∃ λ i → P (f i)
+
+ map-with-∈⁺ : (∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)) → Any P (map-with-∈ xs f)
+ map-with-∈⁻ : Any P (map-with-∈ xs f) → ∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)
+
+ return⁺ : P x → Any P (return x)
+ return⁻ : Any P (return x) → P x
+ ```
+
+* Added proofs to `Data.List.Any.Membership.Properties`
+ ```agda
+ ∈-map⁺ : x ∈ xs → f x ∈ map f xs
+ ∈-map⁻ : y ∈ map f xs → ∃ λ x → x ∈ xs × y ≈ f x
+ ```
+
+* Added proofs to `Data.List.Any.Membership.Propositional.Properties`
+ ```agda
+ ∈-map⁺ : x ∈ xs → f x ∈ map f xs
+ ∈-map⁻ : y ∈ map f xs → ∃ λ x → x ∈ xs × y ≈ f x
+ ```
+
+* Added proofs to `Data.Maybe`:
+ ```agda
+ Eq-refl : Reflexive _≈_ → Reflexive (Eq _≈_)
+ Eq-sym : Symmetric _≈_ → Symmetric (Eq _≈_)
+ Eq-trans : Transitive _≈_ → Transitive (Eq _≈_)
+ Eq-dec : Decidable _≈_ → Decidable (Eq _≈_)
+ Eq-isEquivalence : IsEquivalence _≈_ → IsEquivalence (Eq _≈_)
+ Eq-isDecEquivalence : IsDecEquivalence _≈_ → IsDecEquivalence (Eq _≈_)
+ ```
+
+* Added exponentiation operator `_^_` to `Data.Nat.Base`
+
+* Added proofs to `Data.Nat.Properties`:
+ ```agda
+ suc-injective : suc m ≡ suc n → m ≡ n
+ ≡-isDecEquivalence : IsDecEquivalence (_≡_ {A = ℕ})
+ ≡-decSetoid : DecSetoid _ _
+
+ ≤-reflexive : _≡_ ⇒ _≤_
+ ≤-refl : Reflexive _≤_
+ ≤-trans : Antisymmetric _≡_ _≤_
+ ≤-antisymmetric : Transitive _≤_
+ ≤-total : Total _≤_
+ ≤-isPreorder : IsPreorder _≡_ _≤_
+ ≤-isPartialOrder : IsPartialOrder _≡_ _≤_
+ ≤-isTotalOrder : IsTotalOrder _≡_ _≤_
+ ≤-isDecTotalOrder : IsDecTotalOrder _≡_ _≤_
+
+ _<?_ : Decidable _<_
+ <-irrefl : Irreflexive _≡_ _<_
+ <-asym : Asymmetric _<_
+ <-transʳ : Trans _≤_ _<_ _<_
+ <-transˡ : Trans _<_ _≤_ _<_
+ <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
+ <⇒≤ : _<_ ⇒ _≤_
+ <⇒≢ : _<_ ⇒ _≢_
+ <⇒≱ : _<_ ⇒ _≱_
+ <⇒≯ : _<_ ⇒ _≯_
+ ≰⇒≮ : _≰_ ⇒ _≮_
+ ≰⇒≥ : _≰_ ⇒ _≥_
+ ≮⇒≥ : _≮_ ⇒ _≥_
+ ≤+≢⇒< : m ≤ n → m ≢ n → m < n
+
+ +-identityˡ : LeftIdentity 0 _+_
+ +-identity : Identity 0 _+_
+ +-cancelʳ-≡ : RightCancellative _≡_ _+_
+ +-cancel-≡ : Cancellative _≡_ _+_
+ +-cancelʳ-≤ : RightCancellative _≤_ _+_
+ +-cancel-≤ : Cancellative _≤_ _+_
+ +-isSemigroup : IsSemigroup _≡_ _+_
+ +-monoˡ-< : _+_ Preserves₂ _<_ ⟶ _≤_ ⟶ _<_
+ +-monoʳ-< : _+_ Preserves₂ _≤_ ⟶ _<_ ⟶ _<_
+ +-mono-< : _+_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_
+ m+n≤o⇒m≤o : m + n ≤ o → m ≤ o
+ m+n≤o⇒n≤o : m + n ≤ o → n ≤ o
+ m+n≮n : m + n ≮ n
+
+ *-zeroˡ : LeftZero 0 _*_
+ *-zero : Zero 0 _*_
+ *-identityˡ : LeftIdentity 1 _*_
+ *-identityʳ : RightIdentity 1 _*_
+ *-identity : Identity 1 _*_
+ *-distribˡ-+ : _*_ DistributesOverˡ _+_
+ *-distrib-+ : _*_ DistributesOver _+_
+ *-isSemigroup : IsSemigroup _≡_ _*_
+ *-mono-< : _*_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_
+ *-monoˡ-< : (_* suc n) Preserves _<_ ⟶ _<_
+ *-monoʳ-< : (suc n *_) Preserves _<_ ⟶ _<_
+ *-cancelˡ-≡ : suc k * i ≡ suc k * j → i ≡ j
+
+ ^-distribˡ-+-* : m ^ (n + p) ≡ m ^ n * m ^ p
+ i^j≡0⇒i≡0 : i ^ j ≡ 0 → i ≡ 0
+ i^j≡1⇒j≡0∨i≡1 : i ^ j ≡ 1 → j ≡ 0 ⊎ i ≡ 1
+
+ ⊔-assoc : Associative _⊔_
+ ⊔-comm : Commutative _⊔_
+ ⊔-idem : Idempotent _⊔_
+ ⊔-identityˡ : LeftIdentity 0 _⊔_
+ ⊔-identityʳ : RightIdentity 0 _⊔_
+ ⊔-identity : Identity 0 _⊔_
+ ⊓-assoc : Associative _⊓_
+ ⊓-comm : Commutative _⊓_
+ ⊓-idem : Idempotent _⊓_
+ ⊓-zeroˡ : LeftZero 0 _⊓_
+ ⊓-zeroʳ : RightZero 0 _⊓_
+ ⊓-zero : Zero 0 _⊓_
+ ⊓-distribʳ-⊔ : _⊓_ DistributesOverʳ _⊔_
+ ⊓-distribˡ-⊔ : _⊓_ DistributesOverˡ _⊔_
+ ⊔-abs-⊓ : _⊔_ Absorbs _⊓_
+ ⊓-abs-⊔ : _⊓_ Absorbs _⊔_
+ m⊓n≤n : m ⊓ n ≤ n
+ m≤m⊔n : m ≤ m ⊔ n
+ m⊔n≤m+n : m ⊔ n ≤ m + n
+ m⊓n≤m+n : m ⊓ n ≤ m + n
+ m⊓n≤m⊔n : m ⊔ n ≤ m ⊔ n
+ ⊔-mono-≤ : _⊔_ Preserves₂ _≤_ ⟶ _≤_ ⟶ _≤_
+ ⊔-mono-< : _⊔_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_
+ ⊓-mono-≤ : _⊓_ Preserves₂ _≤_ ⟶ _≤_ ⟶ _≤_
+ ⊓-mono-< : _⊓_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_
+ +-distribˡ-⊔ : _+_ DistributesOverˡ _⊔_
+ +-distribʳ-⊔ : _+_ DistributesOverʳ _⊔_
+ +-distrib-⊔ : _+_ DistributesOver _⊔_
+ +-distribˡ-⊓ : _+_ DistributesOverˡ _⊓_
+ +-distribʳ-⊓ : _+_ DistributesOverʳ _⊓_
+ +-distrib-⊓ : _+_ DistributesOver _⊓_
+ ⊔-isSemigroup : IsSemigroup _≡_ _⊔_
+ ⊓-isSemigroup : IsSemigroup _≡_ _⊓_
+ ⊓-⊔-isLattice : IsLattice _≡_ _⊓_ _⊔_
+
+ ∸-distribʳ-⊔ : _∸_ DistributesOverʳ _⊔_
+ ∸-distribʳ-⊓ : _∸_ DistributesOverʳ _⊓_
+ +-∸-comm : o ≤ m → (m + n) ∸ o ≡ (m ∸ o) + n
+ ```
+
+* Added decidability relation to `Data.Nat.GCD`
+ ```agda
+ gcd? : (m n d : ℕ) → Dec (GCD m n d)
+ ```
+
+* Added "not-divisible-by" relation to `Data.Nat.Divisibility`
+ ```agda
+ m ∤ n = ¬ (m ∣ n)
+ ```
+
+* Added proofs to `Data.Nat.Divisibility`
+ ```agda
+ ∣-reflexive : _≡_ ⇒ _∣_
+ ∣-refl : Reflexive _∣_
+ ∣-trans : Transitive _∣_
+ ∣-antisym : Antisymmetric _≡_ _∣_
+ ∣-isPreorder : IsPreorder _≡_ _∣_
+ ∣-isPartialOrder : IsPartialOrder _≡_ _∣_
+
+ n∣n : n ∣ n
+ ∣m∸n∣n⇒∣m : n ≤ m → i ∣ m ∸ n → i ∣ n → i ∣ m
+ ```
+
+* Added proofs to `Data.Nat.GeneralisedArithmetic`:
+ ```agda
+ fold-+ : fold z s (m + n) ≡ fold (fold z s n) s m
+ fold-k : fold k (s ∘′_) m z ≡ fold (k z) s m
+ fold-* : fold z s (m * n) ≡ fold z (fold id (s ∘_) n) m
+ fold-pull : fold p s m ≡ g (fold z s m) p
+
+ id-is-fold : fold zero suc m ≡ m
+ +-is-fold : fold n suc m ≡ m + n
+ *-is-fold : fold zero (n +_) m ≡ m * n
+ ^-is-fold : fold 1 (m *_) n ≡ m ^ n
+ *+-is-fold : fold p (n +_) m ≡ m * n + p
+ ^*-is-fold : fold p (m *_) n ≡ m ^ n * p
+ ```
+
+* Added syntax for existential quantifiers in `Data.Product`:
+ ```agda
+ ∃-syntax (λ x → B) = ∃[ x ] B
+ ∄-syntax (λ x → B) = ∄[ x ] B
+ ```
+
+* A new module `Data.Rational.Properties` has been added, containing proofs:
+ ```agda
+ ≤-reflexive : _≡_ ⇒ _≤_
+ ≤-refl : Reflexive _≤_
+ ≤-trans : Transitive _≤_
+ ≤-antisym : Antisymmetric _≡_ _≤_
+ ≤-total : Total _≤_
+
+ ≤-isPreorder : IsPreorder _≡_ _≤_
+ ≤-isPartialOrder : IsPartialOrder _≡_ _≤_
+ ≤-isTotalOrder : IsTotalOrder _≡_ _≤_
+ ≤-isDecTotalOrder : IsDecTotalOrder _≡_ _≤_
+ ```
+
+* Added proofs to `Data.Sign.Properties`:
+ ```agda
+ opposite-cong : opposite s ≡ opposite t → s ≡ t
+
+ *-identityˡ : LeftIdentity + _*_
+ *-identityʳ : RightIdentity + _*_
+ *-identity : Identity + _*_
+ *-comm : Commutative _*_
+ *-assoc : Associative _*_
+ cancel-*-left : LeftCancellative _*_
+ *-cancellative : Cancellative _*_
+ s*s≡+ : s * s ≡ +
+ ```
+
+* Added definitions to `Data.Sum`:
+ ```agda
+ From-inj₁ : ∀ {a b} {A : Set a} {B : Set b} → A ⊎ B → Set a
+ from-inj₁ : ∀ {a b} {A : Set a} {B : Set b} (x : A ⊎ B) → From-inj₁ x
+ From-inj₂ : ∀ {a b} {A : Set a} {B : Set b} → A ⊎ B → Set b
+ from-inj₂ : ∀ {a b} {A : Set a} {B : Set b} (x : A ⊎ B) → From-inj₂ x
+ ```
+
+* Added a functor encapsulating `map` in `Data.Vec`:
+ ```agda
+ functor = record { _<$>_ = map}
+ ```
+
+* Added proofs to `Data.Vec.Equality`
+ ```agda
+ to-≅ : xs ≈ ys → xs ≅ ys
+ xs++[]≈xs : xs ++ [] ≈ xs
+ xs++[]≅xs : xs ++ [] ≅ xs
+ ```
+
+* Added proofs to `Data.Vec.Properties`
+ ```agda
+ lookup-map : lookup i (map f xs) ≡ f (lookup i xs)
+ lookup-functor-morphism : Morphism functor IdentityFunctor
+ map-replicate : map f (replicate x) ≡ replicate (f x)
+
+ ⊛-is-zipWith : fs ⊛ xs ≡ zipWith _$_ fs xs
+ map-is-⊛ : map f xs ≡ replicate f ⊛ xs
+ zipWith-is-⊛ : zipWith f xs ys ≡ replicate f ⊛ xs ⊛ ys
+
+ zipWith-replicate₁ : zipWith _⊕_ (replicate x) ys ≡ map (x ⊕_) ys
+ zipWith-replicate₂ : zipWith _⊕_ xs (replicate y) ≡ map (_⊕ y) xs
+ zipWith-map₁ : zipWith _⊕_ (map f xs) ys ≡ zipWith (λ x y → f x ⊕ y) xs ys
+ zipWith-map₂ : zipWith _⊕_ xs (map f ys) ≡ zipWith (λ x y → x ⊕ f y) xs ys
+ ```
+
+* Added proofs to `Data.Vec.All.Properties`
+ ```agda
+ All-++⁺ : All P xs → All P ys → All P (xs ++ ys)
+ All-++ˡ⁻ : All P (xs ++ ys) → All P xs
+ All-++ʳ⁻ : All P (xs ++ ys) → All P ys
+ All-++⁻ : All P (xs ++ ys) → All P xs × All P ys
+
+ All₂-++⁺ : All₂ _~_ ws xs → All₂ _~_ ys zs → All₂ _~_ (ws ++ ys) (xs ++ zs)
+ All₂-++ˡ⁻ : All₂ _~_ (ws ++ ys) (xs ++ zs) → All₂ _~_ ws xs
+ All₂-++ʳ⁻ : All₂ _~_ (ws ++ ys) (xs ++ zs) → All₂ _~_ ys zs
+ All₂-++⁻ : All₂ _~_ (ws ++ ys) (xs ++ zs) → All₂ _~_ ws xs × All₂ _~_ ys zs
+
+ All-concat⁺ : All (All P) xss → All P (concat xss)
+ All-concat⁻ : All P (concat xss) → All (All P) xss
+
+ All₂-concat⁺ : All₂ (All₂ _~_) xss yss → All₂ _~_ (concat xss) (concat yss)
+ All₂-concat⁻ : All₂ _~_ (concat xss) (concat yss) → All₂ (All₂ _~_) xss yss
+ ```
+
+* Added non-dependant versions of the application combinators in `Function` for use
+ cases where the most general one leads to unsolved meta variables:
+ ```agda
+ _$′_ : (A → B) → (A → B)
+ _$!′_ : (A → B) → (A → B)
+ ```
+
+* Added proofs to `Relation.Binary.Consequences`
+ ```agda
+ P-resp⟶¬P-resp : Symmetric _≈_ → P Respects _≈_ → (¬_ ∘ P) Respects _≈_
+ ```
+
+* Added conversion lemmas to `Relation.Binary.HeterogeneousEquality`
+ ```agda
+ ≅-to-type-≡ : {x : A} {y : B} → x ≅ y → A ≡ B
+ ≅-to-subst-≡ : (p : x ≅ y) → subst (λ x → x) (≅-to-type-≡ p) x ≡ y
+ ```
+
Version 0.13
============
diff --git a/GNUmakefile b/GNUmakefile
index 2916a1b..4ba638d 100644
--- a/GNUmakefile
+++ b/GNUmakefile
@@ -1,7 +1,12 @@
AGDA=agda
+# Before running `make test` the `fix-agda-whitespace` program should
+# be installed:
+#
+# cd agda-development-version-path/src/fix-agda-whitespace
+# cabal install
test: Everything.agda
- fix-agda-whitespace --check
+ cabal exec -- fix-agda-whitespace --check
$(AGDA) -i. -isrc README.agda
setup: Everything.agda
@@ -9,7 +14,7 @@ setup: Everything.agda
.PHONY: Everything.agda
Everything.agda:
cabal clean && cabal install
- GenerateEverything
+ cabal exec -- GenerateEverything
.PHONY: listings
listings: Everything.agda
diff --git a/HACKING b/HACKING
deleted file mode 100644
index 1f9edcd..0000000
--- a/HACKING
+++ /dev/null
@@ -1,31 +0,0 @@
-
-* Where to commit changes:
-
- CURRENT_AGDA = current released Agda version, e.g. 2.4.2.5
- AGDA_MAINT = Agda maintenance version, e.g. 2.4.2.6
-
- A. Your change is independent of Agda
-
- 1. Push your commit in the CURRENT_AGDA branch
- 2. Merge the CURRENT_AGDA branch into the AGDA_MAINT branch
- 3. Merge the AGDA_MAINT branch into the master branch
-
- B. Your change is due to a change in the AGDA_MAINT version of Agda
-
- 1. Push your commit in the AGDA_MAINT branch
- 2. Merge the AGDA_MAINT branch into the master branch
-
- C. Your change is due to a change in the master version of Agda
-
- 1. Push your commit in the master branch
-
- This scheme should guarantee that:
-
- a. the stdlib CURRENT_AGDA branch always builds with the current
- released Agda version,
-
- b. the stdlib AGDA_MAINT branch always build with the Agda maint
- branch and
-
- c. the stdlib master branch always builds with the Agda master
- branch.
diff --git a/HACKING.md b/HACKING.md
new file mode 100644
index 0000000..4b6841a
--- /dev/null
+++ b/HACKING.md
@@ -0,0 +1,45 @@
+Testing and documenting your changes
+------------------------------------
+
+When you implement a new feature of fix a bug:
+
+1. Document it in `CHANGELOG.md`.
+
+2. Test your changes by running
+
+ ```
+ make clean
+ make test
+ ```
+
+Where to commit changes
+-----------------------
+
+ CURRENT_AGDA = current released Agda version, e.g. 2.4.2.5
+ AGDA_MAINT = Agda maintenance version, e.g. 2.4.2.6
+
+A. Your change is independent of Agda
+
+ 1. Push your commit in the `CURRENT_AGDA` branch
+ 2. Merge the `CURRENT_AGDA` branch into the `AGDA_MAINT` branch
+ 3. Merge the `AGDA_MAINT` branch into the master branch
+
+B. Your change is due to a change in the `AGDA_MAINT` version of Agda
+
+ 1. Push your commit in the `AGDA_MAINT` branch
+ 2. Merge the `AGDA_MAINT` branch into the master branch
+
+C. Your change is due to a change in the master version of Agda
+
+ 1. Push your commit in the master branch
+
+This scheme should guarantee that:
+
+ a. the stdlib `CURRENT_AGDA` branch always builds with the current
+ released Agda version,
+
+ b. the stdlib `AGDA_MAINT` branch always build with the Agda maint
+ branch and
+
+ c. the stdlib master branch always builds with the Agda master
+ branch.
diff --git a/LICENCE b/LICENCE
index 87bbb49..6b3edac 100644
--- a/LICENCE
+++ b/LICENCE
@@ -1,11 +1,12 @@
-Copyright (c) 2007-2016 Nils Anders Danielsson, Ulf Norell, Shin-Cheng
+Copyright (c) 2007-2017 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, Guilhem Moulin, Noriyuki OHKAWA, Evgeny Kotelnikov,
-James Chapman, Pepijn Kokke and some anonymous contributors.
+James Chapman, Pepijn Kokke, Matthew Daggitt and some anonymous
+contributors.
Permission is hereby granted, free of charge, to any person obtaining a
copy of this software and associated documentation files (the
diff --git a/README.agda b/README.agda
index 260ed75..b3622b8 100644
--- a/README.agda
+++ b/README.agda
@@ -1,20 +1,20 @@
module README where
------------------------------------------------------------------------
--- The Agda standard library, version 0.13
+-- The Agda standard library, version 0.14
--
-- Author: Nils Anders Danielsson, with contributions from Andreas
-- Abel, Stevan Andjelkovic, Jean-Philippe Bernardy, Peter Berry,
-- Joachim Breitner, Samuel Bronson, Daniel Brown, James Chapman,
-- Liang-Ting Chen, Matthew Daggitt, Dominique Devriese, Dan Doel,
-- Érdi Gergő, Helmut Grohne, Simon Foster, Liyang Hu, Patrik Jansson,
--- Alan Jeffrey, Pepijn Kokke, Evgeny Kotelnikov, Eric Mertens, Darin
--- Morrison, Guilhem Moulin, Shin-Cheng Mu, Ulf Norell, Noriyuki
--- OHKAWA, Nicolas Pouillard, Andrés Sicard-Ramírez, Noam Zeilberger
--- and some anonymous contributors.
+-- Alan Jeffrey, Pepijn Kokke, Evgeny Kotelnikov, Sergei Meshveliani
+-- Eric Mertens, Darin Morrison, Guilhem Moulin, Shin-Cheng Mu,
+-- Ulf Norell, Noriyuki OHKAWA, Nicolas Pouillard, Andrés Sicard-Ramírez,
+-- Noam Zeilberger and some anonymous contributors.
-- ----------------------------------------------------------------------
--- This version of the library has been tested using Agda 2.5.2.
+-- This version of the library has been tested using Agda 2.5.3.
-- Note that no guarantees are currently made about forwards or
-- backwards compatibility, the library is still at an experimental
@@ -61,8 +61,6 @@ module README where
-- well-founded induction).
-- • IO
-- Input/output-related functions.
--- • Irrelevance
--- Definitions related to (proscriptive) irrelevance.
-- • Level
-- Universe levels.
-- • Record
@@ -74,6 +72,8 @@ module README where
-- binary relations).
-- • Size
-- Sizes used by the sized types mechanism.
+-- • Strict
+-- Provides access to the builtins relating to strictness.
-- • Universe
-- A definition of universes.
diff --git a/README/AVL.agda b/README/AVL.agda
index 65c1f29..bc96d06 100644
--- a/README/AVL.agda
+++ b/README/AVL.agda
@@ -17,13 +17,11 @@ import Data.AVL
-- total order, and values, which are indexed by keys. Let us use
-- natural numbers as keys and vectors of strings as values.
-import Data.Nat.Properties as ℕ
+open import Data.Nat.Properties using (<-isStrictTotalOrder)
open import Data.String using (String)
open import Data.Vec using (Vec; _∷_; [])
-open import Relation.Binary using (module StrictTotalOrder)
-open Data.AVL (Vec String)
- (StrictTotalOrder.isStrictTotalOrder ℕ.strictTotalOrder)
+open Data.AVL (Vec String) (<-isStrictTotalOrder)
------------------------------------------------------------------------
-- Construction of trees
diff --git a/README/Nat.agda b/README/Nat.agda
index 4a2e978..b6afd17 100644
--- a/README/Nat.agda
+++ b/README/Nat.agda
@@ -24,27 +24,22 @@ ex₂ : 3 + 5 ≡ 2 * 4
ex₂ = refl
-- Data.Nat.Properties contains a number of properties about natural
--- numbers. Algebra defines what a commutative semiring is, among
--- other things.
+-- numbers.
-open import Algebra
import Data.Nat.Properties as Nat
-private
- module CS = CommutativeSemiring Nat.commutativeSemiring
ex₃ : ∀ m n → m * n ≡ n * m
-ex₃ m n = CS.*-comm m n
+ex₃ m n = Nat.*-comm m n
-- The module ≡-Reasoning in Relation.Binary.PropositionalEquality
-- provides some combinators for equational reasoning.
open ≡-Reasoning
-open import Data.Product
ex₄ : ∀ m n → m * (n + 0) ≡ n * m
ex₄ m n = begin
- m * (n + 0) ≡⟨ cong (_*_ m) (proj₂ CS.+-identity n) ⟩
- m * n ≡⟨ CS.*-comm m n ⟩
+ m * (n + 0) ≡⟨ cong (_*_ m) (Nat.+-identityʳ n) ⟩
+ m * n ≡⟨ Nat.*-comm m n ⟩
n * m ∎
-- The module SemiringSolver in Data.Nat.Properties contains a solver
diff --git a/doc/release-notes/future.txt b/doc/release-notes/future.txt
index fb3cf36..b70180c 100644
--- a/doc/release-notes/future.txt
+++ b/doc/release-notes/future.txt
@@ -28,3 +28,5 @@ future release. Don't remove this message please!
** Added BUILTIN binding in Data/Unit.agda.
** Various changes in Reflection.agda
+
+ ** Use COMPILE and FOREIGN compiler pragmas instead of old deprecated ones.
diff --git a/lib.cabal b/lib.cabal
index 949aefe..b000b82 100644
--- a/lib.cabal
+++ b/lib.cabal
@@ -1,23 +1,25 @@
name: lib
-version: 0.13
+version: 0.14
cabal-version: >= 1.10
build-type: Simple
description: Helper programs.
license: MIT
-tested-with: GHC == 7.6.3
- GHC == 7.8.4
- GHC == 7.10.3
- GHC == 8.0.1
+tested-with: GHC == 7.8.4
+ GHC == 7.10.3
+ GHC == 8.0.2
+ GHC == 8.2.1
executable GenerateEverything
hs-source-dirs: .
main-is: GenerateEverything.hs
- build-depends: base >= 4.6.0.1 && < 4.10
+ default-language: Haskell2010
+ build-depends: base >= 4.7.0.2 && < 4.11
, filemanip >= 0.3.6.2 && < 0.4
- , filepath >= 1.3.0.1 && < 1.5
+ , filepath >= 1.3.0.2 && < 1.5
executable AllNonAsciiChars
hs-source-dirs: .
main-is: AllNonAsciiChars.hs
- build-depends: base >= 4.6.0.1 && < 4.10
+ default-language: Haskell2010
+ build-depends: base >= 4.7.0.2 && < 4.11
, filemanip >= 0.3.6.2 && < 0.4
diff --git a/notes/stdlib-releases.txt b/notes/stdlib-releases.txt
index 29004dc..4332cce 100644
--- a/notes/stdlib-releases.txt
+++ b/notes/stdlib-releases.txt
@@ -30,7 +30,7 @@ procedure can be followed:
* Tag version X.Y (do not forget to record the changes above first):
VERSION=X.Y
- git tag -a v$VERSION -m "Agda standard library v$VERSION"
+ git tag -a v$VERSION -m "Agda standard library version $VERSION"
* Removed release-specific information from README.agda.
@@ -40,6 +40,12 @@ procedure can be followed:
git push --follow-tags
+* Update submodule commit for the stable library in Agda:
+
+ cd agda
+ make fast-forward-std-lib
+ record-the-changes-and-push
+
* Update the Agda wiki:
** The standard library page.
diff --git a/src/Algebra/CommutativeMonoidSolver.agda b/src/Algebra/CommutativeMonoidSolver.agda
index b0b3a22..e593f5c 100644
--- a/src/Algebra/CommutativeMonoidSolver.agda
+++ b/src/Algebra/CommutativeMonoidSolver.agda
@@ -12,6 +12,7 @@ open import Data.Fin using (Fin; zero; suc)
open import Data.Maybe as Maybe
using (Maybe; decToMaybe; From-just; from-just)
open import Data.Nat.Base as ℕ using (ℕ; zero; suc; _+_)
+open import Data.Nat.GeneralisedArithmetic using (fold)
open import Data.Product using (_×_; proj₁; proj₂; uncurry)
open import Data.Vec using (Vec; []; _∷_; lookup; replicate)
@@ -69,7 +70,7 @@ Normal n = Vec ℕ n
⟦_⟧⇓ : ∀ {n} → Normal n → Env n → Carrier
⟦ [] ⟧⇓ _ = ε
-⟦ n ∷ v ⟧⇓ (a ∷ ρ) = ℕ.fold (⟦ v ⟧⇓ ρ) (λ b → a ∙ b) n
+⟦ n ∷ v ⟧⇓ (a ∷ ρ) = fold (⟦ v ⟧⇓ ρ) (λ b → a ∙ b) n
------------------------------------------------------------------------
-- Constructions on normal forms
@@ -123,7 +124,7 @@ comp-correct (l ∷ v) (m ∷ w) (a ∷ ρ) = lemma l m (comp-correct v w ρ)
(b ∙ a) ∙ c ≈⟨ assoc _ _ _ ⟩
b ∙ (a ∙ c) ∎
lemma : ∀ l m {d b c} (p : d ≈ b ∙ c) →
- ℕ.fold d (_∙_ a) (l + m) ≈ ℕ.fold b (_∙_ a) l ∙ ℕ.fold c (_∙_ a) m
+ fold d (a ∙_) (l + m) ≈ fold b (a ∙_) l ∙ fold c (a ∙_) m
lemma zero zero p = p
lemma zero (suc m) p = trans (∙-cong refl (lemma zero m p)) (flip12 _ _ _)
lemma (suc l) m p = trans (∙-cong refl (lemma l m p)) (sym (assoc a _ _))
diff --git a/src/Algebra/FunctionProperties.agda b/src/Algebra/FunctionProperties.agda
index 8abf6be..c5e50f5 100644
--- a/src/Algebra/FunctionProperties.agda
+++ b/src/Algebra/FunctionProperties.agda
@@ -92,6 +92,15 @@ Absorptive ∙ ∘ = (∙ Absorbs ∘) × (∘ Absorbs ∙)
Involutive : Op₁ A → Set _
Involutive f = ∀ x → f (f x) ≈ x
+LeftCancellative : Op₂ A → Set _
+LeftCancellative _•_ = ∀ x {y z} → (x • y) ≈ (x • z) → y ≈ z
+
+RightCancellative : Op₂ A → Set _
+RightCancellative _•_ = ∀ {x} y z → (y • x) ≈ (z • x) → y ≈ z
+
+Cancellative : Op₂ A → Set _
+Cancellative _•_ = LeftCancellative _•_ × RightCancellative _•_
+
Congruent₁ : Op₁ A → Set _
Congruent₁ f = f Preserves _≈_ ⟶ _≈_
diff --git a/src/Algebra/FunctionProperties/Consequences.agda b/src/Algebra/FunctionProperties/Consequences.agda
new file mode 100644
index 0000000..6f35771
--- /dev/null
+++ b/src/Algebra/FunctionProperties/Consequences.agda
@@ -0,0 +1,113 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Relations between properties of functions, such as associativity and
+-- commutativity
+------------------------------------------------------------------------
+
+open import Relation.Binary using (Setoid)
+
+module Algebra.FunctionProperties.Consequences
+ {a ℓ} (S : Setoid a ℓ) where
+
+open Setoid S
+open import Algebra.FunctionProperties _≈_
+open import Relation.Binary.EqReasoning S
+open import Data.Sum using (inj₁; inj₂)
+
+------------------------------------------------------------------------
+-- Transposing identity elements
+
+comm+idˡ⇒idʳ : ∀ {_•_} → Commutative _•_ →
+ ∀ {e} → LeftIdentity e _•_ → RightIdentity e _•_
+comm+idˡ⇒idʳ {_•_} comm {e} idˡ x = begin
+ x • e ≈⟨ comm x e ⟩
+ e • x ≈⟨ idˡ x ⟩
+ x ∎
+
+comm+idʳ⇒idˡ : ∀ {_•_} → Commutative _•_ →
+ ∀ {e} → RightIdentity e _•_ → LeftIdentity e _•_
+comm+idʳ⇒idˡ {_•_} comm {e} idʳ x = begin
+ e • x ≈⟨ comm e x ⟩
+ x • e ≈⟨ idʳ x ⟩
+ x ∎
+
+------------------------------------------------------------------------
+-- Transposing zero elements
+
+comm+zeˡ⇒zeʳ : ∀ {_•_} → Commutative _•_ →
+ ∀ {e} → LeftZero e _•_ → RightZero e _•_
+comm+zeˡ⇒zeʳ {_•_} comm {e} zeˡ x = begin
+ x • e ≈⟨ comm x e ⟩
+ e • x ≈⟨ zeˡ x ⟩
+ e ∎
+
+comm+zeʳ⇒zeˡ : ∀ {_•_} → Commutative _•_ →
+ ∀ {e} → RightZero e _•_ → LeftZero e _•_
+comm+zeʳ⇒zeˡ {_•_} comm {e} zeʳ x = begin
+ e • x ≈⟨ comm e x ⟩
+ x • e ≈⟨ zeʳ x ⟩
+ e ∎
+
+------------------------------------------------------------------------
+-- Transposing inverse elements
+
+comm+invˡ⇒invʳ : ∀ {e _⁻¹ _•_} → Commutative _•_ →
+ LeftInverse e _⁻¹ _•_ → RightInverse e _⁻¹ _•_
+comm+invˡ⇒invʳ {e} {_⁻¹} {_•_} comm invˡ x = begin
+ x • (x ⁻¹) ≈⟨ comm x (x ⁻¹) ⟩
+ (x ⁻¹) • x ≈⟨ invˡ x ⟩
+ e ∎
+
+comm+invʳ⇒invˡ : ∀ {e _⁻¹ _•_} → Commutative _•_ →
+ RightInverse e _⁻¹ _•_ → LeftInverse e _⁻¹ _•_
+comm+invʳ⇒invˡ {e} {_⁻¹} {_•_} comm invʳ x = begin
+ (x ⁻¹) • x ≈⟨ comm (x ⁻¹) x ⟩
+ x • (x ⁻¹) ≈⟨ invʳ x ⟩
+ e ∎
+
+------------------------------------------------------------------------
+-- Transposing distributivity
+
+comm+distrˡ⇒distrʳ : ∀ {_•_ _◦_} → Congruent₂ _◦_ → Commutative _•_ →
+ _•_ DistributesOverˡ _◦_ → _•_ DistributesOverʳ _◦_
+comm+distrˡ⇒distrʳ {_•_} {_◦_} cong comm distrˡ x y z = begin
+ (y ◦ z) • x ≈⟨ comm (y ◦ z) x ⟩
+ x • (y ◦ z) ≈⟨ distrˡ x y z ⟩
+ (x • y) ◦ (x • z) ≈⟨ cong (comm x y) (comm x z) ⟩
+ (y • x) ◦ (z • x) ∎
+
+comm+distrʳ⇒distrˡ : ∀ {_•_ _◦_} → Congruent₂ _◦_ → Commutative _•_ →
+ _•_ DistributesOverʳ _◦_ → _•_ DistributesOverˡ _◦_
+comm+distrʳ⇒distrˡ {_•_} {_◦_} cong comm distrˡ x y z = begin
+ x • (y ◦ z) ≈⟨ comm x (y ◦ z) ⟩
+ (y ◦ z) • x ≈⟨ distrˡ x y z ⟩
+ (y • x) ◦ (z • x) ≈⟨ cong (comm y x) (comm z x) ⟩
+ (x • y) ◦ (x • z) ∎
+
+------------------------------------------------------------------------
+-- Transposing cancellativity
+
+comm+cancelˡ⇒cancelʳ : ∀ {_•_} → Commutative _•_ →
+ LeftCancellative _•_ → RightCancellative _•_
+comm+cancelˡ⇒cancelʳ {_•_} comm cancelˡ {x} y z eq = cancelˡ x (begin
+ x • y ≈⟨ comm x y ⟩
+ y • x ≈⟨ eq ⟩
+ z • x ≈⟨ comm z x ⟩
+ x • z ∎)
+
+comm+cancelʳ⇒cancelˡ : ∀ {_•_} → Commutative _•_ →
+ RightCancellative _•_ → LeftCancellative _•_
+comm+cancelʳ⇒cancelˡ {_•_} comm cancelʳ x {y} {z} eq = cancelʳ y z (begin
+ y • x ≈⟨ comm y x ⟩
+ x • y ≈⟨ eq ⟩
+ x • z ≈⟨ comm x z ⟩
+ z • x ∎)
+
+------------------------------------------------------------------------
+-- Selectivity implies idempotence
+
+sel⇒idem : ∀ {_•_} → Selective _•_ → Idempotent _•_
+sel⇒idem sel x with sel x x
+... | inj₁ x•x≈x = x•x≈x
+... | inj₂ x•x≈x = x•x≈x
diff --git a/src/Algebra/Properties/BooleanAlgebra.agda b/src/Algebra/Properties/BooleanAlgebra.agda
index b976159..91663cf 100644
--- a/src/Algebra/Properties/BooleanAlgebra.agda
+++ b/src/Algebra/Properties/BooleanAlgebra.agda
@@ -17,8 +17,10 @@ private
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 Algebra.FunctionProperties _≈_
+open import Algebra.FunctionProperties.Consequences
+ record {isEquivalence = isEquivalence}
+open import Relation.Binary.EqReasoning setoid
open import Relation.Binary
open import Function
open import Function.Equality using (_⟨$⟩_)
@@ -28,23 +30,17 @@ open import Data.Product
------------------------------------------------------------------------
-- Some simple generalisations
+∨-complementˡ : LeftInverse ⊤ ¬_ _∨_
+∨-complementˡ = comm+invʳ⇒invˡ ∨-comm ∨-complementʳ
+
∨-complement : Inverse ⊤ ¬_ _∨_
∨-complement = ∨-complementˡ , ∨-complementʳ
- where
- ∨-complementˡ : LeftInverse ⊤ ¬_ _∨_
- ∨-complementˡ x = begin
- ¬ x ∨ x ≈⟨ ∨-comm _ _ ⟩
- x ∨ ¬ x ≈⟨ ∨-complementʳ _ ⟩
- ⊤ ∎
+
+∧-complementˡ : LeftInverse ⊥ ¬_ _∧_
+∧-complementˡ = comm+invʳ⇒invˡ ∧-comm ∧-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
@@ -52,8 +48,8 @@ open import Data.Product
∧-∨-isBooleanAlgebra : IsBooleanAlgebra _≈_ _∧_ _∨_ ¬_ ⊥ ⊤
∧-∨-isBooleanAlgebra = record
{ isDistributiveLattice = ∧-∨-isDistributiveLattice
- ; ∨-complementʳ = proj₂ ∧-complement
- ; ∧-complementʳ = proj₂ ∨-complement
+ ; ∨-complementʳ = ∧-complementʳ
+ ; ∧-complementʳ = ∨-complementʳ
; ¬-cong = ¬-cong
}
@@ -67,61 +63,106 @@ open import Data.Product
}
------------------------------------------------------------------------
--- (∨, ∧, ⊥, ⊤) is a commutative semiring
+-- (∨, ∧, ⊥, ⊤) and (∧, ∨, ⊤, ⊥) are commutative semirings
+
+∧-identityʳ : RightIdentity ⊤ _∧_
+∧-identityʳ x = begin
+ x ∧ ⊤ ≈⟨ refl ⟨ ∧-cong ⟩ sym (∨-complementʳ _) ⟩
+ x ∧ (x ∨ ¬ x) ≈⟨ proj₂ absorptive _ _ ⟩
+ x ∎
+
+∧-identityˡ : LeftIdentity ⊤ _∧_
+∧-identityˡ = comm+idʳ⇒idˡ ∧-comm ∧-identityʳ
+
+∧-identity : Identity ⊤ _∧_
+∧-identity = ∧-identityˡ , ∧-identityʳ
+
+∨-identityʳ : RightIdentity ⊥ _∨_
+∨-identityʳ x = begin
+ x ∨ ⊥ ≈⟨ refl ⟨ ∨-cong ⟩ sym (∧-complementʳ _) ⟩
+ x ∨ x ∧ ¬ x ≈⟨ proj₁ absorptive _ _ ⟩
+ x ∎
+
+∨-identityˡ : LeftIdentity ⊥ _∨_
+∨-identityˡ = comm+idʳ⇒idˡ ∨-comm ∨-identityʳ
+
+∨-identity : Identity ⊥ _∨_
+∨-identity = ∨-identityˡ , ∨-identityʳ
+
+∧-zeroʳ : RightZero ⊥ _∧_
+∧-zeroʳ x = begin
+ x ∧ ⊥ ≈⟨ refl ⟨ ∧-cong ⟩ sym (∧-complementʳ _) ⟩
+ x ∧ x ∧ ¬ x ≈⟨ sym $ ∧-assoc _ _ _ ⟩
+ (x ∧ x) ∧ ¬ x ≈⟨ ∧-idempotent _ ⟨ ∧-cong ⟩ refl ⟩
+ x ∧ ¬ x ≈⟨ ∧-complementʳ _ ⟩
+ ⊥ ∎
+
+∧-zeroˡ : LeftZero ⊥ _∧_
+∧-zeroˡ = comm+zeʳ⇒zeˡ ∧-comm ∧-zeroʳ
+
+∧-zero : Zero ⊥ _∧_
+∧-zero = ∧-zeroˡ , ∧-zeroʳ
+
+∨-zeroʳ : ∀ x → x ∨ ⊤ ≈ ⊤
+∨-zeroʳ x = begin
+ x ∨ ⊤ ≈⟨ refl ⟨ ∨-cong ⟩ sym (∨-complementʳ _) ⟩
+ x ∨ x ∨ ¬ x ≈⟨ sym $ ∨-assoc _ _ _ ⟩
+ (x ∨ x) ∨ ¬ x ≈⟨ ∨-idempotent _ ⟨ ∨-cong ⟩ refl ⟩
+ x ∨ ¬ x ≈⟨ ∨-complementʳ _ ⟩
+ ⊤ ∎
+
+∨-zeroˡ : LeftZero ⊤ _∨_
+∨-zeroˡ _ = ∨-comm _ _ ⟨ trans ⟩ ∨-zeroʳ _
+
+∨-zero : Zero ⊤ _∨_
+∨-zero = ∨-zeroˡ , ∨-zeroʳ
+
+∨-isSemigroup : IsSemigroup _≈_ _∨_
+∨-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∨-assoc
+ ; ∙-cong = ∨-cong
+ }
-private
+∧-isSemigroup : IsSemigroup _≈_ _∧_
+∧-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∧-assoc
+ ; ∙-cong = ∧-cong
+ }
- ∧-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 _ ⟩
- ⊥ ∎
+∨-⊥-isMonoid : IsMonoid _≈_ _∨_ ⊥
+∨-⊥-isMonoid = record
+ { isSemigroup = ∨-isSemigroup
+ ; identity = ∨-identity
+ }
+
+∧-⊤-isMonoid : IsMonoid _≈_ _∧_ ⊤
+∧-⊤-isMonoid = record
+ { isSemigroup = ∧-isSemigroup
+ ; identity = ∧-identity
+ }
+
+∨-⊥-isCommutativeMonoid : IsCommutativeMonoid _≈_ _∨_ ⊥
+∨-⊥-isCommutativeMonoid = record
+ { isSemigroup = ∨-isSemigroup
+ ; identityˡ = ∨-identityˡ
+ ; comm = ∨-comm
+ }
+
+∧-⊤-isCommutativeMonoid : IsCommutativeMonoid _≈_ _∧_ ⊤
+∧-⊤-isCommutativeMonoid = record
+ { isSemigroup = ∧-isSemigroup
+ ; identityˡ = ∧-identityˡ
+ ; comm = ∧-comm
+ }
∨-∧-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
- }
+ { +-isCommutativeMonoid = ∨-⊥-isCommutativeMonoid
+ ; *-isCommutativeMonoid = ∧-⊤-isCommutativeMonoid
; distribʳ = proj₂ ∧-∨-distrib
- ; zeroˡ = proj₁ ∧-zero
+ ; zeroˡ = ∧-zeroˡ
}
∨-∧-commutativeSemiring : CommutativeSemiring _ _
@@ -133,49 +174,22 @@ private
; 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
- }
+ { +-isCommutativeMonoid = ∧-⊤-isCommutativeMonoid
+ ; *-isCommutativeMonoid = ∨-⊥-isCommutativeMonoid
; distribʳ = proj₂ ∨-∧-distrib
- ; zeroˡ = proj₁ ∨-zero
+ ; zeroˡ = ∨-zeroˡ
}
∧-∨-commutativeSemiring : CommutativeSemiring _ _
-∧-∨-commutativeSemiring =
- record { isCommutativeSemiring = ∧-∨-isCommutativeSemiring }
+∧-∨-commutativeSemiring = record
+ { _+_ = _∧_
+ ; _*_ = _∨_
+ ; 0# = ⊤
+ ; 1# = ⊥
+ ; isCommutativeSemiring = ∧-∨-isCommutativeSemiring
+ }
------------------------------------------------------------------------
-- Some other properties
@@ -187,24 +201,24 @@ private
private
lemma : ∀ x y → x ∧ y ≈ ⊥ → x ∨ y ≈ ⊤ → ¬ x ≈ y
lemma x y x∧y=⊥ x∨y=⊤ = begin
- ¬ x ≈⟨ sym $ proj₂ ∧-identity _ ⟩
+ ¬ x ≈⟨ sym $ ∧-identityʳ _ ⟩
¬ x ∧ ⊤ ≈⟨ refl ⟨ ∧-cong ⟩ sym x∨y=⊤ ⟩
¬ x ∧ (x ∨ y) ≈⟨ proj₁ ∧-∨-distrib _ _ _ ⟩
- ¬ x ∧ x ∨ ¬ x ∧ y ≈⟨ proj₁ ∧-complement _ ⟨ ∨-cong ⟩ refl ⟩
+ ¬ x ∧ x ∨ ¬ x ∧ y ≈⟨ ∧-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 _ ⟩
+ (x ∨ ¬ x) ∧ y ≈⟨ ∨-complementʳ _ ⟨ ∧-cong ⟩ refl ⟩
+ ⊤ ∧ y ≈⟨ ∧-identityˡ _ ⟩
y ∎
¬⊥=⊤ : ¬ ⊥ ≈ ⊤
-¬⊥=⊤ = lemma ⊥ ⊤ (proj₂ ∧-identity _) (proj₂ ∨-zero _)
+¬⊥=⊤ = lemma ⊥ ⊤ (∧-identityʳ _) (∨-zeroʳ _)
¬⊤=⊥ : ¬ ⊤ ≈ ⊥
-¬⊤=⊥ = lemma ⊤ ⊥ (proj₂ ∧-zero _) (proj₂ ∨-identity _)
+¬⊤=⊥ = lemma ⊤ ⊥ (∧-zeroʳ _) (∨-identityʳ _)
¬-involutive : Involutive ¬_
-¬-involutive x = lemma (¬ x) x (proj₁ ∧-complement _) (proj₁ ∨-complement _)
+¬-involutive x = lemma (¬ x) x (∧-complementˡ _) (∨-complementˡ _)
deMorgan₁ : ∀ x y → ¬ (x ∧ y) ≈ ¬ x ∨ ¬ y
deMorgan₁ x y = lemma (x ∧ y) (¬ x ∨ ¬ y) lem₁ lem₂
@@ -213,16 +227,16 @@ deMorgan₁ x y = lemma (x ∧ y) (¬ x ∨ ¬ y) lem₁ lem₂
(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 _ ⟩
+ y ∧ (x ∧ ¬ x) ∨ x ∧ (y ∧ ¬ y) ≈⟨ (refl ⟨ ∧-cong ⟩ ∧-complementʳ _) ⟨ ∨-cong ⟩
+ (refl ⟨ ∧-cong ⟩ ∧-complementʳ _) ⟩
+ (y ∧ ⊥) ∨ (x ∧ ⊥) ≈⟨ ∧-zeroʳ _ ⟨ ∨-cong ⟩ ∧-zeroʳ _ ⟩
+ ⊥ ∨ ⊥ ≈⟨ ∨-identityʳ _ ⟩
⊥ ∎
lem₃ = begin
(x ∧ y) ∨ ¬ x ≈⟨ proj₂ ∨-∧-distrib _ _ _ ⟩
- (x ∨ ¬ x) ∧ (y ∨ ¬ x) ≈⟨ proj₂ ∨-complement _ ⟨ ∧-cong ⟩ refl ⟩
- ⊤ ∧ (y ∨ ¬ x) ≈⟨ proj₁ ∧-identity _ ⟩
+ (x ∨ ¬ x) ∧ (y ∨ ¬ x) ≈⟨ ∨-complementʳ _ ⟨ ∧-cong ⟩ refl ⟩
+ ⊤ ∧ (y ∨ ¬ x) ≈⟨ ∧-identityˡ _ ⟩
y ∨ ¬ x ≈⟨ ∨-comm _ _ ⟩
¬ x ∨ y ∎
@@ -230,8 +244,8 @@ deMorgan₁ x y = lemma (x ∧ y) (¬ x ∨ ¬ y) lem₁ lem₂
(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 _ ⟩
+ ¬ x ∨ (y ∨ ¬ y) ≈⟨ refl ⟨ ∨-cong ⟩ ∨-complementʳ _ ⟩
+ ¬ x ∨ ⊤ ≈⟨ ∨-zeroʳ _ ⟩
⊤ ∎
deMorgan₂ : ∀ x y → ¬ (x ∨ y) ≈ ¬ x ∧ ¬ y
@@ -279,10 +293,24 @@ module XorRing
_⊕_ : 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
+ ⊕-cong : Congruent₂ _⊕_
+ ⊕-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 ∎
+
+ ⊕-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 $ ⊕-def _ _ ⟩
@@ -303,18 +331,10 @@ module XorRing
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 ∧ ¬ x) ∨ (x ∧ ¬ y) ≈⟨ ∧-complementʳ _ ⟨ ∨-cong ⟩ refl ⟩
+ ⊥ ∨ (x ∧ ¬ y) ≈⟨ ∨-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 _ _ ⟩
@@ -329,94 +349,87 @@ module XorRing
¬ (¬ 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 ⟩
+ ⊕-identityˡ : LeftIdentity ⊥ _⊕_
+ ⊕-identityˡ x = begin
+ ⊥ ⊕ x ≈⟨ ⊕-def _ _ ⟩
+ (⊥ ∨ x) ∧ ¬ (⊥ ∧ x) ≈⟨ helper (∨-identityˡ _) (∧-zeroˡ _) ⟩
+ x ∧ ¬ ⊥ ≈⟨ refl ⟨ ∧-cong ⟩ ¬⊥=⊤ ⟩
+ x ∧ ⊤ ≈⟨ ∧-identityʳ _ ⟩
+ x ∎
+
+ ⊕-identityʳ : RightIdentity ⊥ _⊕_
+ ⊕-identityʳ _ = ⊕-comm _ _ ⟨ trans ⟩ ⊕-identityˡ _
+
+ ⊕-identity : Identity ⊥ _⊕_
+ ⊕-identity = ⊕-identityˡ , ⊕-identityʳ
+
+ ⊕-inverseˡ : LeftInverse ⊥ id _⊕_
+ ⊕-inverseˡ x = begin
+ x ⊕ x ≈⟨ ⊕-def _ _ ⟩
+ (x ∨ x) ∧ ¬ (x ∧ x) ≈⟨ helper (∨-idempotent _) (∧-idempotent _) ⟩
+ x ∧ ¬ x ≈⟨ ∧-complementʳ _ ⟩
+ ⊥ ∎
+
+ ⊕-inverseʳ : RightInverse ⊥ id _⊕_
+ ⊕-inverseʳ _ = ⊕-comm _ _ ⟨ trans ⟩ ⊕-inverseˡ _
+
+ ⊕-inverse : Inverse ⊥ id _⊕_
+ ⊕-inverse = ⊕-inverseˡ , ⊕-inverseʳ
+
+ ∧-distribˡ-⊕ : _∧_ DistributesOverˡ _⊕_
+ ∧-distribˡ-⊕ 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 $ ∨-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) ∎
+ ((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 $ ∧-zeroʳ _ ⟩
+ (y ∨ z) ∧ ⊥ ≈⟨ refl ⟨ ∧-cong ⟩ sym (∧-complementʳ _) ⟩
+ (y ∨ z) ∧ (x ∧ ¬ x) ≈⟨ sym $ ∧-assoc _ _ _ ⟩
+ ((y ∨ z) ∧ x) ∧ ¬ x ≈⟨ ∧-comm _ _ ⟨ ∧-cong ⟩ refl ⟩
+ (x ∧ (y ∨ z)) ∧ ¬ x ∎
+
+ ∧-distribʳ-⊕ : _∧_ DistributesOverʳ _⊕_
+ ∧-distribʳ-⊕ = comm+distrˡ⇒distrʳ ⊕-cong ∧-comm ∧-distribˡ-⊕
+
+ ∧-distrib-⊕ : _∧_ DistributesOver _⊕_
+ ∧-distrib-⊕ = ∧-distribˡ-⊕ , ∧-distribʳ-⊕
+
+ private
lemma₂ : ∀ x y u v →
(x ∧ y) ∨ (u ∧ v) ≈
@@ -430,125 +443,129 @@ module XorRing
((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)) ∎
+ ⊕-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 $ ∧-identityˡ _ ⟨ ∧-cong ⟩ ∧-identityʳ _ ⟩
+ (⊤ ∧ (x ∨ ¬ y)) ∧ ((¬ x ∨ y) ∧ ⊤) ≈⟨ sym $
+ (∨-complementˡ _ ⟨ ∧-cong ⟩ ∨-comm _ _)
+ ⟨ ∧-cong ⟩
+ (refl ⟨ ∧-cong ⟩ ∨-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)) ≈⟨ (∨-complementˡ _ ⟨ ∧-cong ⟩ ∨-comm _ _)
+ ⟨ ∧-cong ⟩
+ (refl ⟨ ∧-cong ⟩ ∨-complementˡ _) ⟩
+ (⊤ ∧ (y ∨ ¬ z)) ∧
+ ((¬ y ∨ z) ∧ ⊤) ≈⟨ ∧-identityˡ _ ⟨ ∧-cong ⟩ ∧-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)) ∎
+
+ ⊕-isSemigroup : IsSemigroup _≈_ _⊕_
+ ⊕-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ⊕-assoc
+ ; ∙-cong = ⊕-cong
+ }
- 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) ∎
+ ⊕-⊥-isMonoid : IsMonoid _≈_ _⊕_ ⊥
+ ⊕-⊥-isMonoid = record
+ { isSemigroup = ⊕-isSemigroup
+ ; identity = ⊕-identity
+ }
- 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)) ∎
+ ⊕-⊥-isGroup : IsGroup _≈_ _⊕_ ⊥ id
+ ⊕-⊥-isGroup = record
+ { isMonoid = ⊕-⊥-isMonoid
+ ; inverse = ⊕-inverse
+ ; ⁻¹-cong = id
+ }
+
+ ⊕-⊥-isAbelianGroup : IsAbelianGroup _≈_ _⊕_ ⊥ id
+ ⊕-⊥-isAbelianGroup = record
+ { isGroup = ⊕-⊥-isGroup
+ ; comm = ⊕-comm
+ }
+
+ ⊕-∧-isRing : IsRing _≈_ _⊕_ _∧_ id ⊥ ⊤
+ ⊕-∧-isRing = record
+ { +-isAbelianGroup = ⊕-⊥-isAbelianGroup
+ ; *-isMonoid = ∧-⊤-isMonoid
+ ; distrib = ∧-distrib-⊕
+ }
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-∧-⊕
- }
+ { isRing = ⊕-∧-isRing
; *-comm = ∧-comm
}
diff --git a/src/Algebra/Properties/BooleanAlgebra/Expression.agda b/src/Algebra/Properties/BooleanAlgebra/Expression.agda
index a94397e..f72f579 100644
--- a/src/Algebra/Properties/BooleanAlgebra/Expression.agda
+++ b/src/Algebra/Properties/BooleanAlgebra/Expression.agda
@@ -19,7 +19,7 @@ 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
+open import Data.Product using (_,_; proj₁; proj₂)
import Data.Vec.Properties as VecProp
open import Function
open import Relation.Binary.PropositionalEquality as P using (_≗_)
@@ -75,8 +75,7 @@ module Naturality
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 ρ ⟩
+ op (Vec.lookup x ρ) ≡⟨ P.sym $ VecProp.lookup-map x op ρ ⟩
Vec.lookup x (Vec.map op ρ) ∎
natural (e₁ or e₂) ρ = begin
op (pure₁ _∨_ ⊛₁ ⟦ e₁ ⟧₁ ρ ⊛₁ ⟦ e₂ ⟧₁ ρ) ≡⟨ op-⊛ _ _ ⟩
@@ -108,11 +107,11 @@ lift : ℕ → BooleanAlgebra b b
lift n = record
{ Carrier = Vec Carrier n
; _≈_ = Pointwise _≈_
- ; _∨_ = Vec.zipWith _∨_
- ; _∧_ = Vec.zipWith _∧_
- ; ¬_ = Vec.map ¬_
- ; ⊤ = Vec.replicate ⊤
- ; ⊥ = Vec.replicate ⊥
+ ; _∨_ = zipWith _∨_
+ ; _∧_ = zipWith _∧_
+ ; ¬_ = map ¬_
+ ; ⊤ = pure ⊤
+ ; ⊥ = pure ⊥
; isBooleanAlgebra = record
{ isDistributiveLattice = record
{ isLattice = record
@@ -167,6 +166,9 @@ lift n = record
}
}
where
+ open RawApplicative Vec.applicative
+ using (pure; zipWith) renaming (_<$>_ to map)
+
⟦_⟧Id : ∀ {n} → Expr n → Vec Carrier n → Carrier
⟦_⟧Id = Semantics.⟦_⟧ (RawMonad.rawIApplicative IdentityMonad)
diff --git a/src/Algebra/Properties/DistributiveLattice.agda b/src/Algebra/Properties/DistributiveLattice.agda
index 0c1abcf..5d4f38e 100644
--- a/src/Algebra/Properties/DistributiveLattice.agda
+++ b/src/Algebra/Properties/DistributiveLattice.agda
@@ -16,43 +16,43 @@ private
open module L = Algebra.Properties.Lattice lattice public
hiding (replace-equality)
open import Algebra.Structures
-import Algebra.FunctionProperties as P; open P _≈_
+open import Algebra.FunctionProperties _≈_
open import Relation.Binary
-import Relation.Binary.EqReasoning as EqR; open EqR setoid
+open import Relation.Binary.EqReasoning setoid
open import Function
open import Function.Equality using (_⟨$⟩_)
open import Function.Equivalence using (_⇔_; module Equivalence)
open import Data.Product
+∨-∧-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 ≈⟨ ∨-comm _ _ ⟩
- y ∧ z ∨ x ≈⟨ ∨-∧-distribʳ _ _ _ ⟩
- (y ∨ x) ∧ (z ∨ x) ≈⟨ ∨-comm _ _ ⟨ ∧-cong ⟩ ∨-comm _ _ ⟩
- (x ∨ y) ∧ (x ∨ z) ∎
+
+∧-∨-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 ∎
∧-∨-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.
diff --git a/src/Category/Applicative/Indexed.agda b/src/Category/Applicative/Indexed.agda
index 2012cf0..a96dffd 100644
--- a/src/Category/Applicative/Indexed.agda
+++ b/src/Category/Applicative/Indexed.agda
@@ -9,7 +9,7 @@
module Category.Applicative.Indexed where
-open import Category.Functor
+open import Category.Functor using (RawFunctor)
open import Data.Product
open import Function
open import Level
@@ -61,10 +61,10 @@ record Morphism {i f} {I : Set i} {F₁ F₂ : IFun I f}
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 (f A₁.⊛ x) ≡ (op f A₂.⊛ 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 A₁.<$> x) ≡ (f A₂.<$> 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 ⟩
diff --git a/src/Category/Functor.agda b/src/Category/Functor.agda
index 3fc7676..99a82ee 100644
--- a/src/Category/Functor.agda
+++ b/src/Category/Functor.agda
@@ -11,6 +11,8 @@ module Category.Functor where
open import Function
open import Level
+open import Relation.Binary.PropositionalEquality
+
record RawFunctor {ℓ} (F : Set ℓ → Set ℓ) : Set (suc ℓ) where
infixl 4 _<$>_ _<$_
@@ -19,3 +21,15 @@ record RawFunctor {ℓ} (F : Set ℓ → Set ℓ) : Set (suc ℓ) where
_<$_ : ∀ {A B} → A → F B → F A
x <$ y = const x <$> y
+
+-- A functor morphism from F₁ to F₂ is an operation op such that
+-- op (F₁ f x) ≡ F₂ f (op x)
+
+record Morphism {ℓ} {F₁ F₂ : Set ℓ → Set ℓ}
+ (fun₁ : RawFunctor F₁)
+ (fun₂ : RawFunctor F₂) : Set (suc ℓ) where
+ open RawFunctor
+ field
+ op : ∀{X} → F₁ X → F₂ X
+ op-<$> : ∀{X Y} (f : X → Y) (x : F₁ X) →
+ op (fun₁ ._<$>_ f x) ≡ fun₂ ._<$>_ f (op x)
diff --git a/src/Category/Functor/Identity.agda b/src/Category/Functor/Identity.agda
new file mode 100644
index 0000000..9357eb6
--- /dev/null
+++ b/src/Category/Functor/Identity.agda
@@ -0,0 +1,17 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- The identity functor
+------------------------------------------------------------------------
+
+module Category.Functor.Identity where
+
+open import Category.Functor
+
+Identity : ∀ {f} → Set f → Set f
+Identity A = A
+
+IdentityFunctor : ∀ {f} → RawFunctor (Identity {f})
+IdentityFunctor = record
+ { _<$>_ = λ x → x
+ }
diff --git a/src/Data/Bin.agda b/src/Data/Bin.agda
index 3a05a0e..7807763 100644
--- a/src/Data/Bin.agda
+++ b/src/Data/Bin.agda
@@ -8,43 +8,45 @@ module Data.Bin where
open import Data.Nat as Nat
using (ℕ; zero; z≤n; s≤s) renaming (suc to 1+_)
-open Nat.≤-Reasoning
-import Data.Nat.Properties as NP
-open import Data.Digit
-open import Data.Fin as Fin using (Fin; zero) renaming (suc to 1+_)
+open import Data.Digit using (fromDigits; toDigits; Bit)
+open import Data.Fin as Fin using (Fin; zero)
+ renaming (suc to 1+_)
open import Data.Fin.Properties as FP using (_+′_)
open import Data.List.Base as List hiding (downFrom)
open import Function
-open import Data.Product
-open import Algebra
+open import Data.Product using (uncurry; _,_; _×_)
open import Relation.Binary
-open import Relation.Binary.Consequences
-open import Relation.Binary.PropositionalEquality as PropEq
+open import Relation.Binary.PropositionalEquality
using (_≡_; _≢_; refl; sym)
+open import Relation.Binary.List.StrictLex
open import Relation.Nullary
open import Relation.Nullary.Decidable
-private
- module BitOrd = StrictTotalOrder (FP.strictTotalOrder 2)
------------------------------------------------------------------------
--- The type
+-- Bits
--- A representation of binary natural numbers in which there is
--- exactly one representative for every number.
+pattern 0b = zero
+pattern 1b = 1+ zero
+pattern ⊥b = 1+ 1+ ()
--- The function toℕ below defines the meaning of Bin.
+------------------------------------------------------------------------
+-- The type
-infix 8 _1#
+-- A representation of binary natural numbers in which there is
+-- exactly one representative for every number. The function toℕ below
+-- defines the meaning of Bin.
--- bs stands for the binary number 1<reverse bs>, which is positive.
+-- `bs 1#` stands for the binary number "1<reverse bs>" e.g.
+-- `(0b ∷ []) 1#` represents "10"
+-- `(0b ∷ 1b ∷ 1b ∷ []) 1#` represents "1110"
Bin⁺ : Set
Bin⁺ = List Bit
+infix 8 _1#
+
data Bin : Set where
- -- Zero.
0# : Bin
- -- bs 1# stands for the binary number 1<reverse bs>.
_1# : (bs : Bin⁺) → Bin
------------------------------------------------------------------------
@@ -66,12 +68,12 @@ toℕ = fromDigits ∘ toBits
-- significant one.
fromBits : List Bit → Bin
-fromBits [] = 0#
-fromBits (b ∷ bs) with fromBits bs
-fromBits (b ∷ bs) | bs′ 1# = (b ∷ bs′) 1#
-fromBits (zero ∷ bs) | 0# = 0#
-fromBits ((1+ zero) ∷ bs) | 0# = [] 1#
-fromBits ((1+ 1+ ()) ∷ bs) | _
+fromBits [] = 0#
+fromBits (b ∷ bs) with fromBits bs
+fromBits (b ∷ bs) | bs′ 1# = (b ∷ bs′) 1#
+fromBits (0b ∷ bs) | 0# = 0#
+fromBits (1b ∷ bs) | 0# = [] 1#
+fromBits (⊥b ∷ bs) | _
private
pattern 2+_ n = 1+ 1+ n
@@ -93,113 +95,15 @@ fromℕ : ℕ → Bin
fromℕ n = fromBits $ ntoBits n
------------------------------------------------------------------------
--- (Bin, _≡_, _<_) is a strict total order
+-- Order relation.
-infix 4 _<_
+-- Wrapped so that the parameters can be inferred.
--- Order relation. Wrapped so that the parameters can be inferred.
+infix 4 _<_
data _<_ (b₁ b₂ : Bin) : Set where
less : (lt : (Nat._<_ on toℕ) b₁ b₂) → b₁ < b₂
-private
- <-trans : Transitive _<_
- <-trans (less lt₁) (less lt₂) = less (NP.<-trans lt₁ lt₂)
-
- asym : ∀ {b₁ b₂} → b₁ < b₂ → ¬ b₂ < b₁
- asym {b₁} {b₂} (less lt) (less gt) = tri⟶asym cmp lt gt
- where cmp = StrictTotalOrder.compare NP.strictTotalOrder
-
- irr : ∀ {b₁ b₂} → b₁ < b₂ → b₁ ≢ b₂
- irr lt eq = asym⟶irr (PropEq.resp₂ _<_) sym asym eq lt
-
- irr′ : ∀ {b} → ¬ b < b
- irr′ lt = irr lt refl
-
- ∷∙ : ∀ {b₁ b₂ bs₁ bs₂} →
- bs₁ 1# < bs₂ 1# → (b₁ ∷ bs₁) 1# < (b₂ ∷ bs₂) 1#
- ∷∙ {b₁} {b₂} {bs₁} {bs₂} (less lt) = less (begin
- 1 + (m₁ + n₁ * 2) ≤⟨ ≤-refl {x = 1} +-mono
- (≤-pred (FP.bounded b₁) +-mono ≤-refl) ⟩
- 1 + (1 + n₁ * 2) ≡⟨ refl ⟩
- suc n₁ * 2 ≤⟨ lt *-mono ≤-refl ⟩
- n₂ * 2 ≤⟨ n≤m+n m₂ (n₂ * 2) ⟩
- m₂ + n₂ * 2 ∎
- )
- where
- open Nat; open NP
- open DecTotalOrder decTotalOrder renaming (refl to ≤-refl)
- m₁ = Fin.toℕ b₁; m₂ = Fin.toℕ b₂
- n₁ = toℕ (bs₁ 1#); n₂ = toℕ (bs₂ 1#)
-
- ∙∷ : ∀ {b₁ b₂ bs} →
- Fin._<_ b₁ b₂ → (b₁ ∷ bs) 1# < (b₂ ∷ bs) 1#
- ∙∷ {b₁} {b₂} {bs} lt = less (begin
- 1 + (m₁ + n * 2) ≡⟨ sym (+-assoc 1 m₁ (n * 2)) ⟩
- (1 + m₁) + n * 2 ≤⟨ lt +-mono ≤-refl ⟩
- m₂ + n * 2 ∎)
- where
- open Nat; open NP
- open DecTotalOrder decTotalOrder renaming (refl to ≤-refl)
- open CommutativeSemiring commutativeSemiring using (+-assoc)
- m₁ = Fin.toℕ b₁; m₂ = Fin.toℕ b₂; n = toℕ (bs 1#)
-
- 1<[23] : ∀ {b} → [] 1# < (b ∷ []) 1#
- 1<[23] {b} = less (NP.n≤m+n (Fin.toℕ b) 2)
-
- 1<2+ : ∀ {bs b} → [] 1# < (b ∷ bs) 1#
- 1<2+ {[]} = 1<[23]
- 1<2+ {b ∷ bs} = <-trans 1<[23] (∷∙ {b₁ = b} (1<2+ {bs}))
-
- 0<1 : 0# < [] 1#
- 0<1 = less (s≤s z≤n)
-
- 0<+ : ∀ {bs} → 0# < bs 1#
- 0<+ {[]} = 0<1
- 0<+ {b ∷ bs} = <-trans 0<1 1<2+
-
- compare⁺ : Trichotomous (_≡_ on _1#) (_<_ on _1#)
- compare⁺ [] [] = tri≈ irr′ refl irr′
- compare⁺ [] (b ∷ bs) = tri< 1<2+ (irr 1<2+) (asym 1<2+)
- compare⁺ (b ∷ bs) [] = tri> (asym 1<2+) (irr 1<2+ ∘ sym) 1<2+
- compare⁺ (b₁ ∷ bs₁) (b₂ ∷ bs₂) with compare⁺ bs₁ bs₂
- ... | tri< lt ¬eq ¬gt = tri< (∷∙ lt) (irr (∷∙ lt)) (asym (∷∙ lt))
- ... | tri> ¬lt ¬eq gt = tri> (asym (∷∙ gt)) (irr (∷∙ gt) ∘ sym) (∷∙ gt)
- compare⁺ (b₁ ∷ bs) (b₂ ∷ .bs) | tri≈ ¬lt refl ¬gt with BitOrd.compare b₁ b₂
- compare⁺ (b ∷ bs) (.b ∷ .bs) | tri≈ ¬lt refl ¬gt | tri≈ ¬lt′ refl ¬gt′ =
- tri≈ irr′ refl irr′
- ... | tri< lt′ ¬eq ¬gt′ = tri< (∙∷ lt′) (irr (∙∷ lt′)) (asym (∙∷ lt′))
- ... | tri> ¬lt′ ¬eq gt′ = tri> (asym (∙∷ gt′)) (irr (∙∷ gt′) ∘ sym) (∙∷ gt′)
-
- compare : Trichotomous _≡_ _<_
- compare 0# 0# = tri≈ irr′ refl irr′
- compare 0# (bs₂ 1#) = tri< 0<+ (irr 0<+) (asym 0<+)
- compare (bs₁ 1#) 0# = tri> (asym 0<+) (irr 0<+ ∘ sym) 0<+
- compare (bs₁ 1#) (bs₂ 1#) = compare⁺ bs₁ bs₂
-
-strictTotalOrder : StrictTotalOrder _ _ _
-strictTotalOrder = record
- { Carrier = Bin
- ; _≈_ = _≡_
- ; _<_ = _<_
- ; isStrictTotalOrder = record
- { isEquivalence = PropEq.isEquivalence
- ; trans = <-trans
- ; compare = compare
- }
- }
-
-------------------------------------------------------------------------
--- (Bin, _≡_) is a decidable setoid
-
-decSetoid : DecSetoid _ _
-decSetoid = StrictTotalOrder.decSetoid strictTotalOrder
-
-infix 4 _≟_
-
-_≟_ : Decidable {A = Bin} _≡_
-_≟_ = DecSetoid._≟_ decSetoid
-
------------------------------------------------------------------------
-- Arithmetic
@@ -243,19 +147,19 @@ Carry = Bit
addBits : Carry → Bit → Bit → Carry × Bit
addBits c b₁ b₂ with c +′ (b₁ +′ b₂)
-... | zero = (0b , 0b)
-... | 1+ zero = (0b , 1b)
-... | 1+ 1+ zero = (1b , 0b)
-... | 1+ 1+ 1+ zero = (1b , 1b)
+... | zero = (0b , 0b)
+... | 1+ zero = (0b , 1b)
+... | 1+ 1+ zero = (1b , 0b)
+... | 1+ 1+ 1+ zero = (1b , 1b)
... | 1+ 1+ 1+ 1+ ()
addCarryToBitList : Carry → List Bit → List Bit
-addCarryToBitList zero bs = bs
-addCarryToBitList (1+ zero) [] = 1b ∷ []
-addCarryToBitList (1+ zero) (zero ∷ bs) = 1b ∷ bs
-addCarryToBitList (1+ zero) ((1+ zero) ∷ bs) = 0b ∷ addCarryToBitList 1b bs
-addCarryToBitList (1+ 1+ ()) _
-addCarryToBitList _ ((1+ 1+ ()) ∷ _)
+addCarryToBitList 0b bs = bs
+addCarryToBitList 1b [] = 1b ∷ []
+addCarryToBitList 1b (0b ∷ bs) = 1b ∷ bs
+addCarryToBitList 1b (1b ∷ bs) = 0b ∷ addCarryToBitList 1b bs
+addCarryToBitList ⊥b _
+addCarryToBitList _ (⊥b ∷ _)
addBitLists : Carry → List Bit → List Bit → List Bit
addBitLists c [] bs₂ = addCarryToBitList c bs₂
@@ -276,11 +180,11 @@ _*_ : Bin → Bin → Bin
0# * n = 0#
[] 1# * n = n
-- (b + 2 * bs 1#) * n = b * n + 2 * (bs 1# * n)
-(b ∷ bs) 1# * n with bs 1# * n
-(b ∷ bs) 1# * n | 0# = 0#
-(zero ∷ bs) 1# * n | bs' 1# = (0b ∷ bs') 1#
-((1+ zero) ∷ bs) 1# * n | bs' 1# = n + (0b ∷ bs') 1#
-((1+ 1+ ()) ∷ _) 1# * _ | _
+(b ∷ bs) 1# * n with bs 1# * n
+(b ∷ bs) 1# * n | 0# = 0#
+(0b ∷ bs) 1# * n | bs' 1# = (0b ∷ bs') 1#
+(1b ∷ bs) 1# * n | bs' 1# = n + (0b ∷ bs') 1#
+(⊥b ∷ _) 1# * _ | _
-- Successor.
@@ -295,10 +199,10 @@ suc n = [] 1# + n
-- Predecessor.
pred : Bin⁺ → Bin
-pred [] = 0#
-pred (zero ∷ bs) = pred bs *2+1
-pred ((1+ zero) ∷ bs) = (zero ∷ bs) 1#
-pred ((1+ 1+ ()) ∷ bs)
+pred [] = 0#
+pred (0b ∷ bs) = pred bs *2+1
+pred (1b ∷ bs) = (zero ∷ bs) 1#
+pred (⊥b ∷ bs)
-- downFrom n enumerates all numbers from n - 1 to 0. This function is
-- linear in n. Analysis: fromℕ takes linear time, and the preds used
diff --git a/src/Data/Bin/Properties.agda b/src/Data/Bin/Properties.agda
new file mode 100644
index 0000000..9c95f7b
--- /dev/null
+++ b/src/Data/Bin/Properties.agda
@@ -0,0 +1,141 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties of the binary representation of natural numbers
+------------------------------------------------------------------------
+
+module Data.Bin.Properties where
+
+open import Data.Bin
+open import Data.Digit using (Bit; Expansion)
+import Data.Fin as Fin
+import Data.Fin.Properties as 𝔽ₚ
+open import Data.List.Base using (List; []; _∷_)
+open import Data.List.Properties using (∷-injective)
+open import Data.Nat
+ using (ℕ; zero; z≤n; s≤s; ≤-pred)
+ renaming (suc to 1+_; _+_ to _+ℕ_; _*_ to _*ℕ_; _≤_ to _≤ℕ_)
+import Data.Nat.Properties as ℕₚ
+open import Data.Product using (proj₁; proj₂)
+open import Function using (_∘_)
+open import Relation.Binary
+open import Relation.Binary.Consequences
+open import Relation.Binary.PropositionalEquality
+ using (_≡_; _≢_; refl; sym; isEquivalence; resp₂; decSetoid)
+open import Relation.Nullary using (yes; no)
+
+------------------------------------------------------------------------
+-- (Bin, _≡_) is a decidable setoid
+
+1#-injective : ∀ {as bs} → as 1# ≡ bs 1# → as ≡ bs
+1#-injective refl = refl
+
+infix 4 _≟_ _≟ₑ_
+
+_≟ₑ_ : ∀ {base} → Decidable (_≡_ {A = Expansion base})
+_≟ₑ_ [] [] = yes refl
+_≟ₑ_ [] (_ ∷ _) = no λ()
+_≟ₑ_ (_ ∷ _) [] = no λ()
+_≟ₑ_ (x ∷ xs) (y ∷ ys) with x 𝔽ₚ.≟ y | xs ≟ₑ ys
+... | _ | no xs≢ys = no (xs≢ys ∘ proj₂ ∘ ∷-injective)
+... | no x≢y | _ = no (x≢y ∘ proj₁ ∘ ∷-injective)
+... | yes refl | yes refl = yes refl
+
+_≟_ : Decidable {A = Bin} _≡_
+0# ≟ 0# = yes refl
+0# ≟ bs 1# = no λ()
+as 1# ≟ 0# = no λ()
+as 1# ≟ bs 1# with as ≟ₑ bs
+... | yes refl = yes refl
+... | no as≢bs = no (as≢bs ∘ 1#-injective)
+
+≡-isDecEquivalence : IsDecEquivalence _≡_
+≡-isDecEquivalence = record
+ { isEquivalence = isEquivalence
+ ; _≟_ = _≟_
+ }
+
+≡-decSetoid : DecSetoid _ _
+≡-decSetoid = decSetoid _≟_
+
+------------------------------------------------------------------------
+-- (Bin _≡_ _<_) is a strict total order
+
+<-trans : Transitive _<_
+<-trans (less lt₁) (less lt₂) = less (ℕₚ.<-trans lt₁ lt₂)
+
+<-asym : Asymmetric _<_
+<-asym (less lt) (less gt) = ℕₚ.<-asym lt gt
+
+<-irrefl : Irreflexive _≡_ _<_
+<-irrefl refl (less lt) = ℕₚ.<-irrefl refl lt
+
+∷ʳ-mono-< : ∀ {a b as bs} → as 1# < bs 1# → (a ∷ as) 1# < (b ∷ bs) 1#
+∷ʳ-mono-< {a} {b} {as} {bs} (less lt) = less (begin
+ 1+ (m₁ +ℕ n₁ *ℕ 2) ≤⟨ s≤s (ℕₚ.+-mono-≤ (≤-pred (𝔽ₚ.bounded a)) ℕₚ.≤-refl) ⟩
+ 1+ (1 +ℕ n₁ *ℕ 2) ≡⟨ refl ⟩
+ 1+ n₁ *ℕ 2 ≤⟨ ℕₚ.*-mono-≤ lt ℕₚ.≤-refl ⟩
+ n₂ *ℕ 2 ≤⟨ ℕₚ.n≤m+n m₂ (n₂ *ℕ 2) ⟩
+ m₂ +ℕ n₂ *ℕ 2 ∎)
+ where
+ open ℕₚ.≤-Reasoning
+ m₁ = Fin.toℕ a; m₂ = Fin.toℕ b
+ n₁ = toℕ (as 1#); n₂ = toℕ (bs 1#)
+
+∷ˡ-mono-< : ∀ {a b bs} → a Fin.< b → (a ∷ bs) 1# < (b ∷ bs) 1#
+∷ˡ-mono-< {a} {b} {bs} lt = less (begin
+ 1 +ℕ (m₁ +ℕ n *ℕ 2) ≡⟨ sym (ℕₚ.+-assoc 1 m₁ (n *ℕ 2)) ⟩
+ (1 +ℕ m₁) +ℕ n *ℕ 2 ≤⟨ ℕₚ.+-mono-≤ lt ℕₚ.≤-refl ⟩
+ m₂ +ℕ n *ℕ 2 ∎)
+ where
+ open ℕₚ.≤-Reasoning
+ m₁ = Fin.toℕ a; m₂ = Fin.toℕ b; n = toℕ (bs 1#)
+
+1<[23] : ∀ {b} → [] 1# < (b ∷ []) 1#
+1<[23] {b} = less (ℕₚ.n≤m+n (Fin.toℕ b) 2)
+
+1<2+ : ∀ {b bs} → [] 1# < (b ∷ bs) 1#
+1<2+ {_} {[]} = 1<[23]
+1<2+ {_} {b ∷ bs} = <-trans 1<[23] (∷ʳ-mono-< {a = b} 1<2+)
+
+0<1+ : ∀ {bs} → 0# < bs 1#
+0<1+ {[]} = less (s≤s z≤n)
+0<1+ {b ∷ bs} = <-trans (less (s≤s z≤n)) 1<2+
+
+<⇒≢ : ∀ {a b} → a < b → a ≢ b
+<⇒≢ lt eq = asym⟶irr (resp₂ _<_) sym <-asym eq lt
+
+<-cmp : Trichotomous _≡_ _<_
+<-cmp 0# 0# = tri≈ (<-irrefl refl) refl (<-irrefl refl)
+<-cmp 0# (_ 1#) = tri< 0<1+ (<⇒≢ 0<1+) (<-asym 0<1+)
+<-cmp (_ 1#) 0# = tri> (<-asym 0<1+) (<⇒≢ 0<1+ ∘ sym) 0<1+
+<-cmp ([] 1#) ([] 1#) = tri≈ (<-irrefl refl) refl (<-irrefl refl)
+<-cmp ([] 1#) ((b ∷ bs) 1#) = tri< 1<2+ (<⇒≢ 1<2+) (<-asym 1<2+)
+<-cmp ((a ∷ as) 1#) ([] 1#) = tri> (<-asym 1<2+) (<⇒≢ 1<2+ ∘ sym) 1<2+
+<-cmp ((a ∷ as) 1#) ((b ∷ bs) 1#) with <-cmp (as 1#) (bs 1#)
+... | tri< lt ¬eq ¬gt =
+ tri< (∷ʳ-mono-< lt) (<⇒≢ (∷ʳ-mono-< lt)) (<-asym (∷ʳ-mono-< lt))
+... | tri> ¬lt ¬eq gt =
+ tri> (<-asym (∷ʳ-mono-< gt)) (<⇒≢ (∷ʳ-mono-< gt) ∘ sym) (∷ʳ-mono-< gt)
+... | tri≈ ¬lt refl ¬gt with 𝔽ₚ.cmp a b
+... | tri≈ ¬lt′ refl ¬gt′ =
+ tri≈ (<-irrefl refl) refl (<-irrefl refl)
+... | tri< lt′ ¬eq ¬gt′ =
+ tri< (∷ˡ-mono-< lt′) (<⇒≢ (∷ˡ-mono-< lt′)) (<-asym (∷ˡ-mono-< lt′))
+... | tri> ¬lt′ ¬eq gt′ =
+ tri> (<-asym (∷ˡ-mono-< gt′)) (<⇒≢ (∷ˡ-mono-< gt′) ∘ sym) (∷ˡ-mono-< gt′)
+
+<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
+<-isStrictTotalOrder = record
+ { isEquivalence = isEquivalence
+ ; trans = <-trans
+ ; compare = <-cmp
+ }
+
+<-strictTotalOrder : StrictTotalOrder _ _ _
+<-strictTotalOrder = record
+ { Carrier = Bin
+ ; _≈_ = _≡_
+ ; _<_ = _<_
+ ; isStrictTotalOrder = <-isStrictTotalOrder
+ }
diff --git a/src/Data/BoundedVec.agda b/src/Data/BoundedVec.agda
index 6170f88..6359af2 100644
--- a/src/Data/BoundedVec.agda
+++ b/src/Data/BoundedVec.agda
@@ -11,6 +11,7 @@ module Data.BoundedVec where
open import Data.Nat
open import Data.List.Base as List using (List)
open import Data.Vec as Vec using (Vec)
+import Data.BoundedVec.Inefficient as Ineff
open import Relation.Binary.PropositionalEquality
open import Data.Nat.Properties
open SemiringSolver
@@ -20,15 +21,16 @@ open SemiringSolver
abstract
- data BoundedVec (a : Set) : ℕ → Set where
- bVec : ∀ {m n} (xs : Vec a n) → BoundedVec a (n + m)
+ data BoundedVec {a} (A : Set a) : ℕ → Set a where
+ bVec : ∀ {m n} (xs : Vec A n) → BoundedVec A (n + m)
- [] : ∀ {a n} → BoundedVec a n
+ [] : ∀ {a n} {A : Set a} → BoundedVec A n
[] = bVec Vec.[]
infixr 5 _∷_
- _∷_ : ∀ {a n} → a → BoundedVec a n → BoundedVec a (suc n)
+ _∷_ : ∀ {a n} {A : Set a} →
+ A → BoundedVec A n → BoundedVec A (suc n)
x ∷ bVec xs = bVec (Vec._∷_ x xs)
------------------------------------------------------------------------
@@ -36,13 +38,13 @@ abstract
infixr 5 _∷v_
-data View (a : Set) : ℕ → Set where
- []v : ∀ {n} → View a n
- _∷v_ : ∀ {n} (x : a) (xs : BoundedVec a n) → View a (suc n)
+data View {a} (A : Set a) : ℕ → Set a where
+ []v : ∀ {n} → View A n
+ _∷v_ : ∀ {n} (x : A) (xs : BoundedVec A n) → View A (suc n)
abstract
- view : ∀ {a n} → BoundedVec a n → View a n
+ view : ∀ {a n} {A : Set a} → BoundedVec A n → View A n
view (bVec Vec.[]) = []v
view (bVec (Vec._∷_ x xs)) = x ∷v bVec xs
@@ -51,9 +53,9 @@ abstract
abstract
- ↑ : ∀ {a n} → BoundedVec a n → BoundedVec a (suc n)
- ↑ {a = a} (bVec {m = m} {n = n} xs) =
- subst (BoundedVec a) lemma
+ ↑ : ∀ {a n} {A : Set a} → BoundedVec A n → BoundedVec A (suc n)
+ ↑ {A = A} (bVec {m = m} {n = n} xs) =
+ subst (BoundedVec A) lemma
(bVec {m = suc m} xs)
where
lemma : n + (1 + m) ≡ 1 + (n + m)
@@ -63,15 +65,26 @@ abstract
------------------------------------------------------------------------
-- Conversions
-abstract
+module _ {a} {A : Set a} where
+
+ abstract
- fromList : ∀ {a} → (xs : List a) → BoundedVec a (List.length xs)
- fromList {a = a} xs =
- subst (BoundedVec a) lemma
+ fromList : (xs : List A) → BoundedVec A (List.length xs)
+ fromList xs =
+ subst (BoundedVec A) lemma
(bVec {m = zero} (Vec.fromList xs))
where
lemma : List.length xs + 0 ≡ List.length xs
lemma = solve 1 (λ m → m :+ con 0 := m) refl _
- toList : ∀ {a n} → BoundedVec a n → List a
+ toList : ∀ {n} → BoundedVec A n → List A
toList (bVec xs) = Vec.toList xs
+
+ toInefficient : ∀ {n} → BoundedVec A n → Ineff.BoundedVec A n
+ toInefficient xs with view xs
+ ... | []v = Ineff.[]
+ ... | y ∷v ys = y Ineff.∷ toInefficient ys
+
+ fromInefficient : ∀ {n} → Ineff.BoundedVec A n → BoundedVec A n
+ fromInefficient Ineff.[] = []
+ fromInefficient (x Ineff.∷ xs) = x ∷ fromInefficient xs
diff --git a/src/Data/Char.agda b/src/Data/Char.agda
index 206adf7..be928a5 100644
--- a/src/Data/Char.agda
+++ b/src/Data/Char.agda
@@ -7,7 +7,7 @@
module Data.Char where
open import Data.Nat.Base using (ℕ)
-import Data.Nat.Properties as NatProp
+open import Data.Nat.Properties using (<-strictTotalOrder)
open import Data.Bool.Base using (Bool; true; false)
open import Relation.Nullary
open import Relation.Nullary.Decidable
@@ -62,4 +62,4 @@ decSetoid = PropEq.decSetoid _≟_
-- An ordering induced by the toNat function.
strictTotalOrder : StrictTotalOrder _ _ _
-strictTotalOrder = On.strictTotalOrder NatProp.strictTotalOrder toNat
+strictTotalOrder = On.strictTotalOrder <-strictTotalOrder toNat
diff --git a/src/Data/Colist.agda b/src/Data/Colist.agda
index dacc5ce..c49503e 100644
--- a/src/Data/Colist.agda
+++ b/src/Data/Colist.agda
@@ -43,9 +43,9 @@ data Colist {a} (A : Set a) : Set a where
[] : Colist A
_∷_ : (x : A) (xs : ∞ (Colist A)) → Colist A
-{-# HASKELL type AgdaColist a b = [b] #-}
-{-# COMPILED_DATA Colist MAlonzo.Code.Data.Colist.AgdaColist [] (:) #-}
-{-# COMPILED_DATA_UHC Colist __LIST__ __NIL__ __CONS__ #-}
+{-# FOREIGN GHC type AgdaColist a b = [b] #-}
+{-# COMPILE GHC Colist = data MAlonzo.Code.Data.Colist.AgdaColist ([] | (:)) #-}
+{-# COMPILE UHC Colist = data __LIST__ (__NIL__ | __CONS__) #-}
data Any {a p} {A : Set a} (P : A → Set p) :
Colist A → Set (a ⊔ p) where
diff --git a/src/Data/Colist/Infinite-merge.agda b/src/Data/Colist/Infinite-merge.agda
index 67cde1e..977168e 100644
--- a/src/Data/Colist/Infinite-merge.agda
+++ b/src/Data/Colist/Infinite-merge.agda
@@ -17,7 +17,7 @@ open import Function.Equality using (_⟨$⟩_)
open import Function.Inverse as Inv using (_↔_; module Inverse)
import Function.Related as Related
open import Function.Related.TypeIsomorphisms
-open import Induction.Nat
+open import Induction.Nat using (<′-well-founded)
import Induction.WellFounded as WF
open import Relation.Binary.PropositionalEquality as P using (_≡_)
open import Relation.Binary.Sum
@@ -199,7 +199,7 @@ Any-merge {P = P} = λ xss → record
to : ∀ xss p → Pred (xss , p)
to = λ xss p →
- WF.All.wfRec (WF.Inverse-image.well-founded size <-well-founded) _
+ WF.All.wfRec (WF.Inverse-image.well-founded size <′-well-founded) _
Pred step (xss , p)
where
size : Input → ℕ
diff --git a/src/Data/Container.agda b/src/Data/Container.agda
index 3d1909f..b0897ed 100644
--- a/src/Data/Container.agda
+++ b/src/Data/Container.agda
@@ -38,7 +38,7 @@ open Container public
-- The semantics ("extension") of a container.
-⟦_⟧ : ∀ {ℓ} → Container ℓ → Set ℓ → Set ℓ
+⟦_⟧ : ∀ {ℓ₁ ℓ₂} → Container ℓ₁ → Set ℓ₂ → Set (ℓ₁ ⊔ ℓ₂)
⟦ C ⟧ X = Σ[ s ∈ Shape C ] (Position C s → X)
-- The least and greatest fixpoints of a container.
@@ -93,7 +93,7 @@ setoid C X = record
-- Containers are functors.
-map : ∀ {c} {C : Container c} {X Y} → (X → Y) → ⟦ C ⟧ X → ⟦ C ⟧ Y
+map : ∀ {c ℓ} {C : Container c} {X Y : Set ℓ} → (X → Y) → ⟦ C ⟧ X → ⟦ C ⟧ Y
map f = Prod.map ⟨id⟩ (λ g → f ⟨∘⟩ g)
module Map where
@@ -123,8 +123,8 @@ open _⇒_ public
-- Interpretation of _⇒_.
-⟪_⟫ : ∀ {c} {C₁ C₂ : Container c} →
- C₁ ⇒ C₂ → ∀ {X} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X
+⟪_⟫ : ∀ {c ℓ} {C₁ C₂ : Container c} →
+ C₁ ⇒ C₂ → {X : Set ℓ} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X
⟪ m ⟫ xs = (shape m (proj₁ xs) , proj₂ xs ⟨∘⟩ position m)
module Morphism where
@@ -205,7 +205,7 @@ record _⊸_ {c} (C₁ C₂ : Container c) : Set c where
; position = _⟨$⟩_ (Inverse.to position⊸)
}
- ⟪_⟫⊸ : ∀ {X} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X
+ ⟪_⟫⊸ : ∀ {ℓ} {X : Set ℓ} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X
⟪_⟫⊸ = ⟪ morphism ⟫
open _⊸_ public using (shape⊸; position⊸; ⟪_⟫⊸)
diff --git a/src/Data/Container/Combinator.agda b/src/Data/Container/Combinator.agda
index dbeec2e..02f5283 100644
--- a/src/Data/Container/Combinator.agda
+++ b/src/Data/Container/Combinator.agda
@@ -82,7 +82,7 @@ const[ X ]⟶ C = Π {I = X} (F.const C)
module Identity where
- correct : ∀ {c} {X : Set c} → ⟦ id ⟧ X ↔ F.id X
+ correct : ∀ {c} {X : Set c} → ⟦ id {c} ⟧ X ↔ F.id X
correct {c} = record
{ to = P.→-to-⟶ {a = c} λ xs → proj₂ xs _
; from = P.→-to-⟶ {b₁ = c} λ x → (_ , λ _ → x)
diff --git a/src/Data/Container/Indexed.agda b/src/Data/Container/Indexed.agda
index cd1ab19..184e46e 100644
--- a/src/Data/Container/Indexed.agda
+++ b/src/Data/Container/Indexed.agda
@@ -61,8 +61,8 @@ private
Eq⇒≅ {xs = c , k} {.c , k′} ext (refl , refl , k≈k′) =
H.cong (_,_ c) (ext (λ _ → refl) (λ r → k≈k′ r r refl))
-setoid : ∀ {i o c r} {I : Set i} {O : Set o} →
- Container I O c r → Setoid I (c ⊔ r) _ → Setoid O _ _
+setoid : ∀ {i o c r s} {I : Set i} {O : Set o} →
+ Container I O c r → Setoid I s _ → Setoid O _ _
setoid C X = record
{ Carrier = ⟦ C ⟧ X.Carrier
; _≈_ = _≈_
@@ -99,15 +99,15 @@ map _ f = Prod.map ⟨id⟩ (λ g → f ⟨∘⟩ g)
module Map where
- identity : ∀ {i o c r} {I : Set i} {O : Set o} (C : Container I O c r)
- (X : Setoid I _ _) → let module X = Setoid X in
+ identity : ∀ {i o c r s} {I : Set i} {O : Set o} (C : Container I O c r)
+ (X : Setoid I s _) → let module X = Setoid X in
∀ {o} {xs : ⟦ C ⟧ X.Carrier o} → Eq C X.Carrier X.Carrier
X._≈_ xs (map C {X.Carrier} ⟨id⟩ xs)
identity C X = Setoid.refl (setoid C X)
- composition : ∀ {i o c r ℓ₁ ℓ₂} {I : Set i} {O : Set o}
+ composition : ∀ {i o c r s ℓ₁ ℓ₂} {I : Set i} {O : Set o}
(C : Container I O c r) {X : Pred I ℓ₁} {Y : Pred I ℓ₂}
- (Z : Setoid I _ _) → let module Z = Setoid Z in
+ (Z : Setoid I s _) → let module Z = Setoid Z in
{f : Y ⊆ Z.Carrier} {g : X ⊆ Y} {o : O} {xs : ⟦ C ⟧ X o} →
Eq C Z.Carrier Z.Carrier Z._≈_
(map C {Y} f (map C {X} g xs))
diff --git a/src/Data/Digit.agda b/src/Data/Digit.agda
index 58019b8..8ad6604 100644
--- a/src/Data/Digit.agda
+++ b/src/Data/Digit.agda
@@ -6,39 +6,23 @@
module Data.Digit where
-open import Data.Nat
+open import Data.Nat using (ℕ; zero; suc; pred; _+_; _*_; _≤?_; _≤′_)
open import Data.Nat.Properties
open SemiringSolver
open import Data.Fin as Fin using (Fin; zero; suc; toℕ)
-open import Relation.Nullary.Decidable
open import Data.Char using (Char)
open import Data.List.Base
open import Data.Product
open import Data.Vec as Vec using (Vec; _∷_; [])
-open import Induction.Nat
open import Data.Nat.DivMod
-open ≤-Reasoning
+open import Induction.Nat using (<′-rec; <′-Rec)
+open import Relation.Nullary using (yes; no)
+open import Relation.Nullary.Decidable
+open import Relation.Binary using (Decidable)
open import Relation.Binary.PropositionalEquality as P using (_≡_; refl)
open import Function
------------------------------------------------------------------------
--- A boring lemma
-
-private
-
- lem : ∀ x k r → 2 + x ≤′ r + (1 + x) * (2 + k)
- lem x k r = ≤⇒≤′ $ begin
- 2 + x
- ≤⟨ m≤m+n _ _ ⟩
- 2 + x + (x + (1 + x) * k + r)
- ≡⟨ solve 3 (λ x r k → con 2 :+ x :+ (x :+ (con 1 :+ x) :* k :+ r)
- :=
- r :+ (con 1 :+ x) :* (con 2 :+ k))
- refl x r k ⟩
- r + (1 + x) * (2 + k)
- ∎
-
-------------------------------------------------------------------------
-- Digits
-- Digit b is the type of digits in base b.
@@ -79,11 +63,14 @@ showDigit {base≤16 = base≤16} d =
------------------------------------------------------------------------
-- Digit expansions
+Expansion : ℕ → Set
+Expansion base = List (Fin base)
+
-- fromDigits takes a digit expansion of a natural number, starting
-- with the _least_ significant digit, and returns the corresponding
-- natural number.
-fromDigits : ∀ {base} → List (Fin base) → ℕ
+fromDigits : ∀ {base} → Expansion base → ℕ
fromDigits [] = 0
fromDigits {base} (d ∷ ds) = toℕ d + fromDigits ds * base
@@ -93,10 +80,10 @@ fromDigits {base} (d ∷ ds) = toℕ d + fromDigits ds * base
-- Note that the list of digits is always non-empty.
toDigits : (base : ℕ) {base≥2 : True (2 ≤? base)} (n : ℕ) →
- ∃ λ (ds : List (Fin base)) → fromDigits ds ≡ n
+ ∃ λ (ds : Expansion base) → fromDigits ds ≡ n
toDigits zero {base≥2 = ()} _
toDigits (suc zero) {base≥2 = ()} _
-toDigits (suc (suc k)) n = <-rec Pred helper n
+toDigits (suc (suc k)) n = <′-rec Pred helper n
where
base = suc (suc k)
Pred = λ n → ∃ λ ds → fromDigits ds ≡ n
@@ -104,8 +91,23 @@ toDigits (suc (suc k)) n = <-rec Pred helper n
cons : ∀ {m} (r : Fin base) → Pred m → Pred (toℕ r + m * base)
cons r (ds , eq) = (r ∷ ds , P.cong (λ i → toℕ r + i * base) eq)
- helper : ∀ n → <-Rec Pred n → Pred n
+ open ≤-Reasoning
+
+ lem : ∀ x k r → 2 + x ≤′ r + (1 + x) * (2 + k)
+ lem x k r = ≤⇒≤′ $ begin
+ 2 + x
+ ≤⟨ m≤m+n _ _ ⟩
+ 2 + x + (x + (1 + x) * k + r)
+ ≡⟨ solve 3 (λ x r k → con 2 :+ x :+ (x :+ (con 1 :+ x) :* k :+ r)
+ :=
+ r :+ (con 1 :+ x) :* (con 2 :+ k))
+ refl x r k ⟩
+ r + (1 + x) * (2 + k)
+ ∎
+
+ helper : ∀ n → <′-Rec Pred n → Pred n
helper n rec with n divMod base
helper .(toℕ r + 0 * base) rec | result zero r refl = ([ r ] , refl)
helper .(toℕ r + suc x * base) rec | result (suc x) r refl =
cons r (rec (suc x) (lem (pred (suc x)) k (toℕ r)))
+
diff --git a/src/Data/Empty.agda b/src/Data/Empty.agda
index 5c1fd5d..673a049 100644
--- a/src/Data/Empty.agda
+++ b/src/Data/Empty.agda
@@ -10,8 +10,8 @@ open import Level
data ⊥ : Set where
-{-# HASKELL data AgdaEmpty #-}
-{-# COMPILED_DATA ⊥ MAlonzo.Code.Data.Empty.AgdaEmpty #-}
+{-# FOREIGN GHC data AgdaEmpty #-}
+{-# COMPILE GHC ⊥ = data MAlonzo.Code.Data.Empty.AgdaEmpty () #-}
⊥-elim : ∀ {w} {Whatever : Set w} → ⊥ → Whatever
⊥-elim ()
diff --git a/src/Data/Empty/Irrelevant.agda b/src/Data/Empty/Irrelevant.agda
new file mode 100644
index 0000000..de36a8b
--- /dev/null
+++ b/src/Data/Empty/Irrelevant.agda
@@ -0,0 +1,12 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- An irrelevant version of ⊥-elim
+------------------------------------------------------------------------
+
+module Data.Empty.Irrelevant where
+
+open import Data.Empty hiding (⊥-elim)
+
+⊥-elim : ∀ {w} {Whatever : Set w} → .⊥ → Whatever
+⊥-elim ()
diff --git a/src/Data/Fin.agda b/src/Data/Fin.agda
index ba3f933..12e3950 100644
--- a/src/Data/Fin.agda
+++ b/src/Data/Fin.agda
@@ -10,6 +10,7 @@
module Data.Fin where
+open import Data.Empty using (⊥-elim)
open import Data.Nat as Nat
using (ℕ; zero; suc; z≤n; s≤s)
renaming ( _+_ to _N+_; _∸_ to _N∸_
@@ -19,7 +20,8 @@ import Level
open import Relation.Nullary
open import Relation.Nullary.Decidable
open import Relation.Binary
-open import Relation.Binary.PropositionalEquality as P using (_≡_)
+open import Relation.Binary.PropositionalEquality
+ using (_≡_; _≢_; refl; cong)
------------------------------------------------------------------------
-- Types
@@ -61,9 +63,9 @@ fromℕ≤ (Nat.s≤s (Nat.s≤s m≤n)) = suc (fromℕ≤ (Nat.s≤s m≤n))
-- fromℕ≤″ m _ = "m".
fromℕ≤″ : ∀ m {n} → m Nat.<″ n → Fin n
-fromℕ≤″ zero (Nat.less-than-or-equal P.refl) = zero
-fromℕ≤″ (suc m) (Nat.less-than-or-equal P.refl) =
- suc (fromℕ≤″ m (Nat.less-than-or-equal P.refl))
+fromℕ≤″ zero (Nat.less-than-or-equal refl) = zero
+fromℕ≤″ (suc m) (Nat.less-than-or-equal refl) =
+ suc (fromℕ≤″ m (Nat.less-than-or-equal refl))
-- # m = "m".
@@ -184,6 +186,25 @@ pred : ∀ {n} → Fin n → Fin n
pred zero = zero
pred (suc i) = inject₁ i
+-- The function f(i,j) = if j>i then j-1 else j
+-- This is a variant of the thick function from Conor
+-- McBride's "First-order unification by structural recursion".
+
+punchOut : ∀ {m} {i j : Fin (suc m)} → i ≢ j → Fin m
+punchOut {_} {zero} {zero} i≢j = ⊥-elim (i≢j refl)
+punchOut {_} {zero} {suc j} _ = j
+punchOut {zero} {suc ()}
+punchOut {suc m} {suc i} {zero} _ = zero
+punchOut {suc m} {suc i} {suc j} i≢j = suc (punchOut (i≢j ∘ cong suc))
+
+-- The function f(i,j) = if j≥i then j+1 else j
+
+punchIn : ∀ {m} → Fin (suc m) → Fin m → Fin (suc m)
+punchIn zero j = suc j
+punchIn (suc i) zero = zero
+punchIn (suc i) (suc j) = suc (punchIn i j)
+
+
------------------------------------------------------------------------
-- Order relations
diff --git a/src/Data/Fin/Dec.agda b/src/Data/Fin/Dec.agda
index ef3747d..2a31707 100644
--- a/src/Data/Fin/Dec.agda
+++ b/src/Data/Fin/Dec.agda
@@ -155,6 +155,14 @@ anySubset? {suc n} {P} dec with anySubset? (restrictS inside dec)
extend′ g zero = P0
extend′ g (suc j) = g j
+
+-- When P is a decidable predicate over a finite set the following
+-- lemma can be proved.
+
+¬∀⟶∃¬ : ∀ n {p} (P : Fin n → Set p) → U.Decidable P →
+ ¬ (∀ i → P i) → ∃ λ i → ¬ P i
+¬∀⟶∃¬ n P dec ¬P = Prod.map id proj₁ $ ¬∀⟶∃¬-smallest n P dec ¬P
+
-- Decision procedure for _⊆_ (obtained via the natural lattice
-- order).
diff --git a/src/Data/Fin/Properties.agda b/src/Data/Fin/Properties.agda
index e2724bc..5732040 100644
--- a/src/Data/Fin/Properties.agda
+++ b/src/Data/Fin/Properties.agda
@@ -8,10 +8,12 @@
module Data.Fin.Properties where
open import Algebra
+open import Data.Empty using (⊥-elim)
open import Data.Fin
-open import Data.Nat as N
- using (ℕ; zero; suc; s≤s; z≤n; _∸_)
- renaming (_≤_ to _ℕ≤_; _<_ to _ℕ<_; _+_ to _ℕ+_)
+open import Data.Nat as N using (ℕ; zero; suc; s≤s; z≤n; _∸_) renaming
+ (_≤_ to _ℕ≤_
+ ; _<_ to _ℕ<_
+ ; _+_ to _ℕ+_)
import Data.Nat.Properties as N
open import Data.Product
open import Function
@@ -22,53 +24,47 @@ open import Relation.Nullary
import Relation.Nullary.Decidable as Dec
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as P
- using (_≡_; refl; cong; subst)
+ using (_≡_; _≢_; refl; cong; subst)
open import Category.Functor
open import Category.Applicative
-open DecTotalOrder N.decTotalOrder using () renaming (refl to ℕ≤-refl)
-
------------------------------------------------------------------------
--- Properties
+-- Equality properties
+
+infix 4 _≟_
suc-injective : ∀ {o} {m n : Fin o} → Fin.suc m ≡ suc n → m ≡ n
suc-injective refl = refl
+_≟_ : {n : ℕ} → Decidable {A = Fin n} _≡_
+zero ≟ zero = yes refl
+zero ≟ suc y = no λ()
+suc x ≟ zero = no λ()
+suc x ≟ suc y with x ≟ y
+... | yes x≡y = yes (cong suc x≡y)
+... | no x≢y = no (x≢y ∘ suc-injective)
+
preorder : ℕ → Preorder _ _ _
preorder n = P.preorder (Fin n)
setoid : ℕ → Setoid _ _
setoid n = P.setoid (Fin n)
-cmp : ∀ {n} → Trichotomous _≡_ (_<_ {n})
-cmp zero zero = tri≈ (λ()) refl (λ())
-cmp zero (suc j) = tri< (s≤s z≤n) (λ()) (λ())
-cmp (suc i) zero = tri> (λ()) (λ()) (s≤s z≤n)
-cmp (suc i) (suc j) with cmp i j
-... | tri< lt ¬eq ¬gt = tri< (s≤s lt) (¬eq ∘ suc-injective) (¬gt ∘ N.≤-pred)
-... | tri> ¬lt ¬eq gt = tri> (¬lt ∘ N.≤-pred) (¬eq ∘ suc-injective) (s≤s gt)
-... | tri≈ ¬lt eq ¬gt = tri≈ (¬lt ∘ N.≤-pred) (cong suc eq) (¬gt ∘ N.≤-pred)
-
-strictTotalOrder : ℕ → StrictTotalOrder _ _ _
-strictTotalOrder n = record
- { Carrier = Fin n
- ; _≈_ = _≡_
- ; _<_ = _<_
- ; isStrictTotalOrder = record
- { isEquivalence = P.isEquivalence
- ; trans = N.<-trans
- ; compare = cmp
- }
+isDecEquivalence : ∀ {n} → IsDecEquivalence (_≡_ {A = Fin n})
+isDecEquivalence = record
+ { isEquivalence = P.isEquivalence
+ ; _≟_ = _≟_
}
-
decSetoid : ℕ → DecSetoid _ _
-decSetoid n = StrictTotalOrder.decSetoid (strictTotalOrder n)
-
-infix 4 _≟_
+decSetoid n = record
+ { Carrier = Fin n
+ ; _≈_ = _≡_
+ ; isDecEquivalence = isDecEquivalence
+ }
-_≟_ : {n : ℕ} → Decidable {A = Fin n} _≡_
-_≟_ {n} = DecSetoid._≟_ (decSetoid n)
+------------------------------------------------------------------------
+-- Converting between Fin n and Nat
to-from : ∀ n → toℕ (fromℕ n) ≡ n
to-from zero = refl
@@ -106,9 +102,105 @@ prop-toℕ-≤ (suc {n = suc n} i) = s≤s (prop-toℕ-≤ i)
prop-toℕ-≤′ : ∀ {n} (i : Fin n) → toℕ i ℕ≤ N.pred n
prop-toℕ-≤′ i = N.<⇒≤pred (bounded i)
+fromℕ≤-toℕ : ∀ {m} (i : Fin m) (i<m : toℕ i ℕ< m) → fromℕ≤ i<m ≡ i
+fromℕ≤-toℕ zero (s≤s z≤n) = refl
+fromℕ≤-toℕ (suc i) (s≤s (s≤s m≤n)) = cong suc (fromℕ≤-toℕ i (s≤s m≤n))
+
+toℕ-fromℕ≤ : ∀ {m n} (m<n : m ℕ< n) → toℕ (fromℕ≤ m<n) ≡ m
+toℕ-fromℕ≤ (s≤s z≤n) = refl
+toℕ-fromℕ≤ (s≤s (s≤s m<n)) = cong suc (toℕ-fromℕ≤ (s≤s m<n))
+
+-- fromℕ is a special case of fromℕ≤.
+fromℕ-def : ∀ n → fromℕ n ≡ fromℕ≤ N.≤-refl
+fromℕ-def zero = refl
+fromℕ-def (suc n) = cong suc (fromℕ-def n)
+
+-- fromℕ≤ and fromℕ≤″ give the same result.
+
+fromℕ≤≡fromℕ≤″ :
+ ∀ {m n} (m<n : m N.< n) (m<″n : m N.<″ n) →
+ fromℕ≤ m<n ≡ fromℕ≤″ m m<″n
+fromℕ≤≡fromℕ≤″ (s≤s z≤n) (N.less-than-or-equal refl) = refl
+fromℕ≤≡fromℕ≤″ (s≤s (s≤s m<n)) (N.less-than-or-equal refl) =
+ cong suc (fromℕ≤≡fromℕ≤″ (s≤s m<n) (N.less-than-or-equal refl))
+
+------------------------------------------------------------------------
+-- Ordering properties
+
+-- _≤_ ordering
+
+≤-reflexive : ∀ {n} → _≡_ ⇒ (_≤_ {n})
+≤-reflexive refl = N.≤-refl
+
+≤-refl : ∀ {n} → Reflexive (_≤_ {n})
+≤-refl = ≤-reflexive refl
+
+≤-trans : ∀ {n} → Transitive (_≤_ {n})
+≤-trans = N.≤-trans
+
+≤-antisym : ∀ {n} → Antisymmetric _≡_ (_≤_ {n})
+≤-antisym x≤y y≤x = toℕ-injective (N.≤-antisym x≤y y≤x)
+
+≤-total : ∀ {n} → Total (_≤_ {n})
+≤-total x y = N.≤-total (toℕ x) (toℕ y)
+
+≤-isPreorder : ∀ {n} → IsPreorder _≡_ (_≤_ {n})
+≤-isPreorder = record
+ { isEquivalence = P.isEquivalence
+ ; reflexive = ≤-reflexive
+ ; trans = ≤-trans
+ }
+
+≤-isPartialOrder : ∀ {n} → IsPartialOrder _≡_ (_≤_ {n})
+≤-isPartialOrder = record
+ { isPreorder = ≤-isPreorder
+ ; antisym = ≤-antisym
+ }
+
+≤-isTotalOrder : ∀ {n} → IsTotalOrder _≡_ (_≤_ {n})
+≤-isTotalOrder = record
+ { isPartialOrder = ≤-isPartialOrder
+ ; total = ≤-total
+ }
+
+-- _<_ ordering
+
+<-trans : ∀ {n} → Transitive (_<_ {n})
+<-trans = N.<-trans
+
+cmp : ∀ {n} → Trichotomous _≡_ (_<_ {n})
+cmp zero zero = tri≈ (λ()) refl (λ())
+cmp zero (suc j) = tri< (s≤s z≤n) (λ()) (λ())
+cmp (suc i) zero = tri> (λ()) (λ()) (s≤s z≤n)
+cmp (suc i) (suc j) with cmp i j
+... | tri< lt ¬eq ¬gt = tri< (s≤s lt) (¬eq ∘ suc-injective) (¬gt ∘ N.≤-pred)
+... | tri> ¬lt ¬eq gt = tri> (¬lt ∘ N.≤-pred) (¬eq ∘ suc-injective) (s≤s gt)
+... | tri≈ ¬lt eq ¬gt = tri≈ (¬lt ∘ N.≤-pred) (cong suc eq) (¬gt ∘ N.≤-pred)
+
+_<?_ : ∀ {n} → Decidable (_<_ {n})
+m <? n = suc (toℕ m) N.≤? toℕ n
+
+<-isStrictTotalOrder : ∀ {n} → IsStrictTotalOrder _≡_ (_<_ {n})
+<-isStrictTotalOrder = record
+ { isEquivalence = P.isEquivalence
+ ; trans = <-trans
+ ; compare = cmp
+ }
+
+strictTotalOrder : ℕ → StrictTotalOrder _ _ _
+strictTotalOrder n = record
+ { Carrier = Fin n
+ ; _≈_ = _≡_
+ ; _<_ = _<_
+ ; isStrictTotalOrder = <-isStrictTotalOrder
+ }
+
+------------------------------------------------------------------------
+-- Injection properties
+
-- Lemma: n - i ≤ n.
nℕ-ℕi≤n : ∀ n i → n ℕ-ℕ i ℕ≤ n
-nℕ-ℕi≤n n zero = ℕ≤-refl
+nℕ-ℕi≤n n zero = N.≤-refl
nℕ-ℕi≤n zero (suc ())
nℕ-ℕi≤n (suc n) (suc i) = begin
n ℕ-ℕ i ≤⟨ nℕ-ℕi≤n n i ⟩
@@ -154,35 +246,13 @@ toℕ-raise : ∀ {m} n (i : Fin m) → toℕ (raise n i) ≡ n ℕ+ toℕ i
toℕ-raise zero i = refl
toℕ-raise (suc n) i = cong suc (toℕ-raise n i)
-fromℕ≤-toℕ : ∀ {m} (i : Fin m) (i<m : toℕ i ℕ< m) → fromℕ≤ i<m ≡ i
-fromℕ≤-toℕ zero (s≤s z≤n) = refl
-fromℕ≤-toℕ (suc i) (s≤s (s≤s m≤n)) = cong suc (fromℕ≤-toℕ i (s≤s m≤n))
-
-toℕ-fromℕ≤ : ∀ {m n} (m<n : m ℕ< n) → toℕ (fromℕ≤ m<n) ≡ m
-toℕ-fromℕ≤ (s≤s z≤n) = refl
-toℕ-fromℕ≤ (s≤s (s≤s m<n)) = cong suc (toℕ-fromℕ≤ (s≤s m<n))
-
--- fromℕ is a special case of fromℕ≤.
-fromℕ-def : ∀ n → fromℕ n ≡ fromℕ≤ ℕ≤-refl
-fromℕ-def zero = refl
-fromℕ-def (suc n) = cong suc (fromℕ-def n)
-
--- fromℕ≤ and fromℕ≤″ give the same result.
-
-fromℕ≤≡fromℕ≤″ :
- ∀ {m n} (m<n : m N.< n) (m<″n : m N.<″ n) →
- fromℕ≤ m<n ≡ fromℕ≤″ m m<″n
-fromℕ≤≡fromℕ≤″ (s≤s z≤n) (N.less-than-or-equal refl) = refl
-fromℕ≤≡fromℕ≤″ (s≤s (s≤s m<n)) (N.less-than-or-equal refl) =
- cong suc (fromℕ≤≡fromℕ≤″ (s≤s m<n) (N.less-than-or-equal refl))
-
------------------------------------------------------------------------
-- Operations
infixl 6 _+′_
_+′_ : ∀ {m n} (i : Fin m) (j : Fin n) → Fin (N.pred m ℕ+ n)
-i +′ j = inject≤ (i + j) (N._+-mono_ (prop-toℕ-≤ i) ℕ≤-refl)
+i +′ j = inject≤ (i + j) (N.+-mono-≤ (prop-toℕ-≤ i) N.≤-refl)
-- reverse {n} "i" = "n ∸ 1 ∸ i".
@@ -211,11 +281,10 @@ reverse-involutive {n} i = toℕ-injective (begin
toℕ i ∎)
where
open P.≡-Reasoning
- open CommutativeSemiring N.commutativeSemiring using (+-comm)
lem₁ : ∀ m n → (m ℕ+ n) ∸ (m ℕ+ n ∸ m) ≡ m
lem₁ m n = begin
- m ℕ+ n ∸ (m ℕ+ n ∸ m) ≡⟨ cong (λ ξ → m ℕ+ n ∸ (ξ ∸ m)) (+-comm m n) ⟩
+ m ℕ+ n ∸ (m ℕ+ n ∸ m) ≡⟨ cong (λ ξ → m ℕ+ n ∸ (ξ ∸ m)) (N.+-comm m n) ⟩
m ℕ+ n ∸ (n ℕ+ m ∸ m) ≡⟨ cong (λ ξ → m ℕ+ n ∸ ξ) (N.m+n∸n≡m n m) ⟩
m ℕ+ n ∸ n ≡⟨ N.m+n∸n≡m m n ⟩
m ∎
@@ -284,3 +353,41 @@ private
F (∀ i → P i) → ∀ i → F (P i)
sequence⁻¹ RF F∀iPi i = (λ f → f i) <$> F∀iPi
where open RawFunctor RF
+
+------------------------------------------------------------------------
+
+punchOut-injective : ∀ {m} {i j k : Fin (suc m)}
+ (i≢j : i ≢ j) (i≢k : i ≢ k) →
+ punchOut i≢j ≡ punchOut i≢k → j ≡ k
+punchOut-injective {_} {zero} {zero} {_} i≢j _ _ = ⊥-elim (i≢j refl)
+punchOut-injective {_} {zero} {_} {zero} _ i≢k _ = ⊥-elim (i≢k refl)
+punchOut-injective {_} {zero} {suc j} {suc k} _ _ pⱼ≡pₖ = cong suc pⱼ≡pₖ
+punchOut-injective {zero} {suc ()}
+punchOut-injective {suc n} {suc i} {zero} {zero} _ _ _ = refl
+punchOut-injective {suc n} {suc i} {zero} {suc k} _ _ ()
+punchOut-injective {suc n} {suc i} {suc j} {zero} _ _ ()
+punchOut-injective {suc n} {suc i} {suc j} {suc k} i≢j i≢k pⱼ≡pₖ =
+ cong suc (punchOut-injective (i≢j ∘ cong suc) (i≢k ∘ cong suc) (suc-injective pⱼ≡pₖ))
+
+punchIn-injective : ∀ {m} i (j k : Fin m) →
+ punchIn i j ≡ punchIn i k → j ≡ k
+punchIn-injective zero _ _ refl = refl
+punchIn-injective (suc i) zero zero _ = refl
+punchIn-injective (suc i) zero (suc k) ()
+punchIn-injective (suc i) (suc j) zero ()
+punchIn-injective (suc i) (suc j) (suc k) ↑j+1≡↑k+1 =
+ cong suc (punchIn-injective i j k (suc-injective ↑j+1≡↑k+1))
+
+punchIn-punchOut : ∀ {m} {i j : Fin (suc m)} (i≢j : i ≢ j) →
+ punchIn i (punchOut i≢j) ≡ j
+punchIn-punchOut {_} {zero} {zero} 0≢0 = ⊥-elim (0≢0 refl)
+punchIn-punchOut {_} {zero} {suc j} _ = refl
+punchIn-punchOut {zero} {suc ()}
+punchIn-punchOut {suc m} {suc i} {zero} i≢j = refl
+punchIn-punchOut {suc m} {suc i} {suc j} i≢j =
+ cong suc (punchIn-punchOut (i≢j ∘ cong suc))
+
+punchInᵢ≢i : ∀ {m} i (j : Fin m) → punchIn i j ≢ i
+punchInᵢ≢i zero _ ()
+punchInᵢ≢i (suc i) zero ()
+punchInᵢ≢i (suc i) (suc j) = punchInᵢ≢i i j ∘ suc-injective
diff --git a/src/Data/Fin/Subset.agda b/src/Data/Fin/Subset.agda
index 4cd310b..9b16298 100644
--- a/src/Data/Fin/Subset.agda
+++ b/src/Data/Fin/Subset.agda
@@ -11,15 +11,13 @@ import Algebra.Properties.BooleanAlgebra as BoolAlgProp
import Algebra.Properties.BooleanAlgebra.Expression as BAExpr
import Data.Bool.Properties as BoolProp
open import Data.Fin
-open import Data.List.Base as List using (List)
+open import Data.List.Base using (List; foldr; foldl)
open import Data.Nat
open import Data.Product
open import Data.Vec using (Vec; _∷_; _[_]=_)
import Relation.Binary.Vec.Pointwise as Pointwise
open import Relation.Nullary
-infix 4 _∈_ _∉_ _⊆_ _⊈_
-
------------------------------------------------------------------------
-- Definitions
@@ -36,6 +34,8 @@ Subset = Vec Side
------------------------------------------------------------------------
-- Membership and subset predicates
+infix 4 _∈_ _∉_ _⊆_ _⊈_
+
_∈_ : ∀ {n} → Fin n → Subset n → Set
x ∈ p = p [ x ]= inside
@@ -84,12 +84,12 @@ private
-- N-ary union.
⋃ : ∀ {n} → List (Subset n) → Subset n
-⋃ = List.foldr _∪_ ⊥
+⋃ = foldr _∪_ ⊥
-- N-ary intersection.
⋂ : ∀ {n} → List (Subset n) → Subset n
-⋂ = List.foldr _∩_ ⊤
+⋂ = foldr _∩_ ⊤
------------------------------------------------------------------------
-- Properties
diff --git a/src/Data/Fin/Subset/Properties.agda b/src/Data/Fin/Subset/Properties.agda
index ada8d22..5271457 100644
--- a/src/Data/Fin/Subset/Properties.agda
+++ b/src/Data/Fin/Subset/Properties.agda
@@ -9,10 +9,10 @@ module Data.Fin.Subset.Properties where
open import Algebra
import Algebra.Properties.BooleanAlgebra as BoolProp
open import Data.Empty using (⊥-elim)
-open import Data.Fin using (Fin); open Data.Fin.Fin
+open import Data.Fin using (Fin; suc; zero)
open import Data.Fin.Subset
open import Data.Nat.Base using (ℕ)
-open import Data.Product
+open import Data.Product as Product
open import Data.Sum as Sum
open import Data.Vec hiding (_∈_)
open import Function
@@ -20,7 +20,9 @@ open import Function.Equality using (_⟨$⟩_)
open import Function.Equivalence
using (_⇔_; equivalence; module Equivalence)
open import Relation.Binary
-open import Relation.Binary.PropositionalEquality as P using (_≡_)
+open import Relation.Binary.PropositionalEquality
+ using (_≡_; refl; cong; subst)
+open import Relation.Nullary.Negation using (contradiction)
------------------------------------------------------------------------
-- Constructor mangling
@@ -41,16 +43,13 @@ drop-∷-Empty ¬∃∈ (x , x∈p) = ¬∃∈ (suc x , there x∈p)
∉⊥ (there p) = ∉⊥ p
⊥⊆ : ∀ {n} {p : Subset n} → ⊥ ⊆ p
-⊥⊆ x∈⊥ with ∉⊥ x∈⊥
-... | ()
+⊥⊆ x∈⊥ = contradiction x∈⊥ ∉⊥
-Empty-unique : ∀ {n} {p : Subset n} →
- Empty p → p ≡ ⊥
-Empty-unique {p = []} ¬∃∈ = P.refl
-Empty-unique {p = s ∷ p} ¬∃∈ with Empty-unique (drop-∷-Empty ¬∃∈)
-Empty-unique {p = outside ∷ .⊥} ¬∃∈ | P.refl = P.refl
-Empty-unique {p = inside ∷ .⊥} ¬∃∈ | P.refl =
- ⊥-elim (¬∃∈ (zero , here))
+Empty-unique : ∀ {n} {p : Subset n} → Empty p → p ≡ ⊥
+Empty-unique {_} {[]} ¬∃∈ = refl
+Empty-unique {_} {inside ∷ p} ¬∃∈ = contradiction (zero , here) ¬∃∈
+Empty-unique {_} {outside ∷ p} ¬∃∈ =
+ cong (outside ∷_) (Empty-unique (drop-∷-Empty ¬∃∈))
------------------------------------------------------------------------
-- Properties involving ⊤
@@ -63,48 +62,83 @@ Empty-unique {p = inside ∷ .⊥} ¬∃∈ | P.refl =
⊆⊤ = const ∈⊤
------------------------------------------------------------------------
--- A property involving ⁅_⁆
+-- Properties involving ⁅_⁆
-x∈⁅y⁆⇔x≡y : ∀ {n} {x y : Fin n} → x ∈ ⁅ y ⁆ ⇔ x ≡ y
-x∈⁅y⁆⇔x≡y {x = x} {y} =
- equivalence (to y) (λ x≡y → P.subst (λ y → x ∈ ⁅ y ⁆) x≡y (x∈⁅x⁆ x))
- where
+x∈⁅x⁆ : ∀ {n} (x : Fin n) → x ∈ ⁅ x ⁆
+x∈⁅x⁆ zero = here
+x∈⁅x⁆ (suc x) = there (x∈⁅x⁆ x)
- to : ∀ {n x} (y : Fin n) → x ∈ ⁅ y ⁆ → x ≡ y
- to (suc y) (there p) = P.cong suc (to y p)
- to zero here = P.refl
- to zero (there p) with ∉⊥ p
- ... | ()
+x∈⁅y⁆⇒x≡y : ∀ {n x} (y : Fin n) → x ∈ ⁅ y ⁆ → x ≡ y
+x∈⁅y⁆⇒x≡y zero here = refl
+x∈⁅y⁆⇒x≡y zero (there p) = contradiction p ∉⊥
+x∈⁅y⁆⇒x≡y (suc y) (there p) = cong suc (x∈⁅y⁆⇒x≡y y p)
- x∈⁅x⁆ : ∀ {n} (x : Fin n) → x ∈ ⁅ x ⁆
- x∈⁅x⁆ zero = here
- x∈⁅x⁆ (suc x) = there (x∈⁅x⁆ x)
+x∈⁅y⁆⇔x≡y : ∀ {n} {x y : Fin n} → x ∈ ⁅ y ⁆ ⇔ x ≡ y
+x∈⁅y⁆⇔x≡y {_} {x} {y} = equivalence
+ (x∈⁅y⁆⇒x≡y y)
+ (λ x≡y → subst (λ y → x ∈ ⁅ y ⁆) x≡y (x∈⁅x⁆ x))
------------------------------------------------------------------------
--- A property involving _∪_
-
-∪⇔⊎ : ∀ {n} {p₁ p₂ : Subset n} {x} → x ∈ p₁ ∪ p₂ ⇔ (x ∈ p₁ ⊎ x ∈ p₂)
-∪⇔⊎ = equivalence (to _ _) from
- where
- to : ∀ {n} (p₁ p₂ : Subset n) {x} → x ∈ p₁ ∪ p₂ → x ∈ p₁ ⊎ x ∈ p₂
- to [] [] ()
- to (inside ∷ p₁) (s₂ ∷ p₂) here = inj₁ here
- to (outside ∷ p₁) (inside ∷ p₂) here = inj₂ here
- to (s₁ ∷ p₁) (s₂ ∷ p₂) (there x∈p₁∪p₂) =
- Sum.map there there (to p₁ p₂ x∈p₁∪p₂)
-
- ⊆∪ˡ : ∀ {n p₁} (p₂ : Subset n) → p₁ ⊆ p₁ ∪ p₂
- ⊆∪ˡ [] ()
- ⊆∪ˡ (s ∷ p₂) here = here
- ⊆∪ˡ (s ∷ p₂) (there x∈p₁) = there (⊆∪ˡ p₂ x∈p₁)
-
- ⊆∪ʳ : ∀ {n} (p₁ p₂ : Subset n) → p₂ ⊆ p₁ ∪ p₂
- ⊆∪ʳ p₁ p₂ rewrite BooleanAlgebra.∨-comm (booleanAlgebra _) p₁ p₂
- = ⊆∪ˡ p₁
-
- from : ∀ {n} {p₁ p₂ : Subset n} {x} → x ∈ p₁ ⊎ x ∈ p₂ → x ∈ p₁ ∪ p₂
- from (inj₁ x∈p₁) = ⊆∪ˡ _ x∈p₁
- from (inj₂ x∈p₂) = ⊆∪ʳ _ _ x∈p₂
+-- Properties involving _∪_ and _∩_
+
+module _ {n : ℕ} where
+ open BooleanAlgebra (booleanAlgebra n) public using ()
+ renaming
+ ( ∨-assoc to ∪-assoc
+ ; ∨-comm to ∪-comm
+ ; ∧-assoc to ∩-assoc
+ ; ∧-comm to ∩-comm
+ )
+
+-- _∪_
+
+p⊆p∪q : ∀ {n p} (q : Subset n) → p ⊆ p ∪ q
+p⊆p∪q [] ()
+p⊆p∪q (s ∷ q) here = here
+p⊆p∪q (s ∷ q) (there x∈p) = there (p⊆p∪q q x∈p)
+
+q⊆p∪q : ∀ {n} (p q : Subset n) → q ⊆ p ∪ q
+q⊆p∪q p q rewrite ∪-comm p q = p⊆p∪q p
+
+x∈p∪q⁻ : ∀ {n} (p q : Subset n) {x} → x ∈ p ∪ q → x ∈ p ⊎ x ∈ q
+x∈p∪q⁻ [] [] ()
+x∈p∪q⁻ (inside ∷ p) (t ∷ q) here = inj₁ here
+x∈p∪q⁻ (outside ∷ p) (inside ∷ q) here = inj₂ here
+x∈p∪q⁻ (s ∷ p) (t ∷ q) (there x∈p∪q) =
+ Sum.map there there (x∈p∪q⁻ p q x∈p∪q)
+
+x∈p∪q⁺ : ∀ {n} {p q : Subset n} {x} → x ∈ p ⊎ x ∈ q → x ∈ p ∪ q
+x∈p∪q⁺ (inj₁ x∈p) = p⊆p∪q _ x∈p
+x∈p∪q⁺ (inj₂ x∈q) = q⊆p∪q _ _ x∈q
+
+∪⇔⊎ : ∀ {n} {p q : Subset n} {x} → x ∈ p ∪ q ⇔ (x ∈ p ⊎ x ∈ q)
+∪⇔⊎ = equivalence (x∈p∪q⁻ _ _) x∈p∪q⁺
+
+-- _∩_
+
+p∩q⊆p : ∀ {n} (p q : Subset n) → p ∩ q ⊆ p
+p∩q⊆p [] [] x∈p∩q = x∈p∩q
+p∩q⊆p (inside ∷ p) (inside ∷ q) here = here
+p∩q⊆p (inside ∷ p) (inside ∷ q) (there ∈p∩q) = there (p∩q⊆p p q ∈p∩q)
+p∩q⊆p (outside ∷ p) (inside ∷ q) (there ∈p∩q) = there (p∩q⊆p p q ∈p∩q)
+p∩q⊆p (inside ∷ p) (outside ∷ q) (there ∈p∩q) = there (p∩q⊆p p q ∈p∩q)
+p∩q⊆p (outside ∷ p) (outside ∷ q) (there ∈p∩q) = there (p∩q⊆p p q ∈p∩q)
+
+p∩q⊆q : ∀ {n} (p q : Subset n) → p ∩ q ⊆ q
+p∩q⊆q p q rewrite ∩-comm p q = p∩q⊆p q p
+
+x∈p∩q⁺ : ∀ {n} {p q : Subset n} {x} → x ∈ p × x ∈ q → x ∈ p ∩ q
+x∈p∩q⁺ (here , here) = here
+x∈p∩q⁺ (there x∈p , there x∈q) = there (x∈p∩q⁺ (x∈p , x∈q))
+
+x∈p∩q⁻ : ∀ {n} (p q : Subset n) {x} → x ∈ p ∩ q → x ∈ p × x ∈ q
+x∈p∩q⁻ [] [] ()
+x∈p∩q⁻ (inside ∷ p) (inside ∷ q) here = here , here
+x∈p∩q⁻ (s ∷ p) (t ∷ q) (there x∈p∩q) =
+ Product.map there there (x∈p∩q⁻ p q x∈p∩q)
+
+∩⇔× : ∀ {n} {p q : Subset n} {x} → x ∈ p ∩ q ⇔ (x ∈ p × x ∈ q)
+∩⇔× = equivalence (x∈p∩q⁻ _ _) x∈p∩q⁺
------------------------------------------------------------------------
-- _⊆_ is a partial order
@@ -115,7 +149,7 @@ module NaturalPoset where
private
open module BA {n} = BoolProp (booleanAlgebra n) public
using (poset)
- open module Po {n} = Poset (poset {n = n}) public
+ open module Po {n} = Poset (poset {n = n}) public hiding (refl)
-- _⊆_ is equivalent to the natural lattice order.
@@ -123,16 +157,16 @@ module NaturalPoset where
orders-equivalent = equivalence (to _ _) (from _ _)
where
to : ∀ {n} (p₁ p₂ : Subset n) → p₁ ⊆ p₂ → p₁ ≤ p₂
- to [] [] p₁⊆p₂ = P.refl
+ to [] [] p₁⊆p₂ = refl
to (inside ∷ p₁) (_ ∷ p₂) p₁⊆p₂ with p₁⊆p₂ here
- to (inside ∷ p₁) (.inside ∷ p₂) p₁⊆p₂ | here = P.cong (_∷_ inside) (to p₁ p₂ (drop-∷-⊆ p₁⊆p₂))
- to (outside ∷ p₁) (_ ∷ p₂) p₁⊆p₂ = P.cong (_∷_ outside) (to p₁ p₂ (drop-∷-⊆ p₁⊆p₂))
+ to (inside ∷ p₁) (.inside ∷ p₂) p₁⊆p₂ | here = cong (_∷_ inside) (to p₁ p₂ (drop-∷-⊆ p₁⊆p₂))
+ to (outside ∷ p₁) (_ ∷ p₂) p₁⊆p₂ = cong (_∷_ outside) (to p₁ p₂ (drop-∷-⊆ p₁⊆p₂))
from : ∀ {n} (p₁ p₂ : Subset n) → p₁ ≤ p₂ → p₁ ⊆ p₂
from [] [] p₁≤p₂ x = x
- from (.inside ∷ _) (_ ∷ _) p₁≤p₂ here rewrite P.cong head p₁≤p₂ = here
+ from (.inside ∷ _) (_ ∷ _) p₁≤p₂ here rewrite cong head p₁≤p₂ = here
from (_ ∷ p₁) (_ ∷ p₂) p₁≤p₂ (there xs[i]=x) =
- there (from p₁ p₂ (P.cong tail p₁≤p₂) xs[i]=x)
+ there (from p₁ p₂ (cong tail p₁≤p₂) xs[i]=x)
-- _⊆_ is a partial order.
@@ -143,7 +177,7 @@ poset n = record
; _≤_ = _⊆_
; isPartialOrder = record
{ isPreorder = record
- { isEquivalence = P.isEquivalence
+ { isEquivalence = isEquivalence
; reflexive = λ i≡j → from ⟨$⟩ reflexive i≡j
; trans = λ x⊆y y⊆z → from ⟨$⟩ trans (to ⟨$⟩ x⊆y) (to ⟨$⟩ y⊆z)
}
diff --git a/src/Data/Fin/Substitution/Lemmas.agda b/src/Data/Fin/Substitution/Lemmas.agda
index 225dbb4..cf0bb81 100644
--- a/src/Data/Fin/Substitution/Lemmas.agda
+++ b/src/Data/Fin/Substitution/Lemmas.agda
@@ -55,10 +55,9 @@ record Lemmas₀ (T : ℕ → Set) : Set where
lookup-map-weaken-↑⋆ zero x = refl
lookup-map-weaken-↑⋆ (suc k) zero = refl
lookup-map-weaken-↑⋆ (suc k) (suc x) {ρ} = begin
- lookup x (map weaken (map weaken ρ ↑⋆ k)) ≡⟨ Applicative.Morphism.op-<$> (VecProp.lookup-morphism x) weaken _ ⟩
+ lookup x (map weaken (map weaken ρ ↑⋆ k)) ≡⟨ VecProp.lookup-map x weaken _ ⟩
weaken (lookup x (map weaken ρ ↑⋆ k)) ≡⟨ cong weaken (lookup-map-weaken-↑⋆ k x) ⟩
- weaken (lookup (lift k suc x) ((ρ ↑) ↑⋆ k)) ≡⟨ sym $
- Applicative.Morphism.op-<$> (VecProp.lookup-morphism (lift k suc x)) weaken _ ⟩
+ weaken (lookup (lift k suc x) ((ρ ↑) ↑⋆ k)) ≡⟨ sym $ VecProp.lookup-map (lift k suc x) _ _ ⟩
lookup (lift k suc x) (map weaken ((ρ ↑) ↑⋆ k)) ∎
record Lemmas₁ (T : ℕ → Set) : Set where
@@ -73,7 +72,7 @@ record Lemmas₁ (T : ℕ → Set) : Set where
lookup x ρ ≡ var y →
lookup x (map weaken ρ) ≡ var (suc y)
lookup-map-weaken x {y} {ρ} hyp = begin
- lookup x (map weaken ρ) ≡⟨ Applicative.Morphism.op-<$> (VecProp.lookup-morphism x) weaken ρ ⟩
+ lookup x (map weaken ρ) ≡⟨ VecProp.lookup-map x _ _ ⟩
weaken (lookup x ρ) ≡⟨ cong weaken hyp ⟩
weaken (var y) ≡⟨ weaken-var ⟩
var (suc y) ∎
@@ -141,7 +140,7 @@ record Lemmas₂ (T : ℕ → Set) : Set where
lookup-⊙ : ∀ {m n k} x {ρ₁ : Sub T m n} {ρ₂ : Sub T n k} →
lookup x (ρ₁ ⊙ ρ₂) ≡ lookup x ρ₁ / ρ₂
- lookup-⊙ x = Applicative.Morphism.op-<$> (VecProp.lookup-morphism x) _ _
+ lookup-⊙ x = VecProp.lookup-map x _ _
lookup-⨀ : ∀ {m n} x (ρs : Subs T m n) →
lookup x (⨀ ρs) ≡ var x /✶ ρs
diff --git a/src/Data/Graph/Acyclic.agda b/src/Data/Graph/Acyclic.agda
index 9e48fe8..3a93790 100644
--- a/src/Data/Graph/Acyclic.agda
+++ b/src/Data/Graph/Acyclic.agda
@@ -22,7 +22,7 @@ open import Data.Unit.Base using (⊤; tt)
open import Data.Vec as Vec using (Vec; []; _∷_)
open import Data.List.Base as List using (List; []; _∷_)
open import Function
-open import Induction.Nat
+open import Induction.Nat using (<′-rec; <′-Rec)
open import Relation.Nullary
open import Relation.Binary.PropositionalEquality as P using (_≡_)
@@ -286,11 +286,11 @@ data Tree (N E : Set) : Set where
node : (label : N) (successors : List (E × Tree N E)) → Tree N E
toTree : ∀ {N E n} → Graph N E n → Fin n → Tree N E
-toTree {N} {E} g i = <-rec Pred expand _ (g [ i ])
+toTree {N} {E} g i = <′-rec Pred expand _ (g [ i ])
where
Pred = λ n → Graph N E (suc n) → Tree N E
- expand : (n : ℕ) → <-Rec Pred n → Pred n
+ expand : (n : ℕ) → <′-Rec Pred n → Pred n
expand n rec (c & g) =
node (label c)
(List.map
diff --git a/src/Data/Integer.agda b/src/Data/Integer.agda
index db30cc8..cc6b827 100644
--- a/src/Data/Integer.agda
+++ b/src/Data/Integer.agda
@@ -7,6 +7,7 @@
module Data.Integer where
import Data.Nat as ℕ
+import Data.Nat.Properties as ℕP
import Data.Nat.Show as ℕ
open import Data.Sign as Sign using (Sign)
open import Data.String.Base using (String; _++_)
@@ -15,9 +16,9 @@ open import Data.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
+open import Relation.Binary.PropositionalEquality
+ using (_≡_; refl; sym; subst; cong; cong₂; module ≡-Reasoning)
+open ≡-Reasoning
------------------------------------------------------------------------
-- Integers, basic types and operations
@@ -47,8 +48,7 @@ show i = showSign (sign i) ++ ℕ.show ∣ i ∣
j ∎
signAbs : ∀ i → SignAbs i
-signAbs i = PropEq.subst SignAbs (◃-left-inverse i) $
- sign i ◂ ∣ i ∣
+signAbs i = subst SignAbs (◃-left-inverse i) (sign i ◂ ∣ i ∣)
------------------------------------------------------------------------
-- Equality is decidable
@@ -61,11 +61,8 @@ 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 _≟_
-
------------------------------------------------------------------------
--- Ordering
+-- Ordering is decidable
infix 4 _≤?_
@@ -74,55 +71,3 @@ _≤?_ : Decidable _≤_
-[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 _≡⟨_⟩_)
diff --git a/src/Data/Integer/Addition/Properties.agda b/src/Data/Integer/Addition/Properties.agda
index 35c24be..2da1bdf 100644
--- a/src/Data/Integer/Addition/Properties.agda
+++ b/src/Data/Integer/Addition/Properties.agda
@@ -2,114 +2,26 @@
-- The Agda standard library
--
-- Properties related to addition of integers
+--
+-- This module is DEPRECATED. Please use the corresponding properties in
+-- Data.Integer.Properties directly.
------------------------------------------------------------------------
module Data.Integer.Addition.Properties where
-open import Algebra
-import Algebra.FunctionProperties
-open import Algebra.Structures
-open import Data.Integer hiding (suc)
-open import Data.Nat.Base using (suc; zero) renaming (_+_ to _ℕ+_)
-import Data.Nat.Properties as ℕ
-open import Relation.Binary.PropositionalEquality
-
-open Algebra.FunctionProperties (_≡_ {A = ℤ})
-open CommutativeSemiring ℕ.commutativeSemiring
- using (+-comm; +-assoc; +-identity)
-
-------------------------------------------------------------------------
--- Addition and zero form a commutative monoid
-
-comm : Commutative _+_
-comm -[1+ a ] -[1+ b ] rewrite +-comm a b = refl
-comm (+ a ) (+ b ) rewrite +-comm a b = refl
-comm -[1+ _ ] (+ _ ) = refl
-comm (+ _ ) -[1+ _ ] = refl
-
-identityˡ : LeftIdentity (+ 0) _+_
-identityˡ -[1+ _ ] = refl
-identityˡ (+ _ ) = refl
-
-identityʳ : RightIdentity (+ 0) _+_
-identityʳ x rewrite comm x (+ 0) = identityˡ x
-
-distribˡ-⊖-+-neg : ∀ a b c → b ⊖ c + -[1+ a ] ≡ b ⊖ (suc c ℕ+ a)
-distribˡ-⊖-+-neg _ zero zero = refl
-distribˡ-⊖-+-neg _ zero (suc _) = refl
-distribˡ-⊖-+-neg _ (suc _) zero = refl
-distribˡ-⊖-+-neg a (suc b) (suc c) = distribˡ-⊖-+-neg a b c
-
-distribʳ-⊖-+-neg : ∀ a b c → -[1+ a ] + (b ⊖ c) ≡ b ⊖ (suc a ℕ+ c)
-distribʳ-⊖-+-neg a b c
- rewrite comm -[1+ a ] (b ⊖ c)
- | distribˡ-⊖-+-neg a b c
- | +-comm a c
- = refl
-
-distribˡ-⊖-+-pos : ∀ a b c → b ⊖ c + + a ≡ b ℕ+ a ⊖ c
-distribˡ-⊖-+-pos _ zero zero = refl
-distribˡ-⊖-+-pos _ zero (suc _) = refl
-distribˡ-⊖-+-pos _ (suc _) zero = refl
-distribˡ-⊖-+-pos a (suc b) (suc c) = distribˡ-⊖-+-pos a b c
-
-distribʳ-⊖-+-pos : ∀ a b c → + a + (b ⊖ c) ≡ a ℕ+ b ⊖ c
-distribʳ-⊖-+-pos a b c
- rewrite comm (+ a) (b ⊖ c)
- | distribˡ-⊖-+-pos a b c
- | +-comm a b
- = refl
-
-assoc : Associative _+_
-assoc (+ zero) y z rewrite identityˡ y | identityˡ (y + z) = refl
-assoc x (+ zero) z rewrite identityʳ x | identityˡ z = refl
-assoc x y (+ zero) rewrite identityʳ (x + y) | identityʳ y = refl
-assoc -[1+ a ] -[1+ b ] (+ suc c) = sym (distribʳ-⊖-+-neg a c b)
-assoc -[1+ a ] (+ suc b) (+ suc c) = distribˡ-⊖-+-pos (suc c) b a
-assoc (+ suc a) -[1+ b ] -[1+ c ] = distribˡ-⊖-+-neg c a b
-assoc (+ suc a) -[1+ b ] (+ suc c)
- rewrite distribˡ-⊖-+-pos (suc c) a b
- | distribʳ-⊖-+-pos (suc a) c b
- | sym (+-assoc a 1 c)
- | +-comm a 1
- = refl
-assoc (+ suc a) (+ suc b) -[1+ c ]
- rewrite distribʳ-⊖-+-pos (suc a) b c
- | sym (+-assoc a 1 b)
- | +-comm a 1
- = refl
-assoc -[1+ a ] -[1+ b ] -[1+ c ]
- rewrite sym (+-assoc a 1 (b ℕ+ c))
- | +-comm a 1
- | +-assoc a b c
- = refl
-assoc -[1+ a ] (+ suc b) -[1+ c ]
- rewrite distribʳ-⊖-+-neg a b c
- | distribˡ-⊖-+-neg c b a
- = refl
-assoc (+ suc a) (+ suc b) (+ suc c)
- rewrite +-assoc (suc a) (suc b) (suc c)
- = refl
-
-isSemigroup : IsSemigroup _≡_ _+_
-isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = assoc
- ; ∙-cong = cong₂ _+_
- }
-
-isCommutativeMonoid : IsCommutativeMonoid _≡_ _+_ (+ 0)
-isCommutativeMonoid = record
- { isSemigroup = isSemigroup
- ; identityˡ = identityˡ
- ; comm = comm
- }
-
-commutativeMonoid : CommutativeMonoid _ _
-commutativeMonoid = record
- { Carrier = ℤ
- ; _≈_ = _≡_
- ; _∙_ = _+_
- ; ε = + 0
- ; isCommutativeMonoid = isCommutativeMonoid
- }
+open import Data.Integer.Properties public
+ using
+ ( distribˡ-⊖-+-neg
+ ; distribʳ-⊖-+-neg
+ ; distribˡ-⊖-+-pos
+ ; distribʳ-⊖-+-pos
+ )
+ renaming
+ ( +-comm to comm
+ ; +-identityˡ to identityˡ
+ ; +-identityʳ to identityʳ
+ ; +-assoc to assoc
+ ; +-isSemigroup to isSemigroup
+ ; +-0-isCommutativeMonoid to isCommutativeMonoid
+ ; +-0-commutativeMonoid to commutativeMonoid
+ )
diff --git a/src/Data/Integer/Base.agda b/src/Data/Integer/Base.agda
index 8894f16..5439c63 100644
--- a/src/Data/Integer/Base.agda
+++ b/src/Data/Integer/Base.agda
@@ -18,7 +18,7 @@ open import Relation.Binary.Core using (_≡_; refl)
infix 8 -_
infixl 7 _*_ _⊓_
infixl 6 _+_ _-_ _⊖_ _⊔_
-infix 4 _≤_
+infix 4 _≤_ _≥_ _<_ _>_ _≰_ _≱_ _≮_ _≯_
------------------------------------------------------------------------
-- The types
@@ -146,8 +146,29 @@ _⊓_ : ℤ → ℤ → ℤ
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
+ -≤- : ∀ {m n} → (n≤m : n ℕ.≤ m) → -[1+ m ] ≤ -[1+ n ]
+ +≤+ : ∀ {m n} → (m≤n : m ℕ.≤ n) → + m ≤ + n
+
+_≥_ : Rel ℤ _
+x ≥ y = y ≤ x
+
+_<_ : Rel ℤ _
+x < y = suc x ≤ y
+
+_>_ : Rel ℤ _
+x > y = y < x
+
+_≰_ : Rel ℤ _
+x ≰ y = ¬ (x ≤ y)
+
+_≱_ : Rel ℤ _
+x ≱ y = ¬ (x ≥ y)
+
+_≮_ : Rel ℤ _
+x ≮ y = ¬ (x < y)
+
+_≯_ : Rel ℤ _
+x ≯ y = ¬ (x > y)
drop‿+≤+ : ∀ {m n} → + m ≤ + n → ℕ._≤_ m n
drop‿+≤+ (+≤+ m≤n) = m≤n
diff --git a/src/Data/Integer/Multiplication/Properties.agda b/src/Data/Integer/Multiplication/Properties.agda
index 9fe17ef..f2bfc9a 100644
--- a/src/Data/Integer/Multiplication/Properties.agda
+++ b/src/Data/Integer/Multiplication/Properties.agda
@@ -2,90 +2,20 @@
-- The Agda standard library
--
-- Properties related to multiplication of integers
+--
+-- This module is DEPRECATED. Please use the corresponding properties in
+-- Data.Integer.Properties directly.
------------------------------------------------------------------------
module Data.Integer.Multiplication.Properties where
-open import Algebra
- using (module CommutativeSemiring; CommutativeMonoid)
-import Algebra.FunctionProperties
-open import Algebra.Structures using (IsSemigroup; IsCommutativeMonoid)
-open import Data.Integer
- using (ℤ; -[1+_]; +_; ∣_∣; sign; _◃_) renaming (_*_ to ℤ*)
-open import Data.Nat
- using (suc; zero) renaming (_+_ to _ℕ+_; _*_ to _ℕ*_)
-open import Data.Product using (proj₂)
-import Data.Nat.Properties as ℕ
-open import Data.Sign using () renaming (_*_ to _S*_)
-open import Function using (_∘_)
-open import Relation.Binary.PropositionalEquality
- using (_≡_; refl; cong; cong₂; isEquivalence)
-
-open Algebra.FunctionProperties (_≡_ {A = ℤ})
-open CommutativeSemiring ℕ.commutativeSemiring
- using (+-identity; *-comm) renaming (zero to *-zero)
-
-------------------------------------------------------------------------
--- Multiplication and one form a commutative monoid
-
-private
- lemma : ∀ a b c → c ℕ+ (b ℕ+ a ℕ* suc b) ℕ* suc c
- ≡ c ℕ+ b ℕ* suc c ℕ+ a ℕ* suc (c ℕ+ b ℕ* suc c)
- lemma =
- solve 3 (λ a b c → c :+ (b :+ a :* (con 1 :+ b)) :* (con 1 :+ c)
- := c :+ b :* (con 1 :+ c) :+
- a :* (con 1 :+ (c :+ b :* (con 1 :+ c))))
- refl
- where open ℕ.SemiringSolver
-
-
-identityˡ : LeftIdentity (+ 1) ℤ*
-identityˡ (+ zero ) = refl
-identityˡ -[1+ n ] rewrite proj₂ +-identity n = refl
-identityˡ (+ suc n) rewrite proj₂ +-identity n = refl
-
-comm : Commutative ℤ*
-comm -[1+ a ] -[1+ b ] rewrite *-comm (suc a) (suc b) = refl
-comm -[1+ a ] (+ b ) rewrite *-comm (suc a) b = refl
-comm (+ a ) -[1+ b ] rewrite *-comm a (suc b) = refl
-comm (+ a ) (+ b ) rewrite *-comm a b = refl
-
-assoc : Associative ℤ*
-assoc (+ zero) _ _ = refl
-assoc x (+ zero) _ rewrite proj₂ *-zero ∣ x ∣ = refl
-assoc x y (+ zero) rewrite
- proj₂ *-zero ∣ y ∣
- | proj₂ *-zero ∣ x ∣
- | proj₂ *-zero ∣ sign x S* sign y ◃ ∣ x ∣ ℕ* ∣ y ∣ ∣
- = refl
-assoc -[1+ a ] -[1+ b ] (+ suc c) = cong (+_ ∘ suc) (lemma a b c)
-assoc -[1+ a ] (+ suc b) -[1+ c ] = cong (+_ ∘ suc) (lemma a b c)
-assoc (+ suc a) (+ suc b) (+ suc c) = cong (+_ ∘ suc) (lemma a b c)
-assoc (+ suc a) -[1+ b ] -[1+ c ] = cong (+_ ∘ suc) (lemma a b c)
-assoc -[1+ a ] -[1+ b ] -[1+ c ] = cong -[1+_] (lemma a b c)
-assoc -[1+ a ] (+ suc b) (+ suc c) = cong -[1+_] (lemma a b c)
-assoc (+ suc a) -[1+ b ] (+ suc c) = cong -[1+_] (lemma a b c)
-assoc (+ suc a) (+ suc b) -[1+ c ] = cong -[1+_] (lemma a b c)
-
-isSemigroup : IsSemigroup _ _
-isSemigroup = record
- { isEquivalence = isEquivalence
- ; assoc = assoc
- ; ∙-cong = cong₂ ℤ*
- }
-
-isCommutativeMonoid : IsCommutativeMonoid _≡_ ℤ* (+ 1)
-isCommutativeMonoid = record
- { isSemigroup = isSemigroup
- ; identityˡ = identityˡ
- ; comm = comm
- }
-
-commutativeMonoid : CommutativeMonoid _ _
-commutativeMonoid = record
- { Carrier = ℤ
- ; _≈_ = _≡_
- ; _∙_ = ℤ*
- ; ε = + 1
- ; isCommutativeMonoid = isCommutativeMonoid
- }
+open import Data.Integer.Properties public
+ using ()
+ renaming
+ ( *-comm to comm
+ ; *-identityˡ to identityˡ
+ ; *-assoc to assoc
+ ; *-isSemigroup to isSemigroup
+ ; *-1-isCommutativeMonoid to isCommutativeMonoid
+ ; *-1-commutativeMonoid to commutativeMonoid
+ )
diff --git a/src/Data/Integer/Properties.agda b/src/Data/Integer/Properties.agda
index 718e9f0..543e2b7 100644
--- a/src/Data/Integer/Properties.agda
+++ b/src/Data/Integer/Properties.agda
@@ -8,20 +8,21 @@ module Data.Integer.Properties where
open import Algebra
import Algebra.FunctionProperties
+import Algebra.FunctionProperties.Consequences
import Algebra.Morphism as Morphism
import Algebra.Properties.AbelianGroup
open import Algebra.Structures
-open import Data.Integer hiding (suc; _≤?_)
-import Data.Integer.Addition.Properties as Add
-import Data.Integer.Multiplication.Properties as Mul
+open import Data.Integer renaming (suc to sucℤ)
open import Data.Nat
- using (ℕ; suc; zero; _∸_; _≤?_; _<_; _≥_; _≱_; s≤s; z≤n; ≤-pred)
+ using (ℕ; suc; zero; _∸_; s≤s; z≤n; ≤-pred)
hiding (module ℕ)
- renaming (_+_ to _ℕ+_; _*_ to _ℕ*_)
-open import Data.Nat.Properties as ℕ using (_*-mono_)
+ renaming (_+_ to _ℕ+_; _*_ to _ℕ*_;
+ _<_ to _ℕ<_; _≥_ to _ℕ≥_; _≰_ to _ℕ≰_; _≤?_ to _ℕ≤?_)
+import Data.Nat.Properties as ℕₚ
open import Data.Product using (proj₁; proj₂; _,_)
-open import Data.Sign as Sign using () renaming (_*_ to _S*_)
-import Data.Sign.Properties as SignProp
+open import Data.Sum using (inj₁; inj₂)
+open import Data.Sign as Sign using () renaming (_*_ to _𝕊*_)
+import Data.Sign.Properties as 𝕊ₚ
open import Function using (_∘_; _$_)
open import Relation.Binary
open import Relation.Binary.PropositionalEquality
@@ -29,35 +30,74 @@ open import Relation.Nullary using (yes; no)
open import Relation.Nullary.Negation using (contradiction)
open Algebra.FunctionProperties (_≡_ {A = ℤ})
-open CommutativeMonoid Add.commutativeMonoid
- using ()
- renaming ( assoc to +-assoc; comm to +-comm; identity to +-identity
- ; isCommutativeMonoid to +-isCommutativeMonoid
- ; isMonoid to +-isMonoid
- )
-open CommutativeMonoid Mul.commutativeMonoid
- using ()
- renaming ( assoc to *-assoc; comm to *-comm; identity to *-identity
- ; isCommutativeMonoid to *-isCommutativeMonoid
- ; isMonoid to *-isMonoid
- )
-open CommutativeSemiring ℕ.commutativeSemiring
- using () renaming (zero to *-zero; distrib to *-distrib)
-open DecTotalOrder Data.Nat.decTotalOrder
- using () renaming (refl to ≤-refl)
+open Algebra.FunctionProperties.Consequences (setoid ℤ)
open Morphism.Definitions ℤ ℕ _≡_
-open ℕ.SemiringSolver
+open ℕₚ.SemiringSolver
open ≡-Reasoning
------------------------------------------------------------------------
--- Miscellaneous properties
+-- Equality
--- Some properties relating sign and ∣_∣ to _◃_.
++-injective : ∀ {m n} → + m ≡ + n → m ≡ n
++-injective refl = refl
+
+-[1+-injective : ∀ {m n} → -[1+ m ] ≡ -[1+ n ] → m ≡ n
+-[1+-injective refl = refl
+
+≡-decSetoid : DecSetoid _ _
+≡-decSetoid = decSetoid _≟_
+
+------------------------------------------------------------------------
+-- Properties of -_
+
+doubleNeg : ∀ n → - - n ≡ n
+doubleNeg (+ zero) = refl
+doubleNeg (+ suc n) = refl
+doubleNeg (-[1+ n ]) = refl
+
+neg-injective : ∀ {m n} → - m ≡ - n → m ≡ n
+neg-injective {m} {n} -m≡-n = begin
+ m ≡⟨ sym (doubleNeg m) ⟩
+ - - m ≡⟨ cong -_ -m≡-n ⟩
+ - - n ≡⟨ doubleNeg n ⟩
+ n ∎
+
+------------------------------------------------------------------------
+-- Properties of ∣_∣
+
+∣n∣≡0⇒n≡0 : ∀ {n} → ∣ n ∣ ≡ 0 → n ≡ + 0
+∣n∣≡0⇒n≡0 {+ .zero} refl = refl
+∣n∣≡0⇒n≡0 { -[1+ n ]} ()
+
+∣-n∣≡∣n∣ : ∀ n → ∣ - n ∣ ≡ ∣ n ∣
+∣-n∣≡∣n∣ (+ zero) = refl
+∣-n∣≡∣n∣ (+ suc n) = refl
+∣-n∣≡∣n∣ (-[1+ n ]) = refl
+
+------------------------------------------------------------------------
+-- Properties of sign and _◃_
+
++◃n≡+n : ∀ n → Sign.+ ◃ n ≡ + n
++◃n≡+n zero = refl
++◃n≡+n (suc _) = refl
+
+-◃n≡-n : ∀ n → Sign.- ◃ n ≡ - + n
+-◃n≡-n zero = refl
+-◃n≡-n (suc _) = refl
sign-◃ : ∀ s n → sign (s ◃ suc n) ≡ s
sign-◃ Sign.- _ = refl
sign-◃ Sign.+ _ = refl
+abs-◃ : ∀ s n → ∣ s ◃ n ∣ ≡ n
+abs-◃ _ zero = refl
+abs-◃ Sign.- (suc n) = refl
+abs-◃ Sign.+ (suc n) = refl
+
+signₙ◃∣n∣≡n : ∀ n → sign n ◃ ∣ n ∣ ≡ n
+signₙ◃∣n∣≡n (+ n) = +◃n≡+n n
+signₙ◃∣n∣≡n (-[1+ n ]) = refl
+
sign-cong : ∀ {s₁ s₂ n₁ n₂} →
s₁ ◃ suc n₁ ≡ s₂ ◃ suc n₂ → s₁ ≡ s₂
sign-cong {s₁} {s₂} {n₁} {n₂} eq = begin
@@ -66,11 +106,6 @@ sign-cong {s₁} {s₂} {n₁} {n₂} eq = begin
sign (s₂ ◃ suc n₂) ≡⟨ sign-◃ s₂ n₂ ⟩
s₂ ∎
-abs-◃ : ∀ s n → ∣ s ◃ n ∣ ≡ n
-abs-◃ _ zero = refl
-abs-◃ Sign.- (suc n) = refl
-abs-◃ Sign.+ (suc n) = refl
-
abs-cong : ∀ {s₁ s₂ n₁ n₂} →
s₁ ◃ n₁ ≡ s₂ ◃ n₂ → n₁ ≡ n₂
abs-cong {s₁} {s₂} {n₁} {n₂} eq = begin
@@ -79,48 +114,132 @@ abs-cong {s₁} {s₂} {n₁} {n₂} eq = begin
∣ s₂ ◃ n₂ ∣ ≡⟨ abs-◃ s₂ n₂ ⟩
n₂ ∎
--- ∣_∣ commutes with multiplication.
+∣s◃m∣*∣t◃n∣≡m*n : ∀ s t m n → ∣ s ◃ m ∣ ℕ* ∣ t ◃ n ∣ ≡ m ℕ* n
+∣s◃m∣*∣t◃n∣≡m*n s t m n = cong₂ _ℕ*_ (abs-◃ s m) (abs-◃ t n)
-abs-*-commute : Homomorphic₂ ∣_∣ _*_ _ℕ*_
-abs-*-commute i j = abs-◃ _ _
-
--- Subtraction properties
+------------------------------------------------------------------------
+-- Properties of _⊖_
n⊖n≡0 : ∀ n → n ⊖ n ≡ + 0
n⊖n≡0 zero = refl
n⊖n≡0 (suc n) = n⊖n≡0 n
-sign-⊖-< : ∀ {m n} → m < n → sign (m ⊖ n) ≡ Sign.-
-sign-⊖-< {zero} (s≤s z≤n) = refl
-sign-⊖-< {suc n} (s≤s m<n) = sign-⊖-< m<n
-
-+-⊖-left-cancel : ∀ a b c → (a ℕ+ b) ⊖ (a ℕ+ c) ≡ b ⊖ c
-+-⊖-left-cancel zero b c = refl
-+-⊖-left-cancel (suc a) b c = +-⊖-left-cancel a b c
-
⊖-swap : ∀ a b → a ⊖ b ≡ - (b ⊖ a)
⊖-swap zero zero = refl
⊖-swap (suc _) zero = refl
⊖-swap zero (suc _) = refl
⊖-swap (suc a) (suc b) = ⊖-swap a b
-⊖-≥ : ∀ {m n} → m ≥ n → m ⊖ n ≡ + (m ∸ n)
+⊖-≥ : ∀ {m n} → m ℕ≥ n → m ⊖ n ≡ + (m ∸ n)
⊖-≥ z≤n = refl
⊖-≥ (s≤s n≤m) = ⊖-≥ n≤m
-⊖-< : ∀ {m n} → m < n → m ⊖ n ≡ - + (n ∸ m)
+⊖-< : ∀ {m n} → m ℕ< n → m ⊖ n ≡ - + (n ∸ m)
⊖-< {zero} (s≤s z≤n) = refl
⊖-< {suc m} (s≤s m<n) = ⊖-< m<n
-∣⊖∣-< : ∀ {m n} → m < n → ∣ m ⊖ n ∣ ≡ n ∸ m
+⊖-≰ : ∀ {m n} → n ℕ≰ m → m ⊖ n ≡ - + (n ∸ m)
+⊖-≰ = ⊖-< ∘ ℕₚ.≰⇒>
+
+∣⊖∣-< : ∀ {m n} → m ℕ< n → ∣ m ⊖ n ∣ ≡ n ∸ m
∣⊖∣-< {zero} (s≤s z≤n) = refl
∣⊖∣-< {suc n} (s≤s m<n) = ∣⊖∣-< m<n
+∣⊖∣-≰ : ∀ {m n} → n ℕ≰ m → ∣ m ⊖ n ∣ ≡ n ∸ m
+∣⊖∣-≰ = ∣⊖∣-< ∘ ℕₚ.≰⇒>
+
+sign-⊖-< : ∀ {m n} → m ℕ< n → sign (m ⊖ n) ≡ Sign.-
+sign-⊖-< {zero} (s≤s z≤n) = refl
+sign-⊖-< {suc n} (s≤s m<n) = sign-⊖-< m<n
+
+sign-⊖-≰ : ∀ {m n} → n ℕ≰ m → sign (m ⊖ n) ≡ Sign.-
+sign-⊖-≰ = sign-⊖-< ∘ ℕₚ.≰⇒>
+
+-[n⊖m]≡-m+n : ∀ m n → - (m ⊖ n) ≡ (- (+ m)) + (+ n)
+-[n⊖m]≡-m+n zero zero = refl
+-[n⊖m]≡-m+n zero (suc n) = refl
+-[n⊖m]≡-m+n (suc m) zero = refl
+-[n⊖m]≡-m+n (suc m) (suc n) = sym (⊖-swap n m)
+
++-⊖-left-cancel : ∀ a b c → (a ℕ+ b) ⊖ (a ℕ+ c) ≡ b ⊖ c
++-⊖-left-cancel zero b c = refl
++-⊖-left-cancel (suc a) b c = +-⊖-left-cancel a b c
------------------------------------------------------------------------
--- The integers form a commutative ring
+-- Properties of _+_
+
++-comm : Commutative _+_
++-comm -[1+ a ] -[1+ b ] rewrite ℕₚ.+-comm a b = refl
++-comm (+ a ) (+ b ) rewrite ℕₚ.+-comm a b = refl
++-comm -[1+ _ ] (+ _ ) = refl
++-comm (+ _ ) -[1+ _ ] = refl
+
++-identityˡ : LeftIdentity (+ 0) _+_
++-identityˡ -[1+ _ ] = refl
++-identityˡ (+ _ ) = refl
+
++-identityʳ : RightIdentity (+ 0) _+_
++-identityʳ = comm+idˡ⇒idʳ +-comm +-identityˡ
+
++-identity : Identity (+ 0) _+_
++-identity = +-identityˡ , +-identityʳ
+
+distribˡ-⊖-+-neg : ∀ a b c → b ⊖ c + -[1+ a ] ≡ b ⊖ (suc c ℕ+ a)
+distribˡ-⊖-+-neg _ zero zero = refl
+distribˡ-⊖-+-neg _ zero (suc _) = refl
+distribˡ-⊖-+-neg _ (suc _) zero = refl
+distribˡ-⊖-+-neg a (suc b) (suc c) = distribˡ-⊖-+-neg a b c
+
+distribʳ-⊖-+-neg : ∀ a b c → -[1+ a ] + (b ⊖ c) ≡ b ⊖ (suc a ℕ+ c)
+distribʳ-⊖-+-neg a b c
+ rewrite +-comm -[1+ a ] (b ⊖ c)
+ | distribˡ-⊖-+-neg a b c
+ | ℕₚ.+-comm a c
+ = refl
--- Additive abelian group.
+distribˡ-⊖-+-pos : ∀ a b c → b ⊖ c + + a ≡ b ℕ+ a ⊖ c
+distribˡ-⊖-+-pos _ zero zero = refl
+distribˡ-⊖-+-pos _ zero (suc _) = refl
+distribˡ-⊖-+-pos _ (suc _) zero = refl
+distribˡ-⊖-+-pos a (suc b) (suc c) = distribˡ-⊖-+-pos a b c
+
+distribʳ-⊖-+-pos : ∀ a b c → + a + (b ⊖ c) ≡ a ℕ+ b ⊖ c
+distribʳ-⊖-+-pos a b c
+ rewrite +-comm (+ a) (b ⊖ c)
+ | distribˡ-⊖-+-pos a b c
+ | ℕₚ.+-comm a b
+ = refl
+
++-assoc : Associative _+_
++-assoc (+ zero) y z rewrite +-identityˡ y | +-identityˡ (y + z) = refl
++-assoc x (+ zero) z rewrite +-identityʳ x | +-identityˡ z = refl
++-assoc x y (+ zero) rewrite +-identityʳ (x + y) | +-identityʳ y = refl
++-assoc -[1+ a ] -[1+ b ] (+ suc c) = sym (distribʳ-⊖-+-neg a c b)
++-assoc -[1+ a ] (+ suc b) (+ suc c) = distribˡ-⊖-+-pos (suc c) b a
++-assoc (+ suc a) -[1+ b ] -[1+ c ] = distribˡ-⊖-+-neg c a b
++-assoc (+ suc a) -[1+ b ] (+ suc c)
+ rewrite distribˡ-⊖-+-pos (suc c) a b
+ | distribʳ-⊖-+-pos (suc a) c b
+ | sym (ℕₚ.+-assoc a 1 c)
+ | ℕₚ.+-comm a 1
+ = refl
++-assoc (+ suc a) (+ suc b) -[1+ c ]
+ rewrite distribʳ-⊖-+-pos (suc a) b c
+ | sym (ℕₚ.+-assoc a 1 b)
+ | ℕₚ.+-comm a 1
+ = refl
++-assoc -[1+ a ] -[1+ b ] -[1+ c ]
+ rewrite sym (ℕₚ.+-assoc a 1 (b ℕ+ c))
+ | ℕₚ.+-comm a 1
+ | ℕₚ.+-assoc a b c
+ = refl
++-assoc -[1+ a ] (+ suc b) -[1+ c ]
+ rewrite distribʳ-⊖-+-neg a b c
+ | distribˡ-⊖-+-neg c b a
+ = refl
++-assoc (+ suc a) (+ suc b) (+ suc c)
+ rewrite ℕₚ.+-assoc (suc a) (suc b) (suc c)
+ = refl
inverseˡ : LeftInverse (+ 0) -_ _+_
inverseˡ -[1+ n ] = n⊖n≡0 n
@@ -128,89 +247,232 @@ inverseˡ (+ zero) = refl
inverseˡ (+ suc n) = n⊖n≡0 n
inverseʳ : RightInverse (+ 0) -_ _+_
-inverseʳ i = begin
- i + - i ≡⟨ +-comm i (- i) ⟩
- - i + i ≡⟨ inverseˡ i ⟩
- + 0 ∎
+inverseʳ = comm+invˡ⇒invʳ +-comm inverseˡ
+
++-inverse : Inverse (+ 0) -_ _+_
++-inverse = inverseˡ , inverseʳ
+
++-isSemigroup : IsSemigroup _≡_ _+_
++-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = +-assoc
+ ; ∙-cong = cong₂ _+_
+ }
+
++-0-isMonoid : IsMonoid _≡_ _+_ (+ 0)
++-0-isMonoid = record
+ { isSemigroup = +-isSemigroup
+ ; identity = +-identity
+ }
+
++-0-isCommutativeMonoid : IsCommutativeMonoid _≡_ _+_ (+ 0)
++-0-isCommutativeMonoid = record
+ { isSemigroup = +-isSemigroup
+ ; identityˡ = +-identityˡ
+ ; comm = +-comm
+ }
+
++-0-commutativeMonoid : CommutativeMonoid _ _
++-0-commutativeMonoid = record
+ { Carrier = ℤ
+ ; _≈_ = _≡_
+ ; _∙_ = _+_
+ ; ε = + 0
+ ; isCommutativeMonoid = +-0-isCommutativeMonoid
+ }
+
++-0-isGroup : IsGroup _≡_ _+_ (+ 0) (-_)
++-0-isGroup = record
+ { isMonoid = +-0-isMonoid
+ ; inverse = +-inverse
+ ; ⁻¹-cong = cong (-_)
+ }
+-isAbelianGroup : IsAbelianGroup _≡_ _+_ (+ 0) (-_)
+-isAbelianGroup = record
- { isGroup = record
- { isMonoid = +-isMonoid
- ; inverse = inverseˡ , inverseʳ
- ; ⁻¹-cong = cong (-_)
- }
- ; comm = +-comm
+ { isGroup = +-0-isGroup
+ ; comm = +-comm
}
-open Algebra.Properties.AbelianGroup
- (record { isAbelianGroup = +-isAbelianGroup })
++-0-abelianGroup : AbelianGroup _ _
++-0-abelianGroup = record
+ { Carrier = ℤ
+ ; _≈_ = _≡_
+ ; _∙_ = _+_
+ ; ε = + 0
+ ; _⁻¹ = -_
+ ; isAbelianGroup = +-isAbelianGroup
+ }
+
+open Algebra.Properties.AbelianGroup +-0-abelianGroup
using () renaming (⁻¹-involutive to -‿involutive)
+-- Other properties of _+_
+
+n≢1+n : ∀ {n} → n ≢ sucℤ n
+n≢1+n {+ _} ()
+n≢1+n { -[1+ 0 ]} ()
+n≢1+n { -[1+ suc n ]} ()
+
+1-[1+n]≡-n : ∀ n → sucℤ -[1+ n ] ≡ - (+ n)
+1-[1+n]≡-n zero = refl
+1-[1+n]≡-n (suc n) = refl
+
+neg-distrib-+ : ∀ m n → - (m + n) ≡ (- m) + (- n)
+neg-distrib-+ (+ zero) (+ zero) = refl
+neg-distrib-+ (+ zero) (+ suc n) = refl
+neg-distrib-+ (+ suc m) (+ zero) = cong -[1+_] (ℕₚ.+-identityʳ m)
+neg-distrib-+ (+ suc m) (+ suc n) = cong -[1+_] (ℕₚ.+-suc m n)
+neg-distrib-+ -[1+ m ] -[1+ n ] = cong (λ v → + suc v) (sym (ℕₚ.+-suc m n))
+neg-distrib-+ (+ m) -[1+ n ] = -[n⊖m]≡-m+n m (suc n)
+neg-distrib-+ -[1+ m ] (+ n) =
+ trans (-[n⊖m]≡-m+n n (suc m)) (+-comm (- + n) (+ suc m))
+
+◃-distrib-+ : ∀ s m n → s ◃ (m ℕ+ n) ≡ (s ◃ m) + (s ◃ n)
+◃-distrib-+ Sign.- m n = begin
+ Sign.- ◃ (m ℕ+ n) ≡⟨ -◃n≡-n (m ℕ+ n) ⟩
+ - (+ (m ℕ+ n)) ≡⟨⟩
+ - ((+ m) + (+ n)) ≡⟨ neg-distrib-+ (+ m) (+ n) ⟩
+ (- (+ m)) + (- (+ n)) ≡⟨ sym (cong₂ _+_ (-◃n≡-n m) (-◃n≡-n n)) ⟩
+ (Sign.- ◃ m) + (Sign.- ◃ n) ∎
+◃-distrib-+ Sign.+ m n = begin
+ Sign.+ ◃ (m ℕ+ n) ≡⟨ +◃n≡+n (m ℕ+ n) ⟩
+ + (m ℕ+ n) ≡⟨⟩
+ (+ m) + (+ n) ≡⟨ sym (cong₂ _+_ (+◃n≡+n m) (+◃n≡+n n)) ⟩
+ (Sign.+ ◃ m) + (Sign.+ ◃ n) ∎
--- Distributivity
+
+------------------------------------------------------------------------
+-- Properties of _*_
+
+*-comm : Commutative _*_
+*-comm -[1+ a ] -[1+ b ] rewrite ℕₚ.*-comm (suc a) (suc b) = refl
+*-comm -[1+ a ] (+ b ) rewrite ℕₚ.*-comm (suc a) b = refl
+*-comm (+ a ) -[1+ b ] rewrite ℕₚ.*-comm a (suc b) = refl
+*-comm (+ a ) (+ b ) rewrite ℕₚ.*-comm a b = refl
+
+*-identityˡ : LeftIdentity (+ 1) _*_
+*-identityˡ (+ zero ) = refl
+*-identityˡ -[1+ n ] rewrite ℕₚ.+-identityʳ n = refl
+*-identityˡ (+ suc n) rewrite ℕₚ.+-identityʳ n = refl
+
+*-identityʳ : RightIdentity (+ 1) _*_
+*-identityʳ = comm+idˡ⇒idʳ *-comm *-identityˡ
+
+*-identity : Identity (+ 1) _*_
+*-identity = *-identityˡ , *-identityʳ
+
+*-zeroˡ : LeftZero (+ 0) _*_
+*-zeroˡ n = refl
+
+*-zeroʳ : RightZero (+ 0) _*_
+*-zeroʳ = comm+zeˡ⇒zeʳ *-comm *-zeroˡ
+
+*-zero : Zero (+ 0) _*_
+*-zero = *-zeroˡ , *-zeroʳ
private
+ lemma : ∀ a b c → c ℕ+ (b ℕ+ a ℕ* suc b) ℕ* suc c
+ ≡ c ℕ+ b ℕ* suc c ℕ+ a ℕ* suc (c ℕ+ b ℕ* suc c)
+ lemma =
+ solve 3 (λ a b c → c :+ (b :+ a :* (con 1 :+ b)) :* (con 1 :+ c)
+ := c :+ b :* (con 1 :+ c) :+
+ a :* (con 1 :+ (c :+ b :* (con 1 :+ c))))
+ refl
+ where open ℕₚ.SemiringSolver
+
+*-assoc : Associative _*_
+*-assoc (+ zero) _ _ = refl
+*-assoc x (+ zero) _ rewrite ℕₚ.*-zeroʳ ∣ x ∣ = refl
+*-assoc x y (+ zero) rewrite
+ ℕₚ.*-zeroʳ ∣ y ∣
+ | ℕₚ.*-zeroʳ ∣ x ∣
+ | ℕₚ.*-zeroʳ ∣ sign x 𝕊* sign y ◃ ∣ x ∣ ℕ* ∣ y ∣ ∣
+ = refl
+*-assoc -[1+ a ] -[1+ b ] (+ suc c) = cong (+_ ∘ suc) (lemma a b c)
+*-assoc -[1+ a ] (+ suc b) -[1+ c ] = cong (+_ ∘ suc) (lemma a b c)
+*-assoc (+ suc a) (+ suc b) (+ suc c) = cong (+_ ∘ suc) (lemma a b c)
+*-assoc (+ suc a) -[1+ b ] -[1+ c ] = cong (+_ ∘ suc) (lemma a b c)
+*-assoc -[1+ a ] -[1+ b ] -[1+ c ] = cong -[1+_] (lemma a b c)
+*-assoc -[1+ a ] (+ suc b) (+ suc c) = cong -[1+_] (lemma a b c)
+*-assoc (+ suc a) -[1+ b ] (+ suc c) = cong -[1+_] (lemma a b c)
+*-assoc (+ suc a) (+ suc b) -[1+ c ] = cong -[1+_] (lemma a b c)
+
+*-isSemigroup : IsSemigroup _ _
+*-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = *-assoc
+ ; ∙-cong = cong₂ _*_
+ }
- -- Various lemmas used to prove distributivity.
+*-1-isMonoid : IsMonoid _≡_ _*_ (+ 1)
+*-1-isMonoid = record
+ { isSemigroup = *-isSemigroup
+ ; identity = *-identity
+ }
- sign-⊖-≱ : ∀ {m n} → m ≱ n → sign (m ⊖ n) ≡ Sign.-
- sign-⊖-≱ = sign-⊖-< ∘ ℕ.≰⇒>
+*-1-isCommutativeMonoid : IsCommutativeMonoid _≡_ _*_ (+ 1)
+*-1-isCommutativeMonoid = record
+ { isSemigroup = *-isSemigroup
+ ; identityˡ = *-identityˡ
+ ; comm = *-comm
+ }
- ∣⊖∣-≱ : ∀ {m n} → m ≱ n → ∣ m ⊖ n ∣ ≡ n ∸ m
- ∣⊖∣-≱ = ∣⊖∣-< ∘ ℕ.≰⇒>
+*-1-commutativeMonoid : CommutativeMonoid _ _
+*-1-commutativeMonoid = record
+ { Carrier = ℤ
+ ; _≈_ = _≡_
+ ; _∙_ = _*_
+ ; ε = + 1
+ ; isCommutativeMonoid = *-1-isCommutativeMonoid
+ }
- ⊖-≱ : ∀ {m n} → m ≱ n → m ⊖ n ≡ - + (n ∸ m)
- ⊖-≱ = ⊖-< ∘ ℕ.≰⇒>
+------------------------------------------------------------------------
+-- The integers form a commutative ring
- -- Lemmas working around the fact that _◃_ pattern matches on its
- -- second argument before its first.
+-- Distributivity
- +‿◃ : ∀ n → Sign.+ ◃ n ≡ + n
- +‿◃ zero = refl
- +‿◃ (suc _) = refl
+private
- -‿◃ : ∀ n → Sign.- ◃ n ≡ - + n
- -‿◃ zero = refl
- -‿◃ (suc _) = refl
+ -- lemma used to prove distributivity.
distrib-lemma :
∀ a b c → (c ⊖ b) * -[1+ a ] ≡ a ℕ+ b ℕ* suc a ⊖ (a ℕ+ c ℕ* suc a)
distrib-lemma a b c
rewrite +-⊖-left-cancel a (b ℕ* suc a) (c ℕ* suc a)
| ⊖-swap (b ℕ* suc a) (c ℕ* suc a)
- with b ≤? c
+ with b ℕ≤? c
... | yes b≤c
rewrite ⊖-≥ b≤c
- | ⊖-≥ (b≤c *-mono (≤-refl {x = suc a}))
- | -‿◃ ((c ∸ b) ℕ* suc a)
- | ℕ.*-distrib-∸ʳ (suc a) c b
+ | ⊖-≥ (ℕₚ.*-mono-≤ b≤c (ℕₚ.≤-refl {x = suc a}))
+ | -◃n≡-n ((c ∸ b) ℕ* suc a)
+ | ℕₚ.*-distribʳ-∸ (suc a) c b
= refl
... | no b≰c
- rewrite sign-⊖-≱ b≰c
- | ∣⊖∣-≱ b≰c
- | +‿◃ ((b ∸ c) ℕ* suc a)
- | ⊖-≱ (b≰c ∘ ℕ.cancel-*-right-≤ b c a)
+ rewrite sign-⊖-≰ b≰c
+ | ∣⊖∣-≰ b≰c
+ | +◃n≡+n ((b ∸ c) ℕ* suc a)
+ | ⊖-≰ (b≰c ∘ ℕₚ.*-cancelʳ-≤ b c a)
| -‿involutive (+ (b ℕ* suc a ∸ c ℕ* suc a))
- | ℕ.*-distrib-∸ʳ (suc a) b c
+ | ℕₚ.*-distribʳ-∸ (suc a) b c
= refl
distribʳ : _*_ DistributesOverʳ _+_
distribʳ (+ zero) y z
- rewrite proj₂ *-zero ∣ y ∣
- | proj₂ *-zero ∣ z ∣
- | proj₂ *-zero ∣ y + z ∣
+ rewrite ℕₚ.*-zeroʳ ∣ y ∣
+ | ℕₚ.*-zeroʳ ∣ z ∣
+ | ℕₚ.*-zeroʳ ∣ y + z ∣
= refl
distribʳ x (+ zero) z
- rewrite proj₁ +-identity z
- | proj₁ +-identity (sign z S* sign x ◃ ∣ z ∣ ℕ* ∣ x ∣)
+ rewrite +-identityˡ z
+ | +-identityˡ (sign z 𝕊* sign x ◃ ∣ z ∣ ℕ* ∣ x ∣)
= refl
distribʳ x y (+ zero)
- rewrite proj₂ +-identity y
- | proj₂ +-identity (sign y S* sign x ◃ ∣ y ∣ ℕ* ∣ x ∣)
+ rewrite +-identityʳ y
+ | +-identityʳ (sign y 𝕊* sign x ◃ ∣ y ∣ ℕ* ∣ x ∣)
= refl
distribʳ -[1+ a ] -[1+ b ] -[1+ c ] = cong (+_) $
@@ -243,50 +505,61 @@ distribʳ -[1+ a ] (+ suc b) -[1+ c ] = distrib-lemma a c b
distribʳ (+ suc a) -[1+ b ] (+ suc c)
rewrite +-⊖-left-cancel a (c ℕ* suc a) (b ℕ* suc a)
- with b ≤? c
+ with b ℕ≤? c
... | yes b≤c
rewrite ⊖-≥ b≤c
| +-comm (- (+ (a ℕ+ b ℕ* suc a))) (+ (a ℕ+ c ℕ* suc a))
- | ⊖-≥ (b≤c *-mono ≤-refl {x = suc a})
- | ℕ.*-distrib-∸ʳ (suc a) c b
- | +‿◃ (c ℕ* suc a ∸ b ℕ* suc a)
+ | ⊖-≥ (ℕₚ.*-mono-≤ b≤c (ℕₚ.≤-refl {x = suc a}))
+ | ℕₚ.*-distribʳ-∸ (suc a) c b
+ | +◃n≡+n (c ℕ* suc a ∸ b ℕ* suc a)
= refl
... | no b≰c
- rewrite sign-⊖-≱ b≰c
- | ∣⊖∣-≱ b≰c
- | -‿◃ ((b ∸ c) ℕ* suc a)
- | ⊖-≱ (b≰c ∘ ℕ.cancel-*-right-≤ b c a)
- | ℕ.*-distrib-∸ʳ (suc a) b c
+ rewrite sign-⊖-≰ b≰c
+ | ∣⊖∣-≰ b≰c
+ | -◃n≡-n ((b ∸ c) ℕ* suc a)
+ | ⊖-≰ (b≰c ∘ ℕₚ.*-cancelʳ-≤ b c a)
+ | ℕₚ.*-distribʳ-∸ (suc a) b c
= refl
distribʳ (+ suc c) (+ suc a) -[1+ b ]
rewrite +-⊖-left-cancel c (a ℕ* suc c) (b ℕ* suc c)
- with b ≤? a
+ with b ℕ≤? a
... | yes b≤a
rewrite ⊖-≥ b≤a
- | ⊖-≥ (b≤a *-mono ≤-refl {x = suc c})
- | +‿◃ ((a ∸ b) ℕ* suc c)
- | ℕ.*-distrib-∸ʳ (suc c) a b
+ | ⊖-≥ (ℕₚ.*-mono-≤ b≤a (ℕₚ.≤-refl {x = suc c}))
+ | +◃n≡+n ((a ∸ b) ℕ* suc c)
+ | ℕₚ.*-distribʳ-∸ (suc c) a b
= refl
... | no b≰a
- rewrite sign-⊖-≱ b≰a
- | ∣⊖∣-≱ b≰a
- | ⊖-≱ (b≰a ∘ ℕ.cancel-*-right-≤ b a c)
- | -‿◃ ((b ∸ a) ℕ* suc c)
- | ℕ.*-distrib-∸ʳ (suc c) b a
+ rewrite sign-⊖-≰ b≰a
+ | ∣⊖∣-≰ b≰a
+ | ⊖-≰ (b≰a ∘ ℕₚ.*-cancelʳ-≤ b a c)
+ | -◃n≡-n ((b ∸ a) ℕ* suc c)
+ | ℕₚ.*-distribʳ-∸ (suc c) b a
= refl
--- The IsCommutativeSemiring module contains a proof of
--- distributivity which is used below.
-
isCommutativeSemiring : IsCommutativeSemiring _≡_ _+_ _*_ (+ 0) (+ 1)
isCommutativeSemiring = record
- { +-isCommutativeMonoid = +-isCommutativeMonoid
- ; *-isCommutativeMonoid = *-isCommutativeMonoid
+ { +-isCommutativeMonoid = +-0-isCommutativeMonoid
+ ; *-isCommutativeMonoid = *-1-isCommutativeMonoid
; distribʳ = distribʳ
; zeroˡ = λ _ → refl
}
++-*-isRing : IsRing _≡_ _+_ _*_ -_ (+ 0) (+ 1)
++-*-isRing = record
+ { +-isAbelianGroup = +-isAbelianGroup
+ ; *-isMonoid = *-1-isMonoid
+ ; distrib = IsCommutativeSemiring.distrib
+ isCommutativeSemiring
+ }
+
++-*-isCommutativeRing : IsCommutativeRing _≡_ _+_ _*_ -_ (+ 0) (+ 1)
++-*-isCommutativeRing = record
+ { isRing = +-*-isRing
+ ; *-comm = *-comm
+ }
+
commutativeRing : CommutativeRing _ _
commutativeRing = record
{ Carrier = ℤ
@@ -296,15 +569,7 @@ commutativeRing = record
; -_ = -_
; 0# = + 0
; 1# = + 1
- ; isCommutativeRing = record
- { isRing = record
- { +-isAbelianGroup = +-isAbelianGroup
- ; *-isMonoid = *-isMonoid
- ; distrib = IsCommutativeSemiring.distrib
- isCommutativeSemiring
- }
- ; *-comm = *-comm
- }
+ ; isCommutativeRing = +-*-isCommutativeRing
}
import Algebra.RingSolver.Simple as Solver
@@ -312,36 +577,35 @@ import Algebra.RingSolver.AlmostCommutativeRing as ACR
module RingSolver =
Solver (ACR.fromCommutativeRing commutativeRing) _≟_
-------------------------------------------------------------------------
--- More properties
+-- Other properties of _*_
--- Multiplication is right cancellative for non-zero integers.
+abs-*-commute : Homomorphic₂ ∣_∣ _*_ _ℕ*_
+abs-*-commute i j = abs-◃ _ _
-cancel-*-right : ∀ i j k →
- k ≢ + 0 → i * k ≡ j * k → i ≡ j
+cancel-*-right : ∀ i j k → k ≢ + 0 → i * k ≡ j * k → i ≡ j
cancel-*-right i j k ≢0 eq with signAbs k
cancel-*-right i j .(+ 0) ≢0 eq | s ◂ zero = contradiction refl ≢0
cancel-*-right i j .(s ◃ suc n) ≢0 eq | s ◂ suc n
with ∣ s ◃ suc n ∣ | abs-◃ s (suc n) | sign (s ◃ suc n) | sign-◃ s n
... | .(suc n) | refl | .s | refl =
◃-cong (sign-i≡sign-j i j eq) $
- ℕ.cancel-*-right ∣ i ∣ ∣ j ∣ $ abs-cong eq
+ ℕₚ.*-cancelʳ-≡ ∣ i ∣ ∣ j ∣ $ abs-cong eq
where
sign-i≡sign-j : ∀ i j →
- sign i S* s ◃ ∣ i ∣ ℕ* suc n ≡
- sign j S* s ◃ ∣ j ∣ ℕ* suc n →
+ sign i 𝕊* s ◃ ∣ i ∣ ℕ* suc n ≡
+ sign j 𝕊* s ◃ ∣ j ∣ ℕ* suc n →
sign i ≡ sign j
sign-i≡sign-j i j eq with signAbs i | signAbs j
sign-i≡sign-j .(+ 0) .(+ 0) eq | s₁ ◂ zero | s₂ ◂ zero = refl
sign-i≡sign-j .(+ 0) .(s₂ ◃ suc n₂) eq | s₁ ◂ zero | s₂ ◂ suc n₂
with ∣ s₂ ◃ suc n₂ ∣ | abs-◃ s₂ (suc n₂)
... | .(suc n₂) | refl
- with abs-cong {s₁} {sign (s₂ ◃ suc n₂) S* s} {0} {suc n₂ ℕ* suc n} eq
+ with abs-cong {s₁} {sign (s₂ ◃ suc n₂) 𝕊* s} {0} {suc n₂ ℕ* suc n} eq
... | ()
sign-i≡sign-j .(s₁ ◃ suc n₁) .(+ 0) eq | s₁ ◂ suc n₁ | s₂ ◂ zero
with ∣ s₁ ◃ suc n₁ ∣ | abs-◃ s₁ (suc n₁)
... | .(suc n₁) | refl
- with abs-cong {sign (s₁ ◃ suc n₁) S* s} {s₁} {suc n₁ ℕ* suc n} {0} eq
+ with abs-cong {sign (s₁ ◃ suc n₁) 𝕊* s} {s₁} {suc n₁ ℕ* suc n} {0} eq
... | ()
sign-i≡sign-j .(s₁ ◃ suc n₁) .(s₂ ◃ suc n₂) eq | s₁ ◂ suc n₁ | s₂ ◂ suc n₂
with ∣ s₁ ◃ suc n₁ ∣ | abs-◃ s₁ (suc n₁)
@@ -349,14 +613,11 @@ cancel-*-right i j .(s ◃ suc n) ≢0 eq | s ◂ suc n
| ∣ s₂ ◃ suc n₂ ∣ | abs-◃ s₂ (suc n₂)
| sign (s₂ ◃ suc n₂) | sign-◃ s₂ n₂
... | .(suc n₁) | refl | .s₁ | refl | .(suc n₂) | refl | .s₂ | refl =
- SignProp.cancel-*-right s₁ s₂ (sign-cong eq)
-
--- Multiplication with a positive number is right cancellative (for
--- _≤_).
+ 𝕊ₚ.cancel-*-right s₁ s₂ (sign-cong eq)
cancel-*-+-right-≤ : ∀ m n o → m * + suc o ≤ n * + suc o → m ≤ n
cancel-*-+-right-≤ (-[1+ m ]) (-[1+ n ]) o (-≤- n≤m) =
- -≤- (≤-pred (ℕ.cancel-*-right-≤ (suc n) (suc m) o (s≤s n≤m)))
+ -≤- (≤-pred (ℕₚ.*-cancelʳ-≤ (suc n) (suc m) o (s≤s n≤m)))
cancel-*-+-right-≤ -[1+ _ ] (+ _) _ _ = -≤+
cancel-*-+-right-≤ (+ 0) -[1+ _ ] _ ()
cancel-*-+-right-≤ (+ suc _) -[1+ _ ] _ ()
@@ -364,17 +625,205 @@ cancel-*-+-right-≤ (+ 0) (+ 0) _ _ = +≤+ z≤n
cancel-*-+-right-≤ (+ 0) (+ suc _) _ _ = +≤+ z≤n
cancel-*-+-right-≤ (+ suc _) (+ 0) _ (+≤+ ())
cancel-*-+-right-≤ (+ suc m) (+ suc n) o (+≤+ m≤n) =
- +≤+ (ℕ.cancel-*-right-≤ (suc m) (suc n) o m≤n)
+ +≤+ (ℕₚ.*-cancelʳ-≤ (suc m) (suc n) o m≤n)
--- Multiplication with a positive number is monotone.
-
-*-+-right-mono : ∀ n → (λ x → x * + suc n) Preserves _≤_ ⟶ _≤_
+*-+-right-mono : ∀ n → (_* + suc n) Preserves _≤_ ⟶ _≤_
*-+-right-mono _ (-≤+ {n = 0}) = -≤+
*-+-right-mono _ (-≤+ {n = suc _}) = -≤+
*-+-right-mono x (-≤- n≤m) =
- -≤- (≤-pred (s≤s n≤m *-mono ≤-refl {x = suc x}))
+ -≤- (≤-pred (ℕₚ.*-mono-≤ (s≤s n≤m) (ℕₚ.≤-refl {x = suc x})))
*-+-right-mono _ (+≤+ {m = 0} {n = 0} m≤n) = +≤+ m≤n
*-+-right-mono _ (+≤+ {m = 0} {n = suc _} m≤n) = +≤+ z≤n
*-+-right-mono _ (+≤+ {m = suc _} {n = 0} ())
*-+-right-mono x (+≤+ {m = suc _} {n = suc _} m≤n) =
- +≤+ (m≤n *-mono ≤-refl {x = suc x})
+ +≤+ ((ℕₚ.*-mono-≤ m≤n (ℕₚ.≤-refl {x = suc x})))
+
+-1*n≡-n : ∀ n → -[1+ 0 ] * n ≡ - n
+-1*n≡-n (+ zero) = refl
+-1*n≡-n (+ suc n) = cong -[1+_] (ℕₚ.+-identityʳ n)
+-1*n≡-n -[1+ n ] = cong (λ v → + suc v) (ℕₚ.+-identityʳ n)
+
+◃-distrib-* : ∀ s t m n → (s 𝕊* t) ◃ (m ℕ* n) ≡ (s ◃ m) * (t ◃ n)
+◃-distrib-* s t zero zero = refl
+◃-distrib-* s t zero (suc n) = refl
+◃-distrib-* s t (suc m) zero =
+ trans
+ (cong₂ _◃_ (𝕊ₚ.*-comm s t) (ℕₚ.*-comm m 0))
+ (*-comm (t ◃ zero) (s ◃ suc m))
+◃-distrib-* s t (suc m) (suc n) =
+ sym (cong₂ _◃_
+ (cong₂ _𝕊*_ (sign-◃ s m) (sign-◃ t n))
+ (∣s◃m∣*∣t◃n∣≡m*n s t (suc m) (suc n)))
+
+------------------------------------------------------------------------
+-- Properties _≤_
+
+≤-reflexive : _≡_ ⇒ _≤_
+≤-reflexive { -[1+ n ]} refl = -≤- ℕₚ.≤-refl
+≤-reflexive {+ n} refl = +≤+ ℕₚ.≤-refl
+
+≤-refl : Reflexive _≤_
+≤-refl = ≤-reflexive refl
+
+≤-trans : Transitive _≤_
+≤-trans -≤+ (+≤+ n≤m) = -≤+
+≤-trans (-≤- n≤m) -≤+ = -≤+
+≤-trans (-≤- n≤m) (-≤- k≤n) = -≤- (ℕₚ.≤-trans k≤n n≤m)
+≤-trans (+≤+ m≤n) (+≤+ n≤k) = +≤+ (ℕₚ.≤-trans m≤n n≤k)
+
+≤-antisym : Antisymmetric _≡_ _≤_
+≤-antisym -≤+ ()
+≤-antisym (-≤- n≤m) (-≤- m≤n) = cong -[1+_] $ ℕₚ.≤-antisym m≤n n≤m
+≤-antisym (+≤+ m≤n) (+≤+ n≤m) = cong (+_) $ ℕₚ.≤-antisym m≤n n≤m
+
+≤-total : Total _≤_
+≤-total (-[1+ m ]) (-[1+ n ]) with ℕₚ.≤-total m n
+... | inj₁ m≤n = inj₂ (-≤- m≤n)
+... | inj₂ n≤m = inj₁ (-≤- n≤m)
+≤-total (-[1+ m ]) (+ n ) = inj₁ -≤+
+≤-total (+ m ) (-[1+ n ]) = inj₂ -≤+
+≤-total (+ m ) (+ n ) with ℕₚ.≤-total m n
+... | inj₁ m≤n = inj₁ (+≤+ m≤n)
+... | inj₂ n≤m = inj₂ (+≤+ n≤m)
+
+≤-isPreorder : IsPreorder _≡_ _≤_
+≤-isPreorder = record
+ { isEquivalence = isEquivalence
+ ; reflexive = ≤-reflexive
+ ; trans = ≤-trans
+ }
+
+≤-isPartialOrder : IsPartialOrder _≡_ _≤_
+≤-isPartialOrder = record
+ { isPreorder = ≤-isPreorder
+ ; antisym = ≤-antisym
+ }
+
+≤-poset : Poset _ _ _
+≤-poset = record
+ { Carrier = ℤ
+ ; _≈_ = _≡_
+ ; _≤_ = _≤_
+ ; isPartialOrder = ≤-isPartialOrder
+ }
+
+≤-isTotalOrder : IsTotalOrder _≡_ _≤_
+≤-isTotalOrder = record
+ { isPartialOrder = ≤-isPartialOrder
+ ; total = ≤-total
+ }
+
+≤-isDecTotalOrder : IsDecTotalOrder _≡_ _≤_
+≤-isDecTotalOrder = record
+ { isTotalOrder = ≤-isTotalOrder
+ ; _≟_ = _≟_
+ ; _≤?_ = _≤?_
+ }
+
+≤-decTotalOrder : DecTotalOrder _ _ _
+≤-decTotalOrder = record
+ { Carrier = ℤ
+ ; _≈_ = _≡_
+ ; _≤_ = _≤_
+ ; isDecTotalOrder = ≤-isDecTotalOrder
+ }
+
+import Relation.Binary.PartialOrderReasoning as POR
+module ≤-Reasoning = POR ≤-poset renaming (_≈⟨_⟩_ to _≡⟨_⟩_)
+
+≤-step : ∀ {n m} → n ≤ m → n ≤ sucℤ m
+≤-step -≤+ = -≤+
+≤-step (+≤+ m≤n) = +≤+ (ℕₚ.≤-step m≤n)
+≤-step (-≤- z≤n) = -≤+
+≤-step (-≤- (s≤s n≤m)) = -≤- (ℕₚ.≤-step n≤m)
+
+n≤1+n : ∀ n → n ≤ (+ 1) + n
+n≤1+n n = ≤-step ≤-refl
+
+------------------------------------------------------------------------
+-- Properties _<_
+
+-<+ : ∀ {m n} → -[1+ m ] < + n
+-<+ {0} = +≤+ z≤n
+-<+ {suc _} = -≤+
+
+<-irrefl : Irreflexive _≡_ _<_
+<-irrefl { + n} refl (+≤+ 1+n≤n) = ℕₚ.<-irrefl refl 1+n≤n
+<-irrefl { -[1+ zero ]} refl ()
+<-irrefl { -[1+ suc n ]} refl (-≤- 1+n≤n) = ℕₚ.<-irrefl refl 1+n≤n
+
+<-asym : Asymmetric _<_
+<-asym {+ n} {+ m} (+≤+ n<m) (+≤+ m<n) =
+ ℕₚ.<-asym n<m m<n
+<-asym {+ n} { -[1+ m ]} () _
+<-asym { -[1+ n ]} {+_ n₁} _ ()
+<-asym { -[1+ 0 ]} { -[1+_] _} () _
+<-asym { -[1+ _ ]} { -[1+_] 0} _ ()
+<-asym { -[1+ suc n ]} { -[1+ suc m ]} (-≤- n<m) (-≤- m<n) =
+ ℕₚ.<-asym n<m m<n
+
+<-trans : Transitive _<_
+<-trans { + m} {_} (+≤+ m<n) (+≤+ n<o) =
+ +≤+ (ℕₚ.<-trans m<n n<o)
+<-trans { -[1+ 0 ]} {_} (+≤+ m<n) (+≤+ n<o) = +≤+ z≤n
+<-trans { -[1+ suc m ]} {+ n} m<n (+≤+ m≤n) = -≤+
+<-trans { -[1+ suc m ]} { -[1+ 0 ]} m<n (+≤+ m≤n) = -≤+
+<-trans { -[1+ suc m ]} { -[1+ suc n ]} (-≤- n≤m) -≤+ = -≤+
+<-trans { -[1+ suc m ]} { -[1+ suc n ]} (-≤- n<m) (-≤- o≤n) =
+ -≤- (ℕₚ.≤-trans o≤n (ℕₚ.<⇒≤ n<m))
+
+<-cmp : Trichotomous _≡_ _<_
+<-cmp (+ m) (+ n) with ℕₚ.<-cmp m n
+... | tri< m<n m≢n m≯n =
+ tri< (+≤+ m<n) (m≢n ∘ +-injective) (m≯n ∘ drop‿+≤+)
+... | tri≈ m≮n m≡n m≯n =
+ tri≈ (m≮n ∘ drop‿+≤+) (cong (+_) m≡n) (m≯n ∘ drop‿+≤+)
+... | tri> m≮n m≢n m>n =
+ tri> (m≮n ∘ drop‿+≤+) (m≢n ∘ +-injective) (+≤+ m>n)
+<-cmp (+_ m) -[1+ 0 ] = tri> (λ()) (λ()) (+≤+ z≤n)
+<-cmp (+_ m) -[1+ suc n ] = tri> (λ()) (λ()) -≤+
+<-cmp -[1+ 0 ] (+ n) = tri< (+≤+ z≤n) (λ()) (λ())
+<-cmp -[1+ suc m ] (+ n) = tri< -≤+ (λ()) (λ())
+<-cmp -[1+ 0 ] -[1+ 0 ] = tri≈ (λ()) refl (λ())
+<-cmp -[1+ 0 ] -[1+ suc n ] = tri> (λ()) (λ()) (-≤- z≤n)
+<-cmp -[1+ suc m ] -[1+ 0 ] = tri< (-≤- z≤n) (λ()) (λ())
+<-cmp -[1+ suc m ] -[1+ suc n ] with ℕₚ.<-cmp (suc m) (suc n)
+... | tri< m<n m≢n m≯n =
+ tri> (m≯n ∘ s≤s ∘ drop‿-≤-) (m≢n ∘ -[1+-injective) (-≤- (≤-pred m<n))
+... | tri≈ m≮n m≡n m≯n =
+ tri≈ (m≯n ∘ s≤s ∘ drop‿-≤-) (cong -[1+_] m≡n) (m≮n ∘ s≤s ∘ drop‿-≤-)
+... | tri> m≮n m≢n m>n =
+ tri< (-≤- (≤-pred m>n)) (m≢n ∘ -[1+-injective) (m≮n ∘ s≤s ∘ drop‿-≤-)
+
+<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
+<-isStrictTotalOrder = record
+ { isEquivalence = isEquivalence
+ ; trans = λ {i} → <-trans {i}
+ ; compare = <-cmp
+ }
+
+<-strictTotalOrder : StrictTotalOrder _ _ _
+<-strictTotalOrder = record
+ { Carrier = ℤ
+ ; _≈_ = _≡_
+ ; _<_ = _<_
+ ; isStrictTotalOrder = <-isStrictTotalOrder
+ }
+
+n≮n : ∀ {n} → n ≮ n
+n≮n {+ n} (+≤+ n<n) = contradiction n<n ℕₚ.1+n≰n
+n≮n { -[1+ 0 ]} ()
+n≮n { -[1+ suc n ]} (-≤- n<n) = contradiction n<n ℕₚ.1+n≰n
+
+<⇒≤ : ∀ {m n} → m < n → m ≤ n
+<⇒≤ m<n = ≤-trans (n≤1+n _) m<n
+
+≰→> : ∀ {x y} → x ≰ y → x > y
+≰→> {+ m} {+ n} m≰n = +≤+ (ℕₚ.≰⇒> (m≰n ∘ +≤+))
+≰→> {+ m} { -[1+ n ]} _ = -<+ {n} {m}
+≰→> { -[1+ m ]} {+ _} m≰n = contradiction -≤+ m≰n
+≰→> { -[1+ 0 ]} { -[1+ 0 ]} m≰n = contradiction ≤-refl m≰n
+≰→> { -[1+ suc _ ]} { -[1+ 0 ]} m≰n = contradiction (-≤- z≤n) m≰n
+≰→> { -[1+ m ]} { -[1+ suc n ]} m≰n with m ℕ≤? n
+... | yes m≤n = -≤- m≤n
+... | no m≰n' = contradiction (-≤- (ℕₚ.≰⇒> m≰n')) m≰n
diff --git a/src/Data/List/All.agda b/src/Data/List/All.agda
index 1885ed7..a1caee8 100644
--- a/src/Data/List/All.agda
+++ b/src/Data/List/All.agda
@@ -6,9 +6,9 @@
module Data.List.All where
-open import Data.List.Base as List hiding (map; all)
+open import Data.List.Base as List hiding (map; tabulate; all)
open import Data.List.Any as Any using (here; there)
-open Any.Membership-≡ using (_∈_; _⊆_)
+open import Data.List.Any.Membership.Propositional using (_∈_)
open import Function
open import Level
open import Relation.Nullary
diff --git a/src/Data/List/All/Properties.agda b/src/Data/List/All/Properties.agda
index 8ac7cd0..9830f36 100644
--- a/src/Data/List/All/Properties.agda
+++ b/src/Data/List/All/Properties.agda
@@ -9,11 +9,13 @@ module Data.List.All.Properties where
open import Data.Bool.Base using (Bool; T)
open import Data.Bool.Properties
open import Data.Empty
-open import Data.List.Base as List
-import Data.List.Any as Any; open Any.Membership-≡
+open import Data.Fin using (Fin) renaming (zero to fzero; suc to fsuc)
+open import Data.List.Base
+open import Data.List.Any.Membership.Propositional
open import Data.List.All as All using (All; []; _∷_)
open import Data.List.Any using (Any; here; there)
-open import Data.Product as Prod
+open import Data.Nat using (zero; suc; z≤n; s≤s; _<_)
+open import Data.Product as Prod using (_×_; _,_; uncurry; uncurry′)
open import Function
open import Function.Equality using (_⟨$⟩_)
open import Function.Equivalence using (_⇔_; module Equivalence)
@@ -21,31 +23,83 @@ open import Function.Inverse using (_↔_)
open import Function.Surjection using (_↠_)
open import Relation.Binary.PropositionalEquality as P using (_≡_)
open import Relation.Nullary
-open import Relation.Unary using (Decidable) renaming (_⊆_ to _⋐_)
+open import Relation.Unary
+ using (Decidable; Universal) renaming (_⊆_ to _⋐_)
--- Functions can be shifted between the predicate and the list.
+------------------------------------------------------------------------
+-- When P is universal All P holds
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+ All-universal : Universal P → ∀ xs → All P xs
+ All-universal u [] = []
+ All-universal u (x ∷ xs) = u x ∷ All-universal u xs
+
+------------------------------------------------------------------------
+-- Lemmas relating Any, All and negation.
+
+module _ {a p} {A : Set a} {P : A → Set p} where
-All-map : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
- {f : A → B} {xs} →
- All (P ∘ f) xs → All P (List.map f xs)
-All-map [] = []
-All-map (p ∷ ps) = p ∷ All-map ps
+ ¬Any⇒All¬ : ∀ xs → ¬ Any P xs → All (¬_ ∘ P) xs
+ ¬Any⇒All¬ [] ¬p = []
+ ¬Any⇒All¬ (x ∷ xs) ¬p = ¬p ∘ here ∷ ¬Any⇒All¬ xs (¬p ∘ there)
-map-All : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
- {f : A → B} {xs} →
- All P (List.map f xs) → All (P ∘ f) xs
-map-All {xs = []} [] = []
-map-All {xs = _ ∷ _} (p ∷ ps) = p ∷ map-All ps
+ All¬⇒¬Any : ∀ {xs} → All (¬_ ∘ P) xs → ¬ Any P xs
+ All¬⇒¬Any [] ()
+ All¬⇒¬Any (¬p ∷ _) (here p) = ¬p p
+ All¬⇒¬Any (_ ∷ ¬p) (there p) = All¬⇒¬Any ¬p p
--- A variant of All.map.
+ ¬All⇒Any¬ : Decidable P → ∀ xs → ¬ All P xs → Any (¬_ ∘ P) xs
+ ¬All⇒Any¬ dec [] ¬∀ = ⊥-elim (¬∀ [])
+ ¬All⇒Any¬ dec (x ∷ xs) ¬∀ with dec x
+ ... | yes p = there (¬All⇒Any¬ dec xs (¬∀ ∘ _∷_ p))
+ ... | no ¬p = here ¬p
+
+ Any¬→¬All : ∀ {xs} → Any (¬_ ∘ P) xs → ¬ All P xs
+ Any¬→¬All (here ¬p) = ¬p ∘ All.head
+ Any¬→¬All (there ¬p) = Any¬→¬All ¬p ∘ All.tail
+
+ ¬Any↠All¬ : ∀ {xs} → (¬ Any P xs) ↠ All (¬_ ∘ P) xs
+ ¬Any↠All¬ = record
+ { to = P.→-to-⟶ (¬Any⇒All¬ _)
+ ; surjective = record
+ { from = P.→-to-⟶ All¬⇒¬Any
+ ; right-inverse-of = to∘from
+ }
+ }
+ where
+
+ to∘from : ∀ {xs} (¬p : All (¬_ ∘ P) xs) → ¬Any⇒All¬ xs (All¬⇒¬Any ¬p) ≡ ¬p
+ to∘from [] = P.refl
+ to∘from (¬p ∷ ¬ps) = P.cong₂ _∷_ P.refl (to∘from ¬ps)
+
+ -- If equality of functions were extensional, then the surjection
+ -- could be strengthened to a bijection.
+
+ from∘to : P.Extensionality _ _ →
+ ∀ xs → (¬p : ¬ Any P xs) → All¬⇒¬Any (¬Any⇒All¬ xs ¬p) ≡ ¬p
+ from∘to ext [] ¬p = ext λ ()
+ from∘to ext (x ∷ xs) ¬p = ext λ
+ { (here p) → P.refl
+ ; (there p) → P.cong (λ f → f p) $ from∘to ext xs (¬p ∘ there)
+ }
+
+ Any¬⇔¬All : ∀ {xs} → Decidable P → Any (¬_ ∘ P) xs ⇔ (¬ All P xs)
+ Any¬⇔¬All dec = record
+ { to = P.→-to-⟶ Any¬→¬All
+ ; from = P.→-to-⟶ (¬All⇒Any¬ dec _)
+ }
+ where
+
+ -- If equality of functions were extensional, then the logical
+ -- equivalence could be strengthened to a surjection.
-gmap : ∀ {a b p q}
- {A : Set a} {B : Set b} {P : A → Set p} {Q : B → Set q}
- {f : A → B} →
- P ⋐ Q ∘ f → All P ⋐ All Q ∘ List.map f
-gmap g = All-map ∘ All.map g
+ to∘from : P.Extensionality _ _ →
+ ∀ {xs} (¬∀ : ¬ All P xs) → Any¬→¬All (¬All⇒Any¬ dec xs ¬∀) ≡ ¬∀
+ to∘from ext ¬∀ = ext (⊥-elim ∘ ¬∀)
--- All and all are related via T.
+------------------------------------------------------------------------
+-- Lemmas relating All to ⊤
All-all : ∀ {a} {A : Set a} (p : A → Bool) {xs} →
All (T ∘ p) xs → T (all p xs)
@@ -58,110 +112,140 @@ all-All p [] _ = []
all-All p (x ∷ xs) px∷xs with Equivalence.to (T-∧ {p x}) ⟨$⟩ px∷xs
all-All p (x ∷ xs) px∷xs | (px , pxs) = px ∷ all-All p xs pxs
+------------------------------------------------------------------------
-- All is anti-monotone.
anti-mono : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
xs ⊆ ys → All P ys → All P xs
anti-mono xs⊆ys pys = All.tabulate (All.lookup pys ∘ xs⊆ys)
--- all is anti-monotone.
-
all-anti-mono : ∀ {a} {A : Set a} (p : A → Bool) {xs ys} →
xs ⊆ ys → T (all p ys) → T (all p xs)
all-anti-mono p xs⊆ys = All-all p ∘ anti-mono xs⊆ys ∘ all-All p _
--- All P (xs ++ ys) is isomorphic to All P xs and All P ys.
+------------------------------------------------------------------------
+-- Introduction (⁺) and elimination (⁻) rules for various list functions
+------------------------------------------------------------------------
+-- map
+
+module _{a b} {A : Set a} {B : Set b} where
+
+ All-map : ∀ {p} {P : B → Set p} {f : A → B} {xs} →
+ All (P ∘ f) xs → All P (map f xs)
+ All-map [] = []
+ All-map (p ∷ ps) = p ∷ All-map ps
+
+ map-All : ∀ {p} {P : B → Set p} {f : A → B} {xs} →
+ All P (map f xs) → All (P ∘ f) xs
+ map-All {xs = []} [] = []
+ map-All {xs = _ ∷ _} (p ∷ ps) = p ∷ map-All ps
+
+ -- A variant of All.map.
+
+ gmap : ∀ {p q} {P : A → Set p} {Q : B → Set q} {f : A → B} →
+ P ⋐ Q ∘ f → All P ⋐ All Q ∘ map f
+ gmap g = All-map ∘ All.map g
+
+------------------------------------------------------------------------
+-- _++_
-private
+module _ {a p} {A : Set a} {P : A → Set p} where
- ++⁺ : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys : List A} →
- All P xs → All P ys → All P (xs ++ ys)
+ ++⁺ : ∀ {xs ys} → All P xs → All P ys → All P (xs ++ ys)
++⁺ [] pys = pys
++⁺ (px ∷ pxs) pys = px ∷ ++⁺ pxs pys
- ++⁻ : ∀ {a p} {A : Set a} {P : A → Set p} (xs : List A) {ys} →
- All P (xs ++ ys) → All P xs × All P ys
+ ++⁻ˡ : ∀ xs {ys} → All P (xs ++ ys) → All P xs
+ ++⁻ˡ [] p = []
+ ++⁻ˡ (x ∷ xs) (px ∷ pxs) = px ∷ (++⁻ˡ _ pxs)
+
+ ++⁻ʳ : ∀ xs {ys} → All P (xs ++ ys) → All P ys
+ ++⁻ʳ [] p = p
+ ++⁻ʳ (x ∷ xs) (px ∷ pxs) = ++⁻ʳ xs pxs
+
+ ++⁻ : ∀ xs {ys} → All P (xs ++ ys) → All P xs × All P ys
++⁻ [] p = [] , p
- ++⁻ (x ∷ xs) (px ∷ pxs) = Prod.map (_∷_ px) id $ ++⁻ _ pxs
-
- ++⁺∘++⁻ : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys}
- (p : All P (xs ++ ys)) → uncurry′ ++⁺ (++⁻ xs p) ≡ p
- ++⁺∘++⁻ [] p = P.refl
- ++⁺∘++⁻ (x ∷ xs) (px ∷ pxs) = P.cong (_∷_ px) $ ++⁺∘++⁻ xs pxs
-
- ++⁻∘++⁺ : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys}
- (p : All P xs × All P ys) → ++⁻ xs (uncurry ++⁺ p) ≡ p
- ++⁻∘++⁺ ([] , pys) = P.refl
- ++⁻∘++⁺ (px ∷ pxs , pys) rewrite ++⁻∘++⁺ (pxs , pys) = P.refl
-
-++↔ : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
- (All P xs × All P ys) ↔ All P (xs ++ ys)
-++↔ {P = P} {xs = xs} = record
- { to = P.→-to-⟶ $ uncurry ++⁺
- ; from = P.→-to-⟶ $ ++⁻ xs
- ; inverse-of = record
- { left-inverse-of = ++⁻∘++⁺
- ; right-inverse-of = ++⁺∘++⁻ xs
+ ++⁻ (x ∷ xs) (px ∷ pxs) = Prod.map (px ∷_) id (++⁻ _ pxs)
+
+ ++↔ : ∀ {xs ys} → (All P xs × All P ys) ↔ All P (xs ++ ys)
+ ++↔ {xs} = record
+ { to = P.→-to-⟶ $ uncurry ++⁺
+ ; from = P.→-to-⟶ $ ++⁻ xs
+ ; inverse-of = record
+ { left-inverse-of = ++⁻∘++⁺
+ ; right-inverse-of = ++⁺∘++⁻ xs
+ }
}
- }
+ where
--- Three lemmas relating Any, All and negation.
+ ++⁺∘++⁻ : ∀ xs {ys} (p : All P (xs ++ ys)) →
+ uncurry′ ++⁺ (++⁻ xs p) ≡ p
+ ++⁺∘++⁻ [] p = P.refl
+ ++⁺∘++⁻ (x ∷ xs) (px ∷ pxs) = P.cong (_∷_ px) $ ++⁺∘++⁻ xs pxs
-¬Any↠All¬ : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
- (¬ Any P xs) ↠ All (¬_ ∘ P) xs
-¬Any↠All¬ {P = P} = record
- { to = P.→-to-⟶ (to _)
- ; surjective = record
- { from = P.→-to-⟶ from
- ; right-inverse-of = to∘from
- }
- }
- where
- to : ∀ xs → ¬ Any P xs → All (¬_ ∘ P) xs
- to [] ¬p = []
- to (x ∷ xs) ¬p = ¬p ∘ here ∷ to xs (¬p ∘ there)
-
- from : ∀ {xs} → All (¬_ ∘ P) xs → ¬ Any P xs
- from [] ()
- from (¬p ∷ _) (here p) = ¬p p
- from (_ ∷ ¬p) (there p) = from ¬p p
-
- to∘from : ∀ {xs} (¬p : All (¬_ ∘ P) xs) → to xs (from ¬p) ≡ ¬p
- to∘from [] = P.refl
- to∘from (¬p ∷ ¬ps) = P.cong₂ _∷_ P.refl (to∘from ¬ps)
-
- -- If equality of functions were extensional, then the surjection
- -- could be strengthened to a bijection.
-
- from∘to : P.Extensionality _ _ →
- ∀ xs → (¬p : ¬ Any P xs) → from (to xs ¬p) ≡ ¬p
- from∘to ext [] ¬p = ext λ ()
- from∘to ext (x ∷ xs) ¬p = ext λ
- { (here p) → P.refl
- ; (there p) → P.cong (λ f → f p) $ from∘to ext xs (¬p ∘ there)
- }
+ ++⁻∘++⁺ : ∀ {xs ys} (p : All P xs × All P ys) →
+ ++⁻ xs (uncurry ++⁺ p) ≡ p
+ ++⁻∘++⁺ ([] , pys) = P.refl
+ ++⁻∘++⁺ (px ∷ pxs , pys) rewrite ++⁻∘++⁺ (pxs , pys) = P.refl
-Any¬→¬All : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
- Any (¬_ ∘ P) xs → ¬ All P xs
-Any¬→¬All (here ¬p) = ¬p ∘ All.head
-Any¬→¬All (there ¬p) = Any¬→¬All ¬p ∘ All.tail
-
-Any¬⇔¬All : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
- Decidable P → Any (¬_ ∘ P) xs ⇔ (¬ All P xs)
-Any¬⇔¬All {P = P} dec = record
- { to = P.→-to-⟶ Any¬→¬All
- ; from = P.→-to-⟶ (from _)
- }
- where
- from : ∀ xs → ¬ All P xs → Any (¬_ ∘ P) xs
- from [] ¬∀ = ⊥-elim (¬∀ [])
- from (x ∷ xs) ¬∀ with dec x
- ... | yes p = there (from xs (¬∀ ∘ _∷_ p))
- ... | no ¬p = here ¬p
+------------------------------------------------------------------------
+-- concat
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+ concat⁺ : ∀ {xss} → All (All P) xss → All P (concat xss)
+ concat⁺ [] = []
+ concat⁺ (pxs ∷ pxss) = ++⁺ pxs (concat⁺ pxss)
+
+ concat⁻ : ∀ {xss} → All P (concat xss) → All (All P) xss
+ concat⁻ {[]} [] = []
+ concat⁻ {xs ∷ xss} pxs = ++⁻ˡ xs pxs ∷ concat⁻ (++⁻ʳ xs pxs)
+
+------------------------------------------------------------------------
+-- take and drop
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+ drop⁺ : ∀ {xs} n → All P xs → All P (drop n xs)
+ drop⁺ zero pxs = pxs
+ drop⁺ (suc n) [] = []
+ drop⁺ (suc n) (px ∷ pxs) = drop⁺ n pxs
+
+ take⁺ : ∀ {xs} n → All P xs → All P (take n xs)
+ take⁺ zero pxs = []
+ take⁺ (suc n) [] = []
+ take⁺ (suc n) (px ∷ pxs) = px ∷ take⁺ n pxs
+
+------------------------------------------------------------------------
+-- applyUpTo
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+ applyUpTo⁺₁ : ∀ f n → (∀ {i} → i < n → P (f i)) → All P (applyUpTo f n)
+ applyUpTo⁺₁ f zero Pf = []
+ applyUpTo⁺₁ f (suc n) Pf = Pf (s≤s z≤n) ∷ applyUpTo⁺₁ (f ∘ suc) n (Pf ∘ s≤s)
+
+ applyUpTo⁺₂ : ∀ f n → (∀ i → P (f i)) → All P (applyUpTo f n)
+ applyUpTo⁺₂ f n Pf = applyUpTo⁺₁ f n (λ _ → Pf _)
+
+ applyUpTo⁻ : ∀ f n → All P (applyUpTo f n) → ∀ {i} → i < n → P (f i)
+ applyUpTo⁻ f zero pxs ()
+ applyUpTo⁻ f (suc n) (px ∷ _) (s≤s z≤n) = px
+ applyUpTo⁻ f (suc n) (_ ∷ pxs) (s≤s (s≤s i<n)) =
+ applyUpTo⁻ (f ∘ suc) n pxs (s≤s i<n)
+
+------------------------------------------------------------------------
+-- tabulate
+
+module _ {a p} {A : Set a} {P : A → Set p} where
- -- If equality of functions were extensional, then the logical
- -- equivalence could be strengthened to a surjection.
+ tabulate⁺ : ∀ {n} {f : Fin n → A} →
+ (∀ i → P (f i)) → All P (tabulate f)
+ tabulate⁺ {zero} Pf = []
+ tabulate⁺ {suc n} Pf = Pf fzero ∷ tabulate⁺ (Pf ∘ fsuc)
- to∘from : P.Extensionality _ _ →
- ∀ {xs} (¬∀ : ¬ All P xs) → Any¬→¬All (from xs ¬∀) ≡ ¬∀
- to∘from ext ¬∀ = ext (⊥-elim ∘ ¬∀)
+ tabulate⁻ : ∀ {n} {f : Fin n → A} →
+ All P (tabulate f) → (∀ i → P (f i))
+ tabulate⁻ {zero} pf ()
+ tabulate⁻ {suc n} (px ∷ _) fzero = px
+ tabulate⁻ {suc n} (_ ∷ pf) (fsuc i) = tabulate⁻ pf i
diff --git a/src/Data/List/Any.agda b/src/Data/List/Any.agda
index eefc92c..a9f94da 100644
--- a/src/Data/List/Any.agda
+++ b/src/Data/List/Any.agda
@@ -4,51 +4,38 @@
-- Lists where at least one element satisfies a given property
------------------------------------------------------------------------
-module Data.List.Any where
+module Data.List.Any {a} {A : Set a} where
open import Data.Empty
open import Data.Fin
-open import Function
-open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence as Equiv using (module Equivalence)
-open import Function.Related as Related hiding (_∼[_]_)
open import Data.List.Base as List using (List; []; _∷_)
-open import Data.Product as Prod using (∃; _×_; _,_)
-open import Level using (Level; _⊔_)
-open import Relation.Nullary
+open import Data.Product as Prod using (∃; _,_)
+open import Level using (_⊔_)
+open import Relation.Nullary using (¬_; yes; no)
import Relation.Nullary.Decidable as Dec
open import Relation.Unary using (Decidable) renaming (_⊆_ to _⋐_)
-open import Relation.Binary hiding (Decidable)
-import Relation.Binary.InducedPreorders as Ind
-open import Relation.Binary.List.Pointwise as ListEq using ([]; _∷_)
-open import Relation.Binary.PropositionalEquality as PropEq
- using (_≡_)
-- Any P xs means that at least one element in xs satisfies P.
-data Any {a p} {A : Set a}
- (P : A → Set p) : List A → Set (a ⊔ p) where
+data Any {p} (P : A → Set p) : List A → Set (a ⊔ p) where
here : ∀ {x xs} (px : P x) → Any P (x ∷ xs)
there : ∀ {x xs} (pxs : Any P xs) → Any P (x ∷ xs)
-- Map.
-map : ∀ {a p q} {A : Set a} {P : A → Set p} → {Q : A → Set q} →
- P ⋐ Q → Any P ⋐ Any Q
+map : ∀ {p q} {P : A → Set p} {Q : A → Set q} → P ⋐ Q → Any P ⋐ Any Q
map g (here px) = here (g px)
map g (there pxs) = there (map g pxs)
-- If the head does not satisfy the predicate, then the tail will.
-tail : ∀ {a p} {A : Set a} {x : A} {xs : List A} {P : A → Set p} →
- ¬ P x → Any P (x ∷ xs) → Any P xs
+tail : ∀ {p} {P : A → Set p} {x xs} → ¬ P x → Any P (x ∷ xs) → Any P xs
tail ¬px (here px) = ⊥-elim (¬px px)
tail ¬px (there pxs) = pxs
-- Decides Any.
-any : ∀ {a p} {A : Set a} {P : A → Set p} →
- Decidable P → Decidable (Any P)
+any : ∀ {p} {P : A → Set p} → Decidable P → Decidable (Any P)
any p [] = no λ()
any p (x ∷ xs) with p x
any p (x ∷ xs) | yes px = yes (here px)
@@ -56,167 +43,12 @@ any p (x ∷ xs) | no ¬px = Dec.map′ there (tail ¬px) (any p xs)
-- index x∈xs is the list position (zero-based) which x∈xs points to.
-index : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
- Any P xs → Fin (List.length xs)
+index : ∀ {p} {P : A → Set p} {xs} → Any P xs → Fin (List.length xs)
index (here px) = zero
index (there pxs) = suc (index pxs)
-------------------------------------------------------------------------
--- List membership and some related definitions
-
-module Membership {c ℓ : Level} (S : Setoid c ℓ) where
-
- private
- open module S = Setoid S using (_≈_) renaming (Carrier to A)
- open module LS = Setoid (ListEq.setoid S)
- using () renaming (_≈_ to _≋_)
-
- -- If a predicate P respects the underlying equality then Any P
- -- respects the list equality.
-
- lift-resp : ∀ {p} {P : A → Set p} →
- P Respects _≈_ → Any P Respects _≋_
- lift-resp resp [] ()
- lift-resp resp (x≈y ∷ xs≈ys) (here px) = here (resp x≈y px)
- lift-resp resp (x≈y ∷ xs≈ys) (there pxs) =
- there (lift-resp resp xs≈ys pxs)
-
- -- List membership.
-
- infix 4 _∈_ _∉_
-
- _∈_ : A → List A → Set _
- x ∈ xs = Any (_≈_ x) xs
-
- _∉_ : A → List A → Set _
- x ∉ xs = ¬ x ∈ xs
-
- -- Subsets.
-
- infix 4 _⊆_ _⊈_
-
- _⊆_ : List A → List A → Set _
- xs ⊆ ys = ∀ {x} → x ∈ xs → x ∈ ys
-
- _⊈_ : List A → List A → Set _
- xs ⊈ ys = ¬ xs ⊆ ys
-
- -- Equality is respected by the predicate which is used to define
- -- _∈_.
-
- ∈-resp-≈ : ∀ {x} → (_≈_ x) Respects _≈_
- ∈-resp-≈ = flip S.trans
-
- -- List equality is respected by _∈_.
-
- ∈-resp-list-≈ : ∀ {x} → _∈_ x Respects _≋_
- ∈-resp-list-≈ = lift-resp ∈-resp-≈
-
- -- _⊆_ is a preorder.
-
- ⊆-preorder : Preorder _ _ _
- ⊆-preorder = Ind.InducedPreorder₂ (ListEq.setoid S) _∈_ ∈-resp-list-≈
-
- module ⊆-Reasoning where
- import Relation.Binary.PreorderReasoning as PreR
- open PreR ⊆-preorder public
- renaming (_∼⟨_⟩_ to _⊆⟨_⟩_)
-
- infix 1 _∈⟨_⟩_
-
- _∈⟨_⟩_ : ∀ x {xs ys} → x ∈ xs → xs IsRelatedTo ys → x ∈ ys
- x ∈⟨ x∈xs ⟩ xs⊆ys = (begin xs⊆ys) x∈xs
-
- -- A variant of List.map.
-
- map-with-∈ : ∀ {b} {B : Set b}
- (xs : List A) → (∀ {x} → x ∈ xs → B) → List B
- map-with-∈ [] f = []
- map-with-∈ (x ∷ xs) f = f (here S.refl) ∷ map-with-∈ xs (f ∘ there)
-
- -- Finds an element satisfying the predicate.
-
- find : ∀ {p} {P : A → Set p} {xs} →
- Any P xs → ∃ λ x → x ∈ xs × P x
- find (here px) = (_ , here S.refl , px)
- find (there pxs) = Prod.map id (Prod.map there id) (find pxs)
-
- lose : ∀ {p} {P : A → Set p} {x xs} →
- P Respects _≈_ → x ∈ xs → P x → Any P xs
- lose resp x∈xs px = map (flip resp px) x∈xs
-
--- The code above instantiated (and slightly changed) for
--- propositional equality, along with some additional definitions.
-
-module Membership-≡ where
-
- private
- open module M {a} {A : Set a} = Membership (PropEq.setoid A) public
- hiding (lift-resp; lose; ⊆-preorder; module ⊆-Reasoning)
-
- lose : ∀ {a p} {A : Set a} {P : A → Set p} {x xs} →
- x ∈ xs → P x → Any P xs
- lose {P = P} = M.lose (PropEq.subst P)
-
- -- _⊆_ is a preorder.
-
- ⊆-preorder : ∀ {a} → Set a → Preorder _ _ _
- ⊆-preorder A = Ind.InducedPreorder₂ (PropEq.setoid (List A)) _∈_
- (PropEq.subst (_∈_ _))
-
- -- Set and bag equality and related preorders.
-
- open Related public
- using (Kind; Symmetric-kind)
- renaming ( implication to subset
- ; reverse-implication to superset
- ; equivalence to set
- ; injection to subbag
- ; reverse-injection to superbag
- ; bijection to bag
- )
-
- [_]-Order : Kind → ∀ {a} → Set a → Preorder _ _ _
- [ k ]-Order A = Related.InducedPreorder₂ k (_∈_ {A = A})
-
- [_]-Equality : Symmetric-kind → ∀ {a} → Set a → Setoid _ _
- [ k ]-Equality A = Related.InducedEquivalence₂ k (_∈_ {A = A})
-
- infix 4 _∼[_]_
-
- _∼[_]_ : ∀ {a} {A : Set a} → List A → Kind → List A → Set _
- _∼[_]_ {A = A} xs k ys = Preorder._∼_ ([ k ]-Order A) xs ys
-
- -- Bag equality implies the other relations.
-
- bag-=⇒ : ∀ {k a} {A : Set a} {xs ys : List A} →
- xs ∼[ bag ] ys → xs ∼[ k ] ys
- bag-=⇒ xs≈ys = ↔⇒ xs≈ys
-
- -- "Equational" reasoning for _⊆_.
-
- module ⊆-Reasoning where
- import Relation.Binary.PreorderReasoning as PreR
- private
- open module PR {a} {A : Set a} = PreR (⊆-preorder A) public
- renaming (_∼⟨_⟩_ to _⊆⟨_⟩_; _≈⟨_⟩_ to _≡⟨_⟩_)
-
- infixr 2 _∼⟨_⟩_
- infix 1 _∈⟨_⟩_
-
- _∈⟨_⟩_ : ∀ {a} {A : Set a} x {xs ys : List A} →
- x ∈ xs → xs IsRelatedTo ys → x ∈ ys
- x ∈⟨ x∈xs ⟩ xs⊆ys = (begin xs⊆ys) x∈xs
-
- _∼⟨_⟩_ : ∀ {k a} {A : Set a} xs {ys zs : List A} →
- xs ∼[ ⌊ k ⌋→ ] ys → ys IsRelatedTo zs → xs IsRelatedTo zs
- xs ∼⟨ xs≈ys ⟩ ys≈zs = xs ⊆⟨ ⇒→ xs≈ys ⟩ ys≈zs
-
-------------------------------------------------------------------------
--- Another function
-
-- If any element satisfies P, then P is satisfied.
-satisfied : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
- Any P xs → ∃ P
-satisfied = Prod.map id Prod.proj₂ ∘ Membership-≡.find
+satisfied : ∀ {p} {P : A → Set p} {xs} → Any P xs → ∃ P
+satisfied (here px) = _ , px
+satisfied (there pxs) = satisfied pxs
diff --git a/src/Data/List/Any/BagAndSetEquality.agda b/src/Data/List/Any/BagAndSetEquality.agda
index 983f16d..ad8bbba 100644
--- a/src/Data/List/Any/BagAndSetEquality.agda
+++ b/src/Data/List/Any/BagAndSetEquality.agda
@@ -15,6 +15,7 @@ open import Data.List as List
import Data.List.Properties as LP
open import Data.List.Any as Any using (Any); open Any.Any
open import Data.List.Any.Properties
+open import Data.List.Any.Membership.Propositional
open import Data.Product
open import Data.Sum
open import Function
@@ -30,7 +31,7 @@ open import Relation.Binary.PropositionalEquality as P
open import Relation.Binary.Sum
open import Relation.Nullary
-open Any.Membership-≡
+open import Data.List.Any.Membership.Propositional.Properties
private
module Eq {k a} {A : Set a} = Setoid ([ k ]-Equality A)
module Ord {k a} {A : Set a} = Preorder ([ k ]-Order A)
diff --git a/src/Data/List/Any/Membership.agda b/src/Data/List/Any/Membership.agda
index e2904a9..f46729b 100644
--- a/src/Data/List/Any/Membership.agda
+++ b/src/Data/List/Any/Membership.agda
@@ -1,273 +1,55 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Properties related to list membership
+-- List membership and some related definitions
------------------------------------------------------------------------
--- List membership is defined in Data.List.Any. This module does not
--- treat the general variant of list membership, parametrised on a
--- setoid, only the variant where the equality is fixed to be
--- propositional equality.
+open import Relation.Binary hiding (Decidable)
-module Data.List.Any.Membership where
+module Data.List.Any.Membership {c ℓ} (S : Setoid c ℓ) where
-open import Algebra
-open import Category.Monad
-open import Data.Bool.Base using (Bool; false; true; T)
-open import Data.Empty
-open import Function
-open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence using (module Equivalence)
-import Function.Injection as Inj
-open import Function.Inverse as Inv using (_↔_; module Inverse)
-import Function.Related as Related
-open import Function.Related.TypeIsomorphisms
-open import Data.List as List
-open import Data.List.Any as Any using (Any; here; there)
-open import Data.List.Any.Properties
-open import Data.Nat as Nat
-import Data.Nat.Properties as NatProp
-open import Data.Product as Prod
-open import Data.Sum as Sum
-open import Relation.Binary
-open import Relation.Binary.PropositionalEquality as P
- using (_≡_; refl; _≗_)
-import Relation.Binary.Properties.DecTotalOrder as DTOProperties
-import Relation.Binary.Sigma.Pointwise as Σ
-open import Relation.Unary using (_⟨×⟩_)
-open import Relation.Nullary
-open import Relation.Nullary.Negation
+open import Function using (_∘_; id; flip)
+open import Data.List.Base as List using (List; []; _∷_)
+open import Data.List.Any using (Any; map; here; there)
+open import Data.Product as Prod using (∃; _×_; _,_)
+open import Relation.Nullary using (¬_)
-open Any.Membership-≡
-private
- module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
- open module ListMonad {ℓ} = RawMonad (List.monad {ℓ = ℓ})
+open Setoid S renaming (Carrier to A)
-------------------------------------------------------------------------
--- Properties relating _∈_ to various list functions
-
--- Isomorphisms.
-
-map-∈↔ : ∀ {a b} {A : Set a} {B : Set b} {f : A → B} {y xs} →
- (∃ λ x → x ∈ xs × y ≡ f x) ↔ y ∈ List.map f xs
-map-∈↔ {a} {b} {f = f} {y} {xs} =
- (∃ λ x → x ∈ xs × y ≡ f x) ↔⟨ Any↔ {a = a} {p = b} ⟩
- Any (λ x → y ≡ f x) xs ↔⟨ map↔ {a = a} {b = b} {p = b} ⟩
- y ∈ List.map f xs ∎
- where open Related.EquationalReasoning
-
-concat-∈↔ : ∀ {a} {A : Set a} {x : A} {xss} →
- (∃ λ xs → x ∈ xs × xs ∈ xss) ↔ x ∈ concat xss
-concat-∈↔ {a} {x = x} {xss} =
- (∃ λ xs → x ∈ xs × xs ∈ xss) ↔⟨ Σ.cong {a₁ = a} {b₁ = a} {b₂ = a} Inv.id $ ×⊎.*-comm _ _ ⟩
- (∃ λ xs → xs ∈ xss × x ∈ xs) ↔⟨ Any↔ {a = a} {p = a} ⟩
- Any (Any (_≡_ x)) xss ↔⟨ concat↔ {a = a} {p = a} ⟩
- x ∈ concat xss ∎
- where open Related.EquationalReasoning
-
->>=-∈↔ : ∀ {ℓ} {A B : Set ℓ} {xs} {f : A → List B} {y} →
- (∃ λ x → x ∈ xs × y ∈ f x) ↔ y ∈ (xs >>= f)
->>=-∈↔ {ℓ} {xs = xs} {f} {y} =
- (∃ λ x → x ∈ xs × y ∈ f x) ↔⟨ Any↔ {a = ℓ} {p = ℓ} ⟩
- Any (Any (_≡_ y) ∘ f) xs ↔⟨ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
- y ∈ (xs >>= f) ∎
- where open Related.EquationalReasoning
-
-⊛-∈↔ : ∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) {xs y} →
- (∃₂ λ f x → f ∈ fs × x ∈ xs × y ≡ f x) ↔ y ∈ (fs ⊛ xs)
-⊛-∈↔ {ℓ} fs {xs} {y} =
- (∃₂ λ f x → f ∈ fs × x ∈ xs × y ≡ f x) ↔⟨ Σ.cong {a₁ = ℓ} {b₁ = ℓ} {b₂ = ℓ} Inv.id (∃∃↔∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
- (∃ λ f → f ∈ fs × ∃ λ x → x ∈ xs × y ≡ f x) ↔⟨ Σ.cong {a₁ = ℓ} {b₁ = ℓ} {b₂ = ℓ}
- Inv.id ((_ ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any↔ {a = ℓ} {p = ℓ}) ⟩
- (∃ λ f → f ∈ fs × Any (_≡_ y ∘ f) xs) ↔⟨ Any↔ {a = ℓ} {p = ℓ} ⟩
- Any (λ f → Any (_≡_ y ∘ f) xs) fs ↔⟨ ⊛↔ ⟩
- y ∈ (fs ⊛ xs) ∎
- where open Related.EquationalReasoning
-
-⊗-∈↔ : ∀ {A B : Set} {xs ys} {x : A} {y : B} →
- (x ∈ xs × y ∈ ys) ↔ (x , y) ∈ (xs ⊗ ys)
-⊗-∈↔ {A} {B} {xs} {ys} {x} {y} =
- (x ∈ xs × y ∈ ys) ↔⟨ ⊗↔′ ⟩
- Any (_≡_ x ⟨×⟩ _≡_ y) (xs ⊗ ys) ↔⟨ Any-cong helper (_ ∎) ⟩
- (x , y) ∈ (xs ⊗ ys) ∎
- where
- open Related.EquationalReasoning
-
- helper : (p : A × B) → (x ≡ proj₁ p × y ≡ proj₂ p) ↔ (x , y) ≡ p
- helper (x′ , y′) = record
- { to = P.→-to-⟶ (uncurry $ P.cong₂ _,_)
- ; from = P.→-to-⟶ < P.cong proj₁ , P.cong proj₂ >
- ; inverse-of = record
- { left-inverse-of = λ _ → P.cong₂ _,_ (P.proof-irrelevance _ _)
- (P.proof-irrelevance _ _)
- ; right-inverse-of = λ _ → P.proof-irrelevance _ _
- }
- }
-
--- Other properties.
-
-filter-∈ : ∀ {a} {A : Set a} (p : A → Bool) (xs : List A) {x} →
- x ∈ xs → p x ≡ true → x ∈ filter p xs
-filter-∈ p [] () _
-filter-∈ p (x ∷ xs) (here refl) px≡true rewrite px≡true = here refl
-filter-∈ p (y ∷ xs) (there pxs) px≡true with p y
-... | true = there (filter-∈ p xs pxs px≡true)
-... | false = filter-∈ p xs pxs px≡true
-
-------------------------------------------------------------------------
--- Properties relating _∈_ to various list functions
-
--- Various functions are monotone.
+-- List membership.
-mono : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
- xs ⊆ ys → Any P xs → Any P ys
-mono xs⊆ys =
- _⟨$⟩_ (Inverse.to Any↔) ∘′
- Prod.map id (Prod.map xs⊆ys id) ∘
- _⟨$⟩_ (Inverse.from Any↔)
-
-map-mono : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) {xs ys} →
- xs ⊆ ys → List.map f xs ⊆ List.map f ys
-map-mono f xs⊆ys =
- _⟨$⟩_ (Inverse.to map-∈↔) ∘
- Prod.map id (Prod.map xs⊆ys id) ∘
- _⟨$⟩_ (Inverse.from map-∈↔)
-
-_++-mono_ : ∀ {a} {A : Set a} {xs₁ xs₂ ys₁ ys₂ : List A} →
- xs₁ ⊆ ys₁ → xs₂ ⊆ ys₂ → xs₁ ++ xs₂ ⊆ ys₁ ++ ys₂
-_++-mono_ xs₁⊆ys₁ xs₂⊆ys₂ =
- _⟨$⟩_ (Inverse.to ++↔) ∘
- Sum.map xs₁⊆ys₁ xs₂⊆ys₂ ∘
- _⟨$⟩_ (Inverse.from ++↔)
-
-concat-mono : ∀ {a} {A : Set a} {xss yss : List (List A)} →
- xss ⊆ yss → concat xss ⊆ concat yss
-concat-mono {a} xss⊆yss =
- _⟨$⟩_ (Inverse.to $ concat-∈↔ {a = a}) ∘
- Prod.map id (Prod.map id xss⊆yss) ∘
- _⟨$⟩_ (Inverse.from $ concat-∈↔ {a = a})
-
->>=-mono : ∀ {ℓ} {A B : Set ℓ} (f g : A → List B) {xs ys} →
- xs ⊆ ys → (∀ {x} → f x ⊆ g x) →
- (xs >>= f) ⊆ (ys >>= g)
->>=-mono {ℓ} f g xs⊆ys f⊆g =
- _⟨$⟩_ (Inverse.to $ >>=-∈↔ {ℓ = ℓ}) ∘
- Prod.map id (Prod.map xs⊆ys f⊆g) ∘
- _⟨$⟩_ (Inverse.from $ >>=-∈↔ {ℓ = ℓ})
-
-_⊛-mono_ : ∀ {ℓ} {A B : Set ℓ}
- {fs gs : List (A → B)} {xs ys : List A} →
- fs ⊆ gs → xs ⊆ ys → (fs ⊛ xs) ⊆ (gs ⊛ ys)
-_⊛-mono_ {fs = fs} {gs} fs⊆gs xs⊆ys =
- _⟨$⟩_ (Inverse.to $ ⊛-∈↔ gs) ∘
- Prod.map id (Prod.map id (Prod.map fs⊆gs (Prod.map xs⊆ys id))) ∘
- _⟨$⟩_ (Inverse.from $ ⊛-∈↔ fs)
-
-_⊗-mono_ : {A B : Set} {xs₁ ys₁ : List A} {xs₂ ys₂ : List B} →
- xs₁ ⊆ ys₁ → xs₂ ⊆ ys₂ → (xs₁ ⊗ xs₂) ⊆ (ys₁ ⊗ ys₂)
-xs₁⊆ys₁ ⊗-mono xs₂⊆ys₂ =
- _⟨$⟩_ (Inverse.to ⊗-∈↔) ∘
- Prod.map xs₁⊆ys₁ xs₂⊆ys₂ ∘
- _⟨$⟩_ (Inverse.from ⊗-∈↔)
-
-any-mono : ∀ {a} {A : Set a} (p : A → Bool) →
- ∀ {xs ys} → xs ⊆ ys → T (any p xs) → T (any p ys)
-any-mono {a} p xs⊆ys =
- _⟨$⟩_ (Equivalence.to $ any⇔ {a = a}) ∘
- mono xs⊆ys ∘
- _⟨$⟩_ (Equivalence.from $ any⇔ {a = a})
-
-map-with-∈-mono :
- ∀ {a b} {A : Set a} {B : Set b}
- {xs : List A} {f : ∀ {x} → x ∈ xs → B}
- {ys : List A} {g : ∀ {x} → x ∈ ys → B}
- (xs⊆ys : xs ⊆ ys) → (∀ {x} → f {x} ≗ g ∘ xs⊆ys) →
- map-with-∈ xs f ⊆ map-with-∈ ys g
-map-with-∈-mono {f = f} {g = g} xs⊆ys f≈g {x} =
- _⟨$⟩_ (Inverse.to map-with-∈↔) ∘
- Prod.map id (Prod.map xs⊆ys (λ {x∈xs} x≡fx∈xs → begin
- x ≡⟨ x≡fx∈xs ⟩
- f x∈xs ≡⟨ f≈g x∈xs ⟩
- g (xs⊆ys x∈xs) ∎)) ∘
- _⟨$⟩_ (Inverse.from map-with-∈↔)
- where open P.≡-Reasoning
-
--- Other properties.
-
-filter-⊆ : ∀ {a} {A : Set a} (p : A → Bool) →
- (xs : List A) → filter p xs ⊆ xs
-filter-⊆ _ [] = λ ()
-filter-⊆ p (x ∷ xs) with p x | filter-⊆ p xs
-... | false | hyp = there ∘ hyp
-... | true | hyp =
- λ { (here eq) → here eq
- ; (there ∈filter) → there (hyp ∈filter)
- }
-
-------------------------------------------------------------------------
--- Other properties
+infix 4 _∈_ _∉_
--- Only a finite number of distinct elements can be members of a
--- given list.
+_∈_ : A → List A → Set _
+x ∈ xs = Any (_≈_ x) xs
-finite : ∀ {a} {A : Set a}
- (f : Inj.Injection (P.setoid ℕ) (P.setoid A)) →
- ∀ xs → ¬ (∀ i → Inj.Injection.to f ⟨$⟩ i ∈ xs)
-finite inj [] ∈[] with ∈[] zero
-... | ()
-finite {A = A} inj (x ∷ xs) ∈x∷xs = excluded-middle helper
- where
- open Inj.Injection inj
+_∉_ : A → List A → Set _
+x ∉ xs = ¬ x ∈ xs
- module DTO = DecTotalOrder Nat.decTotalOrder
- module STO = StrictTotalOrder
- (DTOProperties.strictTotalOrder Nat.decTotalOrder)
- module CS = CommutativeSemiring NatProp.commutativeSemiring
+-- Subsets.
- not-x : ∀ {i} → ¬ (to ⟨$⟩ i ≡ x) → to ⟨$⟩ i ∈ xs
- not-x {i} ≢x with ∈x∷xs i
- ... | here ≡x = ⊥-elim (≢x ≡x)
- ... | there ∈xs = ∈xs
+infix 4 _⊆_ _⊈_
- helper : ¬ Dec (∃ λ i → to ⟨$⟩ i ≡ x)
- helper (no ≢x) = finite inj xs (λ i → not-x (≢x ∘ _,_ i))
- helper (yes (i , ≡x)) = finite inj′ xs ∈xs
- where
- open P
+_⊆_ : List A → List A → Set _
+xs ⊆ ys = ∀ {x} → x ∈ xs → x ∈ ys
- f : ℕ → A
- f j with STO.compare i j
- f j | tri< _ _ _ = to ⟨$⟩ suc j
- f j | tri≈ _ _ _ = to ⟨$⟩ suc j
- f j | tri> _ _ _ = to ⟨$⟩ j
+_⊈_ : List A → List A → Set _
+xs ⊈ ys = ¬ xs ⊆ ys
- ∈-if-not-i : ∀ {j} → i ≢ j → to ⟨$⟩ j ∈ xs
- ∈-if-not-i i≢j = not-x (i≢j ∘ injective ∘ trans ≡x ∘ sym)
+-- A variant of List.map.
- lemma : ∀ {k j} → k ≤ j → suc j ≢ k
- lemma 1+j≤j refl = NatProp.1+n≰n 1+j≤j
+map-with-∈ : ∀ {b} {B : Set b}
+ (xs : List A) → (∀ {x} → x ∈ xs → B) → List B
+map-with-∈ [] f = []
+map-with-∈ (x ∷ xs) f = f (here refl) ∷ map-with-∈ xs (f ∘ there)
- ∈xs : ∀ j → f j ∈ xs
- ∈xs j with STO.compare i j
- ∈xs j | tri< (i≤j , _) _ _ = ∈-if-not-i (lemma i≤j ∘ sym)
- ∈xs j | tri> _ i≢j _ = ∈-if-not-i i≢j
- ∈xs .i | tri≈ _ refl _ =
- ∈-if-not-i (NatProp.m≢1+m+n i ∘
- subst (_≡_ i ∘ suc) (sym $ proj₂ CS.+-identity i))
+-- Finds an element satisfying the predicate.
- injective′ : Inj.Injective {B = P.setoid A} (→-to-⟶ f)
- injective′ {j} {k} eq with STO.compare i j | STO.compare i k
- ... | tri< _ _ _ | tri< _ _ _ = cong pred $ injective eq
- ... | tri< _ _ _ | tri≈ _ _ _ = cong pred $ injective eq
- ... | tri< (i≤j , _) _ _ | tri> _ _ (k≤i , _) = ⊥-elim (lemma (DTO.trans k≤i i≤j) $ injective eq)
- ... | tri≈ _ _ _ | tri< _ _ _ = cong pred $ injective eq
- ... | tri≈ _ _ _ | tri≈ _ _ _ = cong pred $ injective eq
- ... | tri≈ _ i≡j _ | tri> _ _ (k≤i , _) = ⊥-elim (lemma (subst (_≤_ k) i≡j k≤i) $ injective eq)
- ... | tri> _ _ (j≤i , _) | tri< (i≤k , _) _ _ = ⊥-elim (lemma (DTO.trans j≤i i≤k) $ sym $ injective eq)
- ... | tri> _ _ (j≤i , _) | tri≈ _ i≡k _ = ⊥-elim (lemma (subst (_≤_ j) i≡k j≤i) $ sym $ injective eq)
- ... | tri> _ _ (j≤i , _) | tri> _ _ (k≤i , _) = injective eq
+find : ∀ {p} {P : A → Set p} {xs} →
+ Any P xs → ∃ λ x → x ∈ xs × P x
+find (here px) = (_ , here refl , px)
+find (there pxs) = Prod.map id (Prod.map there id) (find pxs)
- inj′ = record
- { to = →-to-⟶ {B = P.setoid A} f
- ; injective = injective′
- }
+lose : ∀ {p} {P : A → Set p} {x xs} →
+ P Respects _≈_ → x ∈ xs → P x → Any P xs
+lose resp x∈xs px = map (flip resp px) x∈xs
diff --git a/src/Data/List/Any/Membership/Properties.agda b/src/Data/List/Any/Membership/Properties.agda
new file mode 100644
index 0000000..d49e1dc
--- /dev/null
+++ b/src/Data/List/Any/Membership/Properties.agda
@@ -0,0 +1,72 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties related to propositional list membership
+------------------------------------------------------------------------
+
+open import Data.List
+open import Data.List.Any as Any using (here; there)
+open import Data.List.Any.Properties
+import Data.List.Any.Membership as Membership
+open import Data.Product using (∃; _×_; _,_)
+open import Function using (flip)
+open import Relation.Binary
+open import Relation.Binary.InducedPreorders using (InducedPreorder₂)
+open import Relation.Binary.List.Pointwise as ListEq
+ using () renaming (Rel to ListRel)
+
+
+module Data.List.Any.Membership.Properties where
+
+module SingleSetoid {c ℓ} (S : Setoid c ℓ) where
+
+ open Setoid S
+ open import Data.List.Any.Membership S
+
+ -- Equality is respected by the predicate which is used to define
+ -- _∈_.
+
+ ∈-resp-≈ : ∀ {x} → (x ≈_) Respects _≈_
+ ∈-resp-≈ = flip trans
+
+ -- List equality is respected by _∈_.
+
+ ∈-resp-≋ : ∀ {x} → (x ∈_) Respects (ListRel _≈_)
+ ∈-resp-≋ = lift-resp ∈-resp-≈
+
+ -- _⊆_ is a preorder.
+
+ ⊆-preorder : Preorder _ _ _
+ ⊆-preorder = InducedPreorder₂ (ListEq.setoid S) _∈_ ∈-resp-≋
+
+ module ⊆-Reasoning where
+ import Relation.Binary.PreorderReasoning as PreR
+ open PreR ⊆-preorder public
+ renaming (_∼⟨_⟩_ to _⊆⟨_⟩_)
+
+ infix 1 _∈⟨_⟩_
+
+ _∈⟨_⟩_ : ∀ x {xs ys} → x ∈ xs → xs IsRelatedTo ys → x ∈ ys
+ x ∈⟨ x∈xs ⟩ xs⊆ys = (begin xs⊆ys) x∈xs
+
+open SingleSetoid public
+
+
+module DoubleSetoid {c₁ c₂ ℓ₁ ℓ₂}
+ (S₁ : Setoid c₁ ℓ₁) (S₂ : Setoid c₂ ℓ₂) where
+
+ open Setoid S₁ renaming (Carrier to A₁; _≈_ to _≈₁_; refl to refl₁)
+ open Setoid S₂ renaming (Carrier to A₂; _≈_ to _≈₂_)
+
+ open Membership S₁ using (find) renaming (_∈_ to _∈₁_)
+ open Membership S₂ using () renaming (_∈_ to _∈₂_)
+
+ ∈-map⁺ : ∀ {f} → f Preserves _≈₁_ ⟶ _≈₂_ → ∀ {x xs} →
+ x ∈₁ xs → f x ∈₂ map f xs
+ ∈-map⁺ pres x∈xs = map⁺ (Any.map pres x∈xs)
+
+ ∈-map⁻ : ∀ {y xs f} → y ∈₂ map f xs →
+ ∃ λ x → x ∈₁ xs × y ≈₂ f x
+ ∈-map⁻ x∈map = find (map⁻ x∈map)
+
+open DoubleSetoid public
diff --git a/src/Data/List/Any/Membership/Propositional.agda b/src/Data/List/Any/Membership/Propositional.agda
new file mode 100644
index 0000000..4d29544
--- /dev/null
+++ b/src/Data/List/Any/Membership/Propositional.agda
@@ -0,0 +1,84 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Data.List.Any.Membership instantiated with propositional equality,
+-- along with some additional definitions.
+------------------------------------------------------------------------
+
+module Data.List.Any.Membership.Propositional where
+
+open import Data.Empty
+open import Data.Fin
+open import Function.Inverse using (_↔_)
+open import Function.Related as Related hiding (_∼[_]_)
+open import Data.List.Base as List using (List; []; _∷_)
+open import Data.List.Any using (Any; map)
+import Data.List.Any.Membership as Membership
+open import Data.Product as Prod using (∃; _×_; _,_; uncurry′; proj₂)
+open import Relation.Nullary
+open import Relation.Binary hiding (Decidable)
+import Relation.Binary.InducedPreorders as Ind
+open import Relation.Binary.List.Pointwise as ListEq using ([]; _∷_)
+open import Relation.Binary.PropositionalEquality as PropEq
+ using (_≡_)
+
+private module M {a} {A : Set a} = Membership (PropEq.setoid A)
+open M public hiding (lose)
+
+lose : ∀ {a p} {A : Set a} {P : A → Set p} {x xs} →
+ x ∈ xs → P x → Any P xs
+lose {P = P} = M.lose (PropEq.subst P)
+
+-- _⊆_ is a preorder.
+
+⊆-preorder : ∀ {a} → Set a → Preorder _ _ _
+⊆-preorder A = Ind.InducedPreorder₂ (PropEq.setoid (List A)) _∈_
+ (PropEq.subst (_∈_ _))
+
+-- Set and bag equality and related preorders.
+
+open Related public
+ using (Kind; Symmetric-kind)
+ renaming ( implication to subset
+ ; reverse-implication to superset
+ ; equivalence to set
+ ; injection to subbag
+ ; reverse-injection to superbag
+ ; bijection to bag
+ )
+
+[_]-Order : Kind → ∀ {a} → Set a → Preorder _ _ _
+[ k ]-Order A = Related.InducedPreorder₂ k (_∈_ {A = A})
+
+[_]-Equality : Symmetric-kind → ∀ {a} → Set a → Setoid _ _
+[ k ]-Equality A = Related.InducedEquivalence₂ k (_∈_ {A = A})
+
+infix 4 _∼[_]_
+
+_∼[_]_ : ∀ {a} {A : Set a} → List A → Kind → List A → Set _
+_∼[_]_ {A = A} xs k ys = Preorder._∼_ ([ k ]-Order A) xs ys
+
+-- Bag equality implies the other relations.
+
+bag-=⇒ : ∀ {k a} {A : Set a} {xs ys : List A} →
+ xs ∼[ bag ] ys → xs ∼[ k ] ys
+bag-=⇒ xs≈ys = ↔⇒ xs≈ys
+
+-- "Equational" reasoning for _⊆_.
+
+module ⊆-Reasoning where
+ import Relation.Binary.PreorderReasoning as PreR
+ private
+ open module PR {a} {A : Set a} = PreR (⊆-preorder A) public
+ renaming (_∼⟨_⟩_ to _⊆⟨_⟩_; _≈⟨_⟩_ to _≡⟨_⟩_)
+
+ infixr 2 _∼⟨_⟩_
+ infix 1 _∈⟨_⟩_
+
+ _∈⟨_⟩_ : ∀ {a} {A : Set a} x {xs ys : List A} →
+ x ∈ xs → xs IsRelatedTo ys → x ∈ ys
+ x ∈⟨ x∈xs ⟩ xs⊆ys = (begin xs⊆ys) x∈xs
+
+ _∼⟨_⟩_ : ∀ {k a} {A : Set a} xs {ys zs : List A} →
+ xs ∼[ ⌊ k ⌋→ ] ys → ys IsRelatedTo zs → xs IsRelatedTo zs
+ xs ∼⟨ xs≈ys ⟩ ys≈zs = xs ⊆⟨ ⇒→ xs≈ys ⟩ ys≈zs
diff --git a/src/Data/List/Any/Membership/Propositional/Properties.agda b/src/Data/List/Any/Membership/Propositional/Properties.agda
new file mode 100644
index 0000000..6b86928
--- /dev/null
+++ b/src/Data/List/Any/Membership/Propositional/Properties.agda
@@ -0,0 +1,286 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties related to propositional list membership
+------------------------------------------------------------------------
+
+-- This module does not treat the general variant of list membership,
+-- parametrised on a setoid, only the variant where the equality is
+-- fixed to be propositional equality.
+
+module Data.List.Any.Membership.Propositional.Properties where
+
+open import Algebra
+open import Category.Monad
+open import Data.Bool.Base using (Bool; false; true; T)
+open import Data.Empty
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence using (module Equivalence)
+import Function.Injection as Inj
+open import Function.Inverse as Inv using (_↔_; module Inverse)
+import Function.Related as Related
+open import Function.Related.TypeIsomorphisms
+open import Data.List as List
+open import Data.List.Any as Any using (Any; here; there)
+open import Data.List.Any.Properties
+open import Data.List.Any.Membership.Propositional
+import Data.List.Any.Membership.Properties as Membershipₚ
+open import Data.Nat as Nat
+open import Data.Nat.Properties
+open import Data.Product as Prod
+open import Data.Sum as Sum
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; refl; _≗_)
+import Relation.Binary.Properties.DecTotalOrder as DTOProperties
+import Relation.Binary.Sigma.Pointwise as Σ
+open import Relation.Unary using (_⟨×⟩_)
+open import Relation.Nullary
+open import Relation.Nullary.Negation
+
+private
+ module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
+ open module ListMonad {ℓ} = RawMonad (List.monad {ℓ = ℓ})
+
+------------------------------------------------------------------------
+-- Properties relating _∈_ to various list functions
+------------------------------------------------------------------------
+-- map
+
+module _ {a b} {A : Set a} {B : Set b} {f : A → B} where
+
+ ∈-map⁺ : ∀ {x xs} → x ∈ xs → f x ∈ List.map f xs
+ ∈-map⁺ = Membershipₚ.∈-map⁺ (P.setoid _) (P.setoid _) (P.cong f)
+
+ ∈-map⁻ : ∀ {y xs} → y ∈ List.map f xs → ∃ λ x → x ∈ xs × y ≡ f x
+ ∈-map⁻ = Membershipₚ.∈-map⁻ (P.setoid _) (P.setoid _)
+
+ map-∈↔ : ∀ {y xs} →
+ (∃ λ x → x ∈ xs × y ≡ f x) ↔ y ∈ List.map f xs
+ map-∈↔ {y} {xs} =
+ (∃ λ x → x ∈ xs × y ≡ f x) ↔⟨ Any↔ ⟩
+ Any (λ x → y ≡ f x) xs ↔⟨ map↔ ⟩
+ y ∈ List.map f xs ∎
+ where open Related.EquationalReasoning
+
+------------------------------------------------------------------------
+-- concat
+
+concat-∈↔ : ∀ {a} {A : Set a} {x : A} {xss} →
+ (∃ λ xs → x ∈ xs × xs ∈ xss) ↔ x ∈ concat xss
+concat-∈↔ {a} {x = x} {xss} =
+ (∃ λ xs → x ∈ xs × xs ∈ xss) ↔⟨ Σ.cong {a₁ = a} {b₁ = a} {b₂ = a} Inv.id $ ×⊎.*-comm _ _ ⟩
+ (∃ λ xs → xs ∈ xss × x ∈ xs) ↔⟨ Any↔ {a = a} {p = a} ⟩
+ Any (Any (_≡_ x)) xss ↔⟨ concat↔ {a = a} {p = a} ⟩
+ x ∈ concat xss ∎
+ where open Related.EquationalReasoning
+
+------------------------------------------------------------------------
+-- filter
+
+filter-∈ : ∀ {a} {A : Set a} (p : A → Bool) (xs : List A) {x} →
+ x ∈ xs → p x ≡ true → x ∈ filter p xs
+filter-∈ p [] () _
+filter-∈ p (x ∷ xs) (here refl) px≡true rewrite px≡true = here refl
+filter-∈ p (y ∷ xs) (there pxs) px≡true with p y
+... | true = there (filter-∈ p xs pxs px≡true)
+... | false = filter-∈ p xs pxs px≡true
+
+------------------------------------------------------------------------
+-- Other monad functions
+
+>>=-∈↔ : ∀ {ℓ} {A B : Set ℓ} {xs} {f : A → List B} {y} →
+ (∃ λ x → x ∈ xs × y ∈ f x) ↔ y ∈ (xs >>= f)
+>>=-∈↔ {ℓ} {xs = xs} {f} {y} =
+ (∃ λ x → x ∈ xs × y ∈ f x) ↔⟨ Any↔ {a = ℓ} {p = ℓ} ⟩
+ Any (Any (_≡_ y) ∘ f) xs ↔⟨ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
+ y ∈ (xs >>= f) ∎
+ where open Related.EquationalReasoning
+
+⊛-∈↔ : ∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) {xs y} →
+ (∃₂ λ f x → f ∈ fs × x ∈ xs × y ≡ f x) ↔ y ∈ (fs ⊛ xs)
+⊛-∈↔ {ℓ} fs {xs} {y} =
+ (∃₂ λ f x → f ∈ fs × x ∈ xs × y ≡ f x) ↔⟨ Σ.cong {a₁ = ℓ} {b₁ = ℓ} {b₂ = ℓ} Inv.id (∃∃↔∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
+ (∃ λ f → f ∈ fs × ∃ λ x → x ∈ xs × y ≡ f x) ↔⟨ Σ.cong {a₁ = ℓ} {b₁ = ℓ} {b₂ = ℓ}
+ Inv.id ((_ ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any↔ {a = ℓ} {p = ℓ}) ⟩
+ (∃ λ f → f ∈ fs × Any (_≡_ y ∘ f) xs) ↔⟨ Any↔ {a = ℓ} {p = ℓ} ⟩
+ Any (λ f → Any (_≡_ y ∘ f) xs) fs ↔⟨ ⊛↔ ⟩
+ y ∈ (fs ⊛ xs) ∎
+ where open Related.EquationalReasoning
+
+⊗-∈↔ : ∀ {A B : Set} {xs ys} {x : A} {y : B} →
+ (x ∈ xs × y ∈ ys) ↔ (x , y) ∈ (xs ⊗ ys)
+⊗-∈↔ {A} {B} {xs} {ys} {x} {y} =
+ (x ∈ xs × y ∈ ys) ↔⟨ ⊗↔′ ⟩
+ Any (_≡_ x ⟨×⟩ _≡_ y) (xs ⊗ ys) ↔⟨ Any-cong helper (_ ∎) ⟩
+ (x , y) ∈ (xs ⊗ ys) ∎
+ where
+ open Related.EquationalReasoning
+
+ helper : (p : A × B) → (x ≡ proj₁ p × y ≡ proj₂ p) ↔ (x , y) ≡ p
+ helper (x′ , y′) = record
+ { to = P.→-to-⟶ (uncurry $ P.cong₂ _,_)
+ ; from = P.→-to-⟶ < P.cong proj₁ , P.cong proj₂ >
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.cong₂ _,_ (P.proof-irrelevance _ _)
+ (P.proof-irrelevance _ _)
+ ; right-inverse-of = λ _ → P.proof-irrelevance _ _
+ }
+ }
+
+------------------------------------------------------------------------
+-- Properties relating _∈_ to various list functions
+
+-- Various functions are monotone.
+
+mono : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
+ xs ⊆ ys → Any P xs → Any P ys
+mono xs⊆ys =
+ _⟨$⟩_ (Inverse.to Any↔) ∘′
+ Prod.map id (Prod.map xs⊆ys id) ∘
+ _⟨$⟩_ (Inverse.from Any↔)
+
+map-mono : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) {xs ys} →
+ xs ⊆ ys → List.map f xs ⊆ List.map f ys
+map-mono f xs⊆ys =
+ _⟨$⟩_ (Inverse.to map-∈↔) ∘
+ Prod.map id (Prod.map xs⊆ys id) ∘
+ _⟨$⟩_ (Inverse.from map-∈↔)
+
+_++-mono_ : ∀ {a} {A : Set a} {xs₁ xs₂ ys₁ ys₂ : List A} →
+ xs₁ ⊆ ys₁ → xs₂ ⊆ ys₂ → xs₁ ++ xs₂ ⊆ ys₁ ++ ys₂
+_++-mono_ xs₁⊆ys₁ xs₂⊆ys₂ =
+ _⟨$⟩_ (Inverse.to ++↔) ∘
+ Sum.map xs₁⊆ys₁ xs₂⊆ys₂ ∘
+ _⟨$⟩_ (Inverse.from ++↔)
+
+concat-mono : ∀ {a} {A : Set a} {xss yss : List (List A)} →
+ xss ⊆ yss → concat xss ⊆ concat yss
+concat-mono {a} xss⊆yss =
+ _⟨$⟩_ (Inverse.to $ concat-∈↔ {a = a}) ∘
+ Prod.map id (Prod.map id xss⊆yss) ∘
+ _⟨$⟩_ (Inverse.from $ concat-∈↔ {a = a})
+
+>>=-mono : ∀ {ℓ} {A B : Set ℓ} (f g : A → List B) {xs ys} →
+ xs ⊆ ys → (∀ {x} → f x ⊆ g x) →
+ (xs >>= f) ⊆ (ys >>= g)
+>>=-mono {ℓ} f g xs⊆ys f⊆g =
+ _⟨$⟩_ (Inverse.to $ >>=-∈↔ {ℓ = ℓ}) ∘
+ Prod.map id (Prod.map xs⊆ys f⊆g) ∘
+ _⟨$⟩_ (Inverse.from $ >>=-∈↔ {ℓ = ℓ})
+
+_⊛-mono_ : ∀ {ℓ} {A B : Set ℓ}
+ {fs gs : List (A → B)} {xs ys : List A} →
+ fs ⊆ gs → xs ⊆ ys → (fs ⊛ xs) ⊆ (gs ⊛ ys)
+_⊛-mono_ {fs = fs} {gs} fs⊆gs xs⊆ys =
+ _⟨$⟩_ (Inverse.to $ ⊛-∈↔ gs) ∘
+ Prod.map id (Prod.map id (Prod.map fs⊆gs (Prod.map xs⊆ys id))) ∘
+ _⟨$⟩_ (Inverse.from $ ⊛-∈↔ fs)
+
+_⊗-mono_ : {A B : Set} {xs₁ ys₁ : List A} {xs₂ ys₂ : List B} →
+ xs₁ ⊆ ys₁ → xs₂ ⊆ ys₂ → (xs₁ ⊗ xs₂) ⊆ (ys₁ ⊗ ys₂)
+xs₁⊆ys₁ ⊗-mono xs₂⊆ys₂ =
+ _⟨$⟩_ (Inverse.to ⊗-∈↔) ∘
+ Prod.map xs₁⊆ys₁ xs₂⊆ys₂ ∘
+ _⟨$⟩_ (Inverse.from ⊗-∈↔)
+
+any-mono : ∀ {a} {A : Set a} (p : A → Bool) →
+ ∀ {xs ys} → xs ⊆ ys → T (any p xs) → T (any p ys)
+any-mono {a} p xs⊆ys =
+ _⟨$⟩_ (Equivalence.to $ any⇔ {a = a}) ∘
+ mono xs⊆ys ∘
+ _⟨$⟩_ (Equivalence.from $ any⇔ {a = a})
+
+map-with-∈-mono :
+ ∀ {a b} {A : Set a} {B : Set b}
+ {xs : List A} {f : ∀ {x} → x ∈ xs → B}
+ {ys : List A} {g : ∀ {x} → x ∈ ys → B}
+ (xs⊆ys : xs ⊆ ys) → (∀ {x} → f {x} ≗ g ∘ xs⊆ys) →
+ map-with-∈ xs f ⊆ map-with-∈ ys g
+map-with-∈-mono {f = f} {g = g} xs⊆ys f≈g {x} =
+ _⟨$⟩_ (Inverse.to map-with-∈↔) ∘
+ Prod.map id (Prod.map xs⊆ys (λ {x∈xs} x≡fx∈xs → begin
+ x ≡⟨ x≡fx∈xs ⟩
+ f x∈xs ≡⟨ f≈g x∈xs ⟩
+ g (xs⊆ys x∈xs) ∎)) ∘
+ _⟨$⟩_ (Inverse.from map-with-∈↔)
+ where open P.≡-Reasoning
+
+-- Other properties.
+
+filter-⊆ : ∀ {a} {A : Set a} (p : A → Bool) →
+ (xs : List A) → filter p xs ⊆ xs
+filter-⊆ _ [] = λ ()
+filter-⊆ p (x ∷ xs) with p x | filter-⊆ p xs
+... | false | hyp = there ∘ hyp
+... | true | hyp =
+ λ { (here eq) → here eq
+ ; (there ∈filter) → there (hyp ∈filter)
+ }
+
+------------------------------------------------------------------------
+-- Other properties
+
+-- Only a finite number of distinct elements can be members of a
+-- given list.
+
+finite : ∀ {a} {A : Set a}
+ (f : Inj.Injection (P.setoid ℕ) (P.setoid A)) →
+ ∀ xs → ¬ (∀ i → Inj.Injection.to f ⟨$⟩ i ∈ xs)
+finite inj [] ∈[] with ∈[] zero
+... | ()
+finite {A = A} inj (x ∷ xs) ∈x∷xs = excluded-middle helper
+ where
+ open Inj.Injection inj
+
+ module STO = StrictTotalOrder
+ (DTOProperties.strictTotalOrder ≤-decTotalOrder)
+
+ not-x : ∀ {i} → ¬ (to ⟨$⟩ i ≡ x) → to ⟨$⟩ i ∈ xs
+ not-x {i} ≢x with ∈x∷xs i
+ ... | here ≡x = ⊥-elim (≢x ≡x)
+ ... | there ∈xs = ∈xs
+
+ helper : ¬ Dec (∃ λ i → to ⟨$⟩ i ≡ x)
+ helper (no ≢x) = finite inj xs (λ i → not-x (≢x ∘ _,_ i))
+ helper (yes (i , ≡x)) = finite inj′ xs ∈xs
+ where
+ open P
+
+ f : ℕ → A
+ f j with STO.compare i j
+ f j | tri< _ _ _ = to ⟨$⟩ suc j
+ f j | tri≈ _ _ _ = to ⟨$⟩ suc j
+ f j | tri> _ _ _ = to ⟨$⟩ j
+
+ ∈-if-not-i : ∀ {j} → i ≢ j → to ⟨$⟩ j ∈ xs
+ ∈-if-not-i i≢j = not-x (i≢j ∘ injective ∘ trans ≡x ∘ sym)
+
+ lemma : ∀ {k j} → k ≤ j → suc j ≢ k
+ lemma 1+j≤j refl = 1+n≰n 1+j≤j
+
+ ∈xs : ∀ j → f j ∈ xs
+ ∈xs j with STO.compare i j
+ ∈xs j | tri< (i≤j , _) _ _ = ∈-if-not-i (lemma i≤j ∘ sym)
+ ∈xs j | tri> _ i≢j _ = ∈-if-not-i i≢j
+ ∈xs .i | tri≈ _ refl _ =
+ ∈-if-not-i (m≢1+m+n i ∘
+ subst (_≡_ i ∘ suc) (sym (+-identityʳ i)))
+
+ injective′ : Inj.Injective {B = P.setoid A} (→-to-⟶ f)
+ injective′ {j} {k} eq with STO.compare i j | STO.compare i k
+ ... | tri< _ _ _ | tri< _ _ _ = cong pred $ injective eq
+ ... | tri< _ _ _ | tri≈ _ _ _ = cong pred $ injective eq
+ ... | tri< (i≤j , _) _ _ | tri> _ _ (k≤i , _) = ⊥-elim (lemma (≤-trans k≤i i≤j) $ injective eq)
+ ... | tri≈ _ _ _ | tri< _ _ _ = cong pred $ injective eq
+ ... | tri≈ _ _ _ | tri≈ _ _ _ = cong pred $ injective eq
+ ... | tri≈ _ i≡j _ | tri> _ _ (k≤i , _) = ⊥-elim (lemma (subst (_≤_ k) i≡j k≤i) $ injective eq)
+ ... | tri> _ _ (j≤i , _) | tri< (i≤k , _) _ _ = ⊥-elim (lemma (≤-trans j≤i i≤k) $ sym $ injective eq)
+ ... | tri> _ _ (j≤i , _) | tri≈ _ i≡k _ = ⊥-elim (lemma (subst (_≤_ j) i≡k j≤i) $ sym $ injective eq)
+ ... | tri> _ _ (j≤i , _) | tri> _ _ (k≤i , _) = injective eq
+
+ inj′ = record
+ { to = →-to-⟶ {B = P.setoid A} f
+ ; injective = injective′
+ }
diff --git a/src/Data/List/Any/Properties.agda b/src/Data/List/Any/Properties.agda
index 3586ecb..4cf1fa1 100644
--- a/src/Data/List/Any/Properties.agda
+++ b/src/Data/List/Any/Properties.agda
@@ -10,40 +10,51 @@
module Data.List.Any.Properties where
open import Algebra
-import Algebra.FunctionProperties as FP
open import Category.Monad
open import Data.Bool.Base using (Bool; false; true; T)
open import Data.Bool.Properties
-open import Data.Empty
+open import Data.Empty using (⊥)
+open import Data.Fin using (Fin) renaming (zero to fzero; suc to fsuc)
open import Data.List as List
open import Data.List.Any as Any using (Any; here; there)
-open import Data.Product as Prod hiding (swap)
+open import Data.List.Any.Membership.Propositional
+open import Data.Nat using (zero; suc; _<_; z≤n; s≤s)
+open import Data.Product as Prod
+ using (_×_; _,_; ∃; ∃₂; proj₁; proj₂; uncurry′)
open import Data.Sum as Sum using (_⊎_; inj₁; inj₂; [_,_]′)
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 using (Related)
-open import Function.Related.TypeIsomorphisms
-open import Level
open import Relation.Binary
-import Relation.Binary.HeterogeneousEquality as H
open import Relation.Binary.Product.Pointwise
open import Relation.Binary.PropositionalEquality as P
- using (_≡_; refl; inspect) renaming ([_] to P[_])
-open import Relation.Unary using (_⟨×⟩_; _⟨→⟩_) renaming (_⊆_ to _⋐_)
+ using (_≡_; refl; inspect)
+open import Relation.Unary
+ using (Pred; _⟨×⟩_; _⟨→⟩_) renaming (_⊆_ to _⋐_)
import Relation.Binary.Sigma.Pointwise as Σ
open import Relation.Binary.Sum
+open import Relation.Binary.List.Pointwise
+ using ([]; _∷_) renaming (Rel to ListRel)
-open Any.Membership-≡
open Related.EquationalReasoning
private
- module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
open module ListMonad {ℓ} = RawMonad (List.monad {ℓ = ℓ})
------------------------------------------------------------------------
--- Some lemmas related to map, find and lose
+-- If a predicate P respects the underlying equality then Any P
+-- respects the list equality.
+
+lift-resp : ∀ {a p ℓ} {A : Set a} {P : A → Set p} {_≈_ : Rel A ℓ} →
+ P Respects _≈_ → Any P Respects (ListRel _≈_)
+lift-resp resp [] ()
+lift-resp resp (x≈y ∷ xs≈ys) (here px) = here (resp x≈y px)
+lift-resp resp (x≈y ∷ xs≈ys) (there pxs) =
+ there (lift-resp resp xs≈ys pxs)
+------------------------------------------------------------------------
+-- Some lemmas related to map, find and lose
-- Any.map is functorial.
map-id : ∀ {a p} {A : Set a} {P : A → Set p} (f : P ⋐ P) {xs} →
@@ -84,29 +95,31 @@ find-∈ : ∀ {a} {A : Set a} {x : A} {xs : List A} (x∈xs : x ∈ xs) →
find-∈ (here refl) = refl
find-∈ (there x∈xs) rewrite find-∈ x∈xs = refl
-private
-
- -- find and lose are inverses (more or less).
+-- find and lose are inverses (more or less).
- lose∘find : ∀ {a p} {A : Set a} {P : A → Set p} {xs : List A}
- (p : Any P xs) →
- uncurry′ lose (proj₂ (find p)) ≡ p
- lose∘find p = map∘find p P.refl
+lose∘find : ∀ {a p} {A : Set a} {P : A → Set p} {xs : List A}
+ (p : Any P xs) →
+ uncurry′ lose (proj₂ (find p)) ≡ p
+lose∘find p = map∘find p P.refl
- find∘lose : ∀ {a p} {A : Set a} (P : A → Set p) {x xs}
- (x∈xs : x ∈ xs) (pp : P x) →
- find {P = P} (lose x∈xs pp) ≡ (x , x∈xs , pp)
- find∘lose P x∈xs p
- rewrite find∘map x∈xs (flip (P.subst P) p)
- | find-∈ x∈xs
- = refl
+find∘lose : ∀ {a p} {A : Set a} (P : A → Set p) {x xs}
+ (x∈xs : x ∈ xs) (pp : P x) →
+ find {P = P} (lose x∈xs pp) ≡ (x , x∈xs , pp)
+find∘lose P x∈xs p
+ rewrite find∘map x∈xs (flip (P.subst P) p)
+ | find-∈ x∈xs
+ = refl
-- Any can be expressed using _∈_.
+∃∈-Any : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
+ (∃ λ x → x ∈ xs × P x) → Any P xs
+∃∈-Any = uncurry′ lose ∘ proj₂
+
Any↔ : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
(∃ λ x → x ∈ xs × P x) ↔ Any P xs
Any↔ {P = P} {xs} = record
- { to = P.→-to-⟶ to
+ { to = P.→-to-⟶ ∃∈-Any
; from = P.→-to-⟶ (find {P = P})
; inverse-of = record
{ left-inverse-of = λ p →
@@ -114,9 +127,6 @@ Any↔ {P = P} {xs} = record
; right-inverse-of = lose∘find
}
}
- where
- to : (∃ λ x → x ∈ xs × P x) → Any P xs
- to = uncurry′ lose ∘ proj₂
------------------------------------------------------------------------
-- Any is a congruence
@@ -135,22 +145,36 @@ Any-cong {P₁ = P₁} {P₂} {xs₁} {xs₂} P₁↔P₂ xs₁≈xs₂ =
-- Nested occurrences of Any can sometimes be swapped. See also ×↔.
-swap : ∀ {ℓ} {A B : Set ℓ} {P : A → B → Set ℓ} {xs ys} →
+swap : ∀ {a b ℓ} {A : Set a} {B : Set b} {P : A → B → Set ℓ} {xs ys} →
+ Any (λ x → Any (P x) ys) xs → Any (λ y → Any (flip P y) xs) ys
+swap (here pys) = Any.map here pys
+swap (there pxys) = Any.map there (swap pxys)
+
+swap-there : ∀ {a b ℓ} {A : Set a} {B : Set b} {P : A → B → Set ℓ}
+ {x xs ys} → (any : Any (λ x → Any (P x) ys) xs) →
+ swap (Any.map (there {x = x}) any) ≡ there (swap any)
+swap-there (here pys) = refl
+swap-there (there pxys) = P.cong (Any.map there) (swap-there pxys)
+
+swap-invol : ∀ {a b ℓ} {A : Set a} {B : Set b} {P : A → B → Set ℓ}
+ {xs ys} → (any : Any (λ x → Any (P x) ys) xs) →
+ swap (swap any) ≡ any
+swap-invol (here (here px)) = refl
+swap-invol (here (there pys)) =
+ P.cong (Any.map there) (swap-invol (here pys))
+swap-invol (there pxys) =
+ P.trans (swap-there (swap pxys)) (P.cong there (swap-invol pxys))
+
+swap↔ : ∀ {ℓ} {A B : Set ℓ} {P : A → B → Set ℓ} {xs ys} →
Any (λ x → Any (P x) ys) xs ↔ Any (λ y → Any (flip P y) xs) ys
-swap {ℓ} {P = P} {xs} {ys} =
- Any (λ x → Any (P x) ys) xs ↔⟨ sym $ Any↔ {a = ℓ} {p = ℓ} ⟩
- (∃ λ x → x ∈ xs × Any (P x) ys) ↔⟨ sym $ Σ.cong Inv.id (λ {x} → (x ∈ xs ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any↔ {a = ℓ} {p = ℓ}) ⟩
- (∃ λ x → x ∈ xs × ∃ λ y → y ∈ ys × P x y) ↔⟨ Σ.cong {a₁ = ℓ} Inv.id (∃∃↔∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
- (∃₂ λ x y → x ∈ xs × y ∈ ys × P x y) ↔⟨ ∃∃↔∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _ ⟩
- (∃₂ λ y x → x ∈ xs × y ∈ ys × P x y) ↔⟨ Σ.cong Inv.id (λ {y} → Σ.cong Inv.id (λ {x} →
- (x ∈ xs × y ∈ ys × P x y) ↔⟨ sym $ ×⊎.*-assoc _ _ _ ⟩
- ((x ∈ xs × y ∈ ys) × P x y) ↔⟨ ×⊎.*-comm (x ∈ xs) (y ∈ ys) ⟨ ×⊎.*-cong ⟩ (P x y ∎) ⟩
- ((y ∈ ys × x ∈ xs) × P x y) ↔⟨ ×⊎.*-assoc _ _ _ ⟩
- (y ∈ ys × x ∈ xs × P x y) ∎)) ⟩
- (∃₂ λ y x → y ∈ ys × x ∈ xs × P x y) ↔⟨ Σ.cong {a₁ = ℓ} Inv.id (∃∃↔∃∃ {a = ℓ} {b = ℓ} {p = ℓ} _) ⟩
- (∃ λ y → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ↔⟨ Σ.cong Inv.id (λ {y} → (y ∈ ys ∎) ⟨ ×⊎.*-cong {ℓ = ℓ} ⟩ Any↔ {a = ℓ} {p = ℓ}) ⟩
- (∃ λ y → y ∈ ys × Any (flip P y) xs) ↔⟨ Any↔ {a = ℓ} {p = ℓ} ⟩
- Any (λ y → Any (flip P y) xs) ys ∎
+swap↔ {P = P} = record
+ { to = P.→-to-⟶ swap
+ ; from = P.→-to-⟶ swap
+ ; inverse-of = record
+ { left-inverse-of = swap-invol
+ ; right-inverse-of = swap-invol
+ }
+ }
------------------------------------------------------------------------
-- Lemmas relating Any to ⊥
@@ -180,165 +204,182 @@ swap {ℓ} {P = P} {xs} {ys} =
}
------------------------------------------------------------------------
--- Lemmas relating Any to sums and products
+-- Lemmas relating Any to ⊤
--- Sums commute with Any (for a fixed list).
+-- These introduction and elimination rules are not inverses, though.
-⊎↔ : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} {xs} →
- (Any P xs ⊎ Any Q xs) ↔ Any (λ x → P x ⊎ Q x) xs
-⊎↔ {P = P} {Q} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = from∘to
- ; right-inverse-of = to∘from
+module _ {a} {A : Set a} where
+
+ any⁺ : ∀ (p : A → Bool) {xs} → Any (T ∘ p) xs → T (any p xs)
+ any⁺ p (here px) = Equivalence.from T-∨ ⟨$⟩ inj₁ px
+ any⁺ p (there {x = x} pxs) with p x
+ ... | true = _
+ ... | false = any⁺ p pxs
+
+ any⁻ : ∀ (p : A → Bool) xs → T (any p xs) → Any (T ∘ p) xs
+ any⁻ p [] ()
+ any⁻ p (x ∷ xs) px∷xs with p x | inspect p x
+ any⁻ p (x ∷ xs) px∷xs | true | P.[ eq ] = here (Equivalence.from T-≡ ⟨$⟩ eq)
+ any⁻ p (x ∷ xs) px∷xs | false | _ = there (any⁻ p xs px∷xs)
+
+ any⇔ : ∀ {p : A → Bool} {xs} → Any (T ∘ p) xs ⇔ T (any p xs)
+ any⇔ = Eq.equivalence (any⁺ _) (any⁻ _ _)
+
+------------------------------------------------------------------------
+-- Sums commute with Any
+
+module _ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} where
+
+ Any-⊎⁺ : ∀ {xs} → Any P xs ⊎ Any Q xs → Any (λ x → P x ⊎ Q x) xs
+ Any-⊎⁺ = [ Any.map inj₁ , Any.map inj₂ ]′
+
+ Any-⊎⁻ : ∀ {xs} → Any (λ x → P x ⊎ Q x) xs → Any P xs ⊎ Any Q xs
+ Any-⊎⁻ (here (inj₁ p)) = inj₁ (here p)
+ Any-⊎⁻ (here (inj₂ q)) = inj₂ (here q)
+ Any-⊎⁻ (there p) = Sum.map there there (Any-⊎⁻ p)
+
+ ⊎↔ : ∀ {xs} → (Any P xs ⊎ Any Q xs) ↔ Any (λ x → P x ⊎ Q x) xs
+ ⊎↔ = record
+ { to = P.→-to-⟶ Any-⊎⁺
+ ; from = P.→-to-⟶ Any-⊎⁻
+ ; inverse-of = record
+ { left-inverse-of = from∘to
+ ; right-inverse-of = to∘from
+ }
}
- }
- where
- to : ∀ {xs} → Any P xs ⊎ Any Q xs → Any (λ x → P x ⊎ Q x) xs
- to = [ Any.map inj₁ , Any.map inj₂ ]′
-
- from : ∀ {xs} → Any (λ x → P x ⊎ Q x) xs → Any P xs ⊎ Any Q xs
- from (here (inj₁ p)) = inj₁ (here p)
- from (here (inj₂ q)) = inj₂ (here q)
- from (there p) = Sum.map there there (from p)
-
- from∘to : ∀ {xs} (p : Any P xs ⊎ Any Q xs) → from (to p) ≡ p
- from∘to (inj₁ (here p)) = P.refl
- from∘to (inj₁ (there p)) rewrite from∘to (inj₁ p) = P.refl
- from∘to (inj₂ (here q)) = P.refl
- from∘to (inj₂ (there q)) rewrite from∘to (inj₂ q) = P.refl
-
- to∘from : ∀ {xs} (p : Any (λ x → P x ⊎ Q x) xs) →
- to (from p) ≡ p
- to∘from (here (inj₁ p)) = P.refl
- to∘from (here (inj₂ q)) = P.refl
- to∘from (there p) with from p | to∘from p
- to∘from (there .(Any.map inj₁ p)) | inj₁ p | P.refl = P.refl
- to∘from (there .(Any.map inj₂ q)) | inj₂ q | P.refl = P.refl
+ where
+ from∘to : ∀ {xs} (p : Any P xs ⊎ Any Q xs) → Any-⊎⁻ (Any-⊎⁺ p) ≡ p
+ from∘to (inj₁ (here p)) = P.refl
+ from∘to (inj₁ (there p)) rewrite from∘to (inj₁ p) = P.refl
+ from∘to (inj₂ (here q)) = P.refl
+ from∘to (inj₂ (there q)) rewrite from∘to (inj₂ q) = P.refl
+
+ to∘from : ∀ {xs} (p : Any (λ x → P x ⊎ Q x) xs) →
+ Any-⊎⁺ (Any-⊎⁻ p) ≡ p
+ to∘from (here (inj₁ p)) = P.refl
+ to∘from (here (inj₂ q)) = P.refl
+ to∘from (there p) with Any-⊎⁻ p | to∘from p
+ to∘from (there .(Any.map inj₁ p)) | inj₁ p | P.refl = P.refl
+ to∘from (there .(Any.map inj₂ q)) | inj₂ q | P.refl = P.refl
+
+------------------------------------------------------------------------
-- Products "commute" with Any.
-×↔ : {A B : Set} {P : A → Set} {Q : B → Set}
- {xs : List A} {ys : List B} →
- (Any P xs × Any Q ys) ↔ Any (λ x → Any (λ y → P x × Q y) ys) xs
-×↔ {P = P} {Q} {xs} {ys} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = from∘to
- ; right-inverse-of = to∘from
- }
- }
- where
- to : Any P xs × Any Q ys → Any (λ x → Any (λ y → P x × Q y) ys) xs
- to (p , q) = Any.map (λ p → Any.map (λ q → (p , q)) q) p
+module _ {a b p q} {A : Set a} {B : Set b}
+ {P : Pred A p} {Q : Pred B q} where
+
+ Any-×⁺ : ∀ {xs ys} → Any P xs × Any Q ys →
+ Any (λ x → Any (λ y → P x × Q y) ys) xs
+ Any-×⁺ (p , q) = Any.map (λ p → Any.map (λ q → (p , q)) q) p
- from : Any (λ x → Any (λ y → P x × Q y) ys) xs → Any P xs × Any Q ys
- from pq with Prod.map id (Prod.map id find) (find pq)
+ Any-×⁻ : ∀ {xs ys} → Any (λ x → Any (λ y → P x × Q y) ys) xs →
+ Any P xs × Any Q ys
+ Any-×⁻ pq with Prod.map id (Prod.map id find) (find pq)
... | (x , x∈xs , y , y∈ys , p , q) = (lose x∈xs p , lose y∈ys q)
- from∘to : ∀ pq → from (to pq) ≡ pq
- from∘to (p , q)
- rewrite find∘map {Q = λ x → Any (λ y → P x × Q y) ys}
- p (λ p → Any.map (λ q → (p , q)) q)
- | find∘map {Q = λ y → P (proj₁ (find p)) × Q y}
- q (λ q → proj₂ (proj₂ (find p)) , q)
- | lose∘find p
- | lose∘find q
- = refl
-
- to∘from : ∀ pq → to (from pq) ≡ pq
- to∘from pq
- with find pq
- | (λ (f : _≡_ (proj₁ (find pq)) ⋐ _) → map∘find pq {f})
- ... | (x , x∈xs , pq′) | lem₁
- with find pq′
- | (λ (f : _≡_ (proj₁ (find pq′)) ⋐ _) → map∘find pq′ {f})
- ... | (y , y∈ys , p , q) | lem₂
- rewrite P.sym $ map-∘ {R = λ x → Any (λ y → P x × Q y) ys}
- (λ p → Any.map (λ q → p , q) (lose y∈ys q))
- (λ y → P.subst P y p)
- x∈xs
- = lem₁ _ helper
+ ×↔ : ∀ {xs ys} →
+ (Any P xs × Any Q ys) ↔ Any (λ x → Any (λ y → P x × Q y) ys) xs
+ ×↔ {xs} {ys} = record
+ { to = P.→-to-⟶ Any-×⁺
+ ; from = P.→-to-⟶ Any-×⁻
+ ; inverse-of = record
+ { left-inverse-of = from∘to
+ ; right-inverse-of = to∘from
+ }
+ }
where
- helper : Any.map (λ q → p , q) (lose y∈ys q) ≡ pq′
- helper rewrite P.sym $ map-∘ {R = λ y → P x × Q y}
- (λ q → p , q)
- (λ y → P.subst Q y q)
- y∈ys
- = lem₂ _ refl
+ from∘to : ∀ pq → Any-×⁻ (Any-×⁺ pq) ≡ pq
+ from∘to (p , q)
+ rewrite find∘map {Q = λ x → Any (λ y → P x × Q y) ys}
+ p (λ p → Any.map (λ q → (p , q)) q)
+ | find∘map {Q = λ y → P (proj₁ (find p)) × Q y}
+ q (λ q → proj₂ (proj₂ (find p)) , q)
+ | lose∘find p
+ | lose∘find q
+ = refl
+
+ to∘from : ∀ pq → Any-×⁺ (Any-×⁻ pq) ≡ pq
+ to∘from pq
+ with find pq
+ | (λ (f : (proj₁ (find pq) ≡_) ⋐ _) → map∘find pq {f})
+ ... | (x , x∈xs , pq′) | lem₁
+ with find pq′
+ | (λ (f : (proj₁ (find pq′) ≡_) ⋐ _) → map∘find pq′ {f})
+ ... | (y , y∈ys , p , q) | lem₂
+ rewrite P.sym $ map-∘ {R = λ x → Any (λ y → P x × Q y) ys}
+ (λ p → Any.map (λ q → p , q) (lose y∈ys q))
+ (λ y → P.subst P y p)
+ x∈xs
+ = lem₁ _ helper
+ where
+ helper : Any.map (λ q → p , q) (lose y∈ys q) ≡ pq′
+ helper rewrite P.sym $ map-∘ {R = λ y → P x × Q y}
+ (λ q → p , q)
+ (λ y → P.subst Q y q)
+ y∈ys
+ = lem₂ _ refl
------------------------------------------------------------------------
-- Invertible introduction (⁺) and elimination (⁻) rules for various
-- list functions
+------------------------------------------------------------------------
+-- map
--- map.
-
-private
+module _ {a b} {A : Set a} {B : Set b} where
- map⁺ : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
- {f : A → B} {xs} →
+ map⁺ : ∀ {p} {P : B → Set p} {f : A → B} {xs} →
Any (P ∘ f) xs → Any P (List.map f xs)
map⁺ (here p) = here p
map⁺ (there p) = there $ map⁺ p
- map⁻ : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
- {f : A → B} {xs} →
+ map⁻ : ∀ {p} {P : B → Set p} {f : A → B} {xs} →
Any P (List.map f xs) → Any (P ∘ f) xs
map⁻ {xs = []} ()
map⁻ {xs = x ∷ xs} (here p) = here p
map⁻ {xs = x ∷ xs} (there p) = there $ map⁻ p
- map⁺∘map⁻ : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
- {f : A → B} {xs} →
- (p : Any P (List.map f xs)) →
- map⁺ (map⁻ p) ≡ p
+ map⁺∘map⁻ : ∀ {p} {P : B → Set p} {f : A → B} {xs} →
+ (p : Any P (List.map f xs)) → map⁺ (map⁻ p) ≡ p
map⁺∘map⁻ {xs = []} ()
map⁺∘map⁻ {xs = x ∷ xs} (here p) = refl
map⁺∘map⁻ {xs = x ∷ xs} (there p) = P.cong there (map⁺∘map⁻ p)
- map⁻∘map⁺ : ∀ {a b p} {A : Set a} {B : Set b} (P : B → Set p)
- {f : A → B} {xs} →
- (p : Any (P ∘ f) xs) →
- map⁻ {P = P} (map⁺ p) ≡ p
+ map⁻∘map⁺ : ∀ {p} (P : B → Set p) {f : A → B} {xs} →
+ (p : Any (P ∘ f) xs) → map⁻ {P = P} (map⁺ p) ≡ p
map⁻∘map⁺ P (here p) = refl
map⁻∘map⁺ P (there p) = P.cong there (map⁻∘map⁺ P p)
-map↔ : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
- {f : A → B} {xs} →
- Any (P ∘ f) xs ↔ Any P (List.map f xs)
-map↔ {P = P} {f = f} = record
- { to = P.→-to-⟶ $ map⁺ {P = P} {f = f}
- ; from = P.→-to-⟶ $ map⁻ {P = P} {f = f}
- ; inverse-of = record
- { left-inverse-of = map⁻∘map⁺ P
- ; right-inverse-of = map⁺∘map⁻
+ map↔ : ∀ {p} {P : B → Set p} {f : A → B} {xs} →
+ Any (P ∘ f) xs ↔ Any P (List.map f xs)
+ map↔ {P = P} {f = f} = record
+ { to = P.→-to-⟶ $ map⁺ {P = P} {f = f}
+ ; from = P.→-to-⟶ $ map⁻ {P = P} {f = f}
+ ; inverse-of = record
+ { left-inverse-of = map⁻∘map⁺ P
+ ; right-inverse-of = map⁺∘map⁻
+ }
}
- }
--- _++_.
+------------------------------------------------------------------------
+-- _++_
-private
+module _ {a p} {A : Set a} {P : A → Set p} where
- ++⁺ˡ : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
- Any P xs → Any P (xs ++ ys)
+ ++⁺ˡ : ∀ {xs ys} → Any P xs → Any P (xs ++ ys)
++⁺ˡ (here p) = here p
++⁺ˡ (there p) = there (++⁺ˡ p)
- ++⁺ʳ : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys} →
- Any P ys → Any P (xs ++ ys)
+ ++⁺ʳ : ∀ xs {ys} → Any P ys → Any P (xs ++ ys)
++⁺ʳ [] p = p
++⁺ʳ (x ∷ xs) p = there (++⁺ʳ xs p)
- ++⁻ : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys} →
- Any P (xs ++ ys) → Any P xs ⊎ Any P ys
+ ++⁻ : ∀ xs {ys} → Any P (xs ++ ys) → Any P xs ⊎ Any P ys
++⁻ [] p = inj₂ p
++⁻ (x ∷ xs) (here p) = inj₁ (here p)
++⁻ (x ∷ xs) (there p) = Sum.map there id (++⁻ xs p)
- ++⁺∘++⁻ : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys}
- (p : Any P (xs ++ ys)) →
+ ++⁺∘++⁻ : ∀ xs {ys} (p : Any P (xs ++ ys)) →
[ ++⁺ˡ , ++⁺ʳ xs ]′ (++⁻ xs p) ≡ p
++⁺∘++⁻ [] p = refl
++⁺∘++⁻ (x ∷ xs) (here p) = refl
@@ -346,8 +387,7 @@ private
++⁺∘++⁻ (x ∷ xs) (there p) | inj₁ p′ | ih = P.cong there ih
++⁺∘++⁻ (x ∷ xs) (there p) | inj₂ p′ | ih = P.cong there ih
- ++⁻∘++⁺ : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys}
- (p : Any P xs ⊎ Any P ys) →
+ ++⁻∘++⁺ : ∀ xs {ys} (p : Any P xs ⊎ Any P ys) →
++⁻ xs ([ ++⁺ˡ , ++⁺ʳ xs ]′ p) ≡ p
++⁻∘++⁺ [] (inj₁ ())
++⁻∘++⁺ [] (inj₂ p) = refl
@@ -355,119 +395,216 @@ private
++⁻∘++⁺ (x ∷ xs) {ys} (inj₁ (there p)) rewrite ++⁻∘++⁺ xs {ys} (inj₁ p) = refl
++⁻∘++⁺ (x ∷ xs) (inj₂ p) rewrite ++⁻∘++⁺ xs (inj₂ p) = refl
-++ˡ = ++⁺ˡ
-++ʳ = ++⁺ʳ
-
-++↔ : ∀ {a p} {A : Set a} {P : A → Set p} {xs ys} →
- (Any P xs ⊎ Any P ys) ↔ Any P (xs ++ ys)
-++↔ {P = P} {xs = xs} = record
- { to = P.→-to-⟶ [ ++⁺ˡ {P = P}, ++⁺ʳ {P = P} xs ]′
- ; from = P.→-to-⟶ $ ++⁻ {P = P} xs
- ; inverse-of = record
- { left-inverse-of = ++⁻∘++⁺ xs
- ; right-inverse-of = ++⁺∘++⁻ xs
+ ++↔ : ∀ {xs ys} → (Any P xs ⊎ Any P ys) ↔ Any P (xs ++ ys)
+ ++↔ {xs = xs} = record
+ { to = P.→-to-⟶ [ ++⁺ˡ , ++⁺ʳ xs ]′
+ ; from = P.→-to-⟶ $ ++⁻ xs
+ ; inverse-of = record
+ { left-inverse-of = ++⁻∘++⁺ xs
+ ; right-inverse-of = ++⁺∘++⁻ xs
+ }
}
- }
-
--- return.
-
-private
-
- return⁺ : ∀ {a p} {A : Set a} {P : A → Set p} {x} →
- P x → Any P (return x)
- return⁺ = here
-
- return⁻ : ∀ {a p} {A : Set a} {P : A → Set p} {x} →
- Any P (return x) → P x
- return⁻ (here p) = p
- return⁻ (there ())
-
- return⁺∘return⁻ : ∀ {a p} {A : Set a} {P : A → Set p} {x}
- (p : Any P (return x)) →
- return⁺ (return⁻ p) ≡ p
- return⁺∘return⁻ (here p) = refl
- return⁺∘return⁻ (there ())
- return⁻∘return⁺ : ∀ {a p} {A : Set a} (P : A → Set p) {x} (p : P x) →
- return⁻ {P = P} (return⁺ p) ≡ p
- return⁻∘return⁺ P p = refl
+ ++-comm : ∀ xs ys → Any P (xs ++ ys) → Any P (ys ++ xs)
+ ++-comm xs ys = [ ++⁺ʳ ys , ++⁺ˡ ]′ ∘ ++⁻ xs
-return↔ : ∀ {a p} {A : Set a} {P : A → Set p} {x} →
- P x ↔ Any P (return x)
-return↔ {P = P} = record
- { to = P.→-to-⟶ $ return⁺ {P = P}
- ; from = P.→-to-⟶ $ return⁻ {P = P}
- ; inverse-of = record
- { left-inverse-of = return⁻∘return⁺ P
- ; right-inverse-of = return⁺∘return⁻
+ ++-comm∘++-comm : ∀ xs {ys} (p : Any P (xs ++ ys)) →
+ ++-comm ys xs (++-comm xs ys p) ≡ p
+ ++-comm∘++-comm [] {ys} p
+ rewrite ++⁻∘++⁺ ys {ys = []} (inj₁ p) = P.refl
+ ++-comm∘++-comm (x ∷ xs) {ys} (here p)
+ rewrite ++⁻∘++⁺ ys {ys = x ∷ xs} (inj₂ (here p)) = P.refl
+ ++-comm∘++-comm (x ∷ xs) (there p) with ++⁻ xs p | ++-comm∘++-comm xs p
+ ++-comm∘++-comm (x ∷ xs) {ys} (there .([ ++⁺ʳ xs , ++⁺ˡ ]′ (++⁻ ys (++⁺ʳ ys p))))
+ | inj₁ p | P.refl
+ rewrite ++⁻∘++⁺ ys (inj₂ p)
+ | ++⁻∘++⁺ ys (inj₂ $ there {x = x} p) = P.refl
+ ++-comm∘++-comm (x ∷ xs) {ys} (there .([ ++⁺ʳ xs , ++⁺ˡ ]′ (++⁻ ys (++⁺ˡ p))))
+ | inj₂ p | P.refl
+ rewrite ++⁻∘++⁺ ys {ys = xs} (inj₁ p)
+ | ++⁻∘++⁺ ys {ys = x ∷ xs} (inj₁ p) = P.refl
+
+ ++↔++ : ∀ xs ys → Any P (xs ++ ys) ↔ Any P (ys ++ xs)
+ ++↔++ xs ys = record
+ { to = P.→-to-⟶ $ ++-comm xs ys
+ ; from = P.→-to-⟶ $ ++-comm ys xs
+ ; inverse-of = record
+ { left-inverse-of = ++-comm∘++-comm xs
+ ; right-inverse-of = ++-comm∘++-comm ys
+ }
}
- }
-
--- _∷_.
-∷↔ : ∀ {a p} {A : Set a} (P : A → Set p) {x xs} →
- (P x ⊎ Any P xs) ↔ Any P (x ∷ xs)
-∷↔ P {x} {xs} =
- (P x ⊎ Any P xs) ↔⟨ return↔ {P = P} ⊎-cong (Any P xs ∎) ⟩
- (Any P [ x ] ⊎ Any P xs) ↔⟨ ++↔ {P = P} {xs = [ x ]} ⟩
- Any P (x ∷ xs) ∎
-
--- concat.
+------------------------------------------------------------------------
+-- concat
-private
+module _ {a p} {A : Set a} {P : A → Set p} where
- concat⁺ : ∀ {a p} {A : Set a} {P : A → Set p} {xss} →
- Any (Any P) xss → Any P (concat xss)
+ concat⁺ : ∀ {xss} → Any (Any P) xss → Any P (concat xss)
concat⁺ (here p) = ++⁺ˡ p
concat⁺ (there {x = xs} p) = ++⁺ʳ xs (concat⁺ p)
- concat⁻ : ∀ {a p} {A : Set a} {P : A → Set p} xss →
- Any P (concat xss) → Any (Any P) xss
+ concat⁻ : ∀ xss → Any P (concat xss) → Any (Any P) xss
concat⁻ [] ()
concat⁻ ([] ∷ xss) p = there $ concat⁻ xss p
concat⁻ ((x ∷ xs) ∷ xss) (here p) = here (here p)
- concat⁻ ((x ∷ xs) ∷ xss) (there p)
- with concat⁻ (xs ∷ xss) p
+ concat⁻ ((x ∷ xs) ∷ xss) (there p) with concat⁻ (xs ∷ xss) p
... | here p′ = here (there p′)
... | there p′ = there p′
- concat⁻∘++⁺ˡ : ∀ {a p} {A : Set a} {P : A → Set p} {xs} xss (p : Any P xs) →
+ concat⁻∘++⁺ˡ : ∀ {xs} xss (p : Any P xs) →
concat⁻ (xs ∷ xss) (++⁺ˡ p) ≡ here p
concat⁻∘++⁺ˡ xss (here p) = refl
concat⁻∘++⁺ˡ xss (there p) rewrite concat⁻∘++⁺ˡ xss p = refl
- concat⁻∘++⁺ʳ : ∀ {a p} {A : Set a} {P : A → Set p} xs xss (p : Any P (concat xss)) →
- concat⁻ (xs ∷ xss) (++⁺ʳ xs p) ≡ there (concat⁻ xss p)
+ concat⁻∘++⁺ʳ : ∀ xs xss (p : Any P (concat xss)) →
+ concat⁻ (xs ∷ xss) (++⁺ʳ xs p) ≡ there (concat⁻ xss p)
concat⁻∘++⁺ʳ [] xss p = refl
concat⁻∘++⁺ʳ (x ∷ xs) xss p rewrite concat⁻∘++⁺ʳ xs xss p = refl
- concat⁺∘concat⁻ : ∀ {a p} {A : Set a} {P : A → Set p} xss (p : Any P (concat xss)) →
- concat⁺ (concat⁻ xss p) ≡ p
+ concat⁺∘concat⁻ : ∀ xss (p : Any P (concat xss)) →
+ concat⁺ (concat⁻ xss p) ≡ p
concat⁺∘concat⁻ [] ()
concat⁺∘concat⁻ ([] ∷ xss) p = concat⁺∘concat⁻ xss p
concat⁺∘concat⁻ ((x ∷ xs) ∷ xss) (here p) = refl
concat⁺∘concat⁻ ((x ∷ xs) ∷ xss) (there p)
- with concat⁻ (xs ∷ xss) p | concat⁺∘concat⁻ (xs ∷ xss) p
+ with concat⁻ (xs ∷ xss) p | concat⁺∘concat⁻ (xs ∷ xss) p
concat⁺∘concat⁻ ((x ∷ xs) ∷ xss) (there .(++⁺ˡ p′)) | here p′ | refl = refl
concat⁺∘concat⁻ ((x ∷ xs) ∷ xss) (there .(++⁺ʳ xs (concat⁺ p′))) | there p′ | refl = refl
- concat⁻∘concat⁺ : ∀ {a p} {A : Set a} {P : A → Set p} {xss} (p : Any (Any P) xss) →
- concat⁻ xss (concat⁺ p) ≡ p
+ concat⁻∘concat⁺ : ∀ {xss} (p : Any (Any P) xss) → concat⁻ xss (concat⁺ p) ≡ p
concat⁻∘concat⁺ (here p) = concat⁻∘++⁺ˡ _ p
concat⁻∘concat⁺ (there {x = xs} {xs = xss} p)
rewrite concat⁻∘++⁺ʳ xs xss (concat⁺ p) =
P.cong there $ concat⁻∘concat⁺ p
-concat↔ : ∀ {a p} {A : Set a} {P : A → Set p} {xss} →
- Any (Any P) xss ↔ Any P (concat xss)
-concat↔ {P = P} {xss = xss} = record
- { to = P.→-to-⟶ $ concat⁺ {P = P}
- ; from = P.→-to-⟶ $ concat⁻ {P = P} xss
- ; inverse-of = record
- { left-inverse-of = concat⁻∘concat⁺
- ; right-inverse-of = concat⁺∘concat⁻ xss
+ concat↔ : ∀ {xss} → Any (Any P) xss ↔ Any P (concat xss)
+ concat↔ {xss = xss} = record
+ { to = P.→-to-⟶ $ concat⁺
+ ; from = P.→-to-⟶ $ concat⁻ xss
+ ; inverse-of = record
+ { left-inverse-of = concat⁻∘concat⁺
+ ; right-inverse-of = concat⁺∘concat⁻ xss
+ }
}
- }
+
+------------------------------------------------------------------------
+-- applyUpTo
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+ applyUpTo⁺ : ∀ f {i n} → P (f i) → i < n → Any P (applyUpTo f n)
+ applyUpTo⁺ _ p (s≤s z≤n) = here p
+ applyUpTo⁺ f p (s≤s (s≤s i<n)) =
+ there (applyUpTo⁺ (f ∘ suc) p (s≤s i<n))
+
+ applyUpTo⁻ : ∀ f {n} → Any P (applyUpTo f n) →
+ ∃ λ i → i < n × P (f i)
+ applyUpTo⁻ f {zero} ()
+ applyUpTo⁻ f {suc n} (here p) = zero , s≤s z≤n , p
+ applyUpTo⁻ f {suc n} (there p) with applyUpTo⁻ (f ∘ suc) p
+ ... | i , i<n , q = suc i , s≤s i<n , q
+
+------------------------------------------------------------------------
+-- tabulate
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+ tabulate⁺ : ∀ {n} {f : Fin n → A} i → P (f i) → Any P (tabulate f)
+ tabulate⁺ fzero p = here p
+ tabulate⁺ (fsuc i) p = there (tabulate⁺ i p)
+
+ tabulate⁻ : ∀ {n} {f : Fin n → A} →
+ Any P (tabulate f) → ∃ λ i → P (f i)
+ tabulate⁻ {zero} ()
+ tabulate⁻ {suc n} (here p) = fzero , p
+ tabulate⁻ {suc n} (there p) = Prod.map fsuc id (tabulate⁻ p)
+
+------------------------------------------------------------------------
+-- map-with-∈.
+
+module _ {a b p} {A : Set a} {B : Set b} {P : B → Set p} where
+
+ map-with-∈⁺ : ∀ {xs : List A} (f : ∀ {x} → x ∈ xs → B) →
+ (∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)) →
+ Any P (map-with-∈ xs f)
+ map-with-∈⁺ f (_ , here refl , p) = here p
+ map-with-∈⁺ f (_ , there x∈xs , p) =
+ there $ map-with-∈⁺ (f ∘ there) (_ , x∈xs , p)
+
+ map-with-∈⁻ : ∀ (xs : List A) (f : ∀ {x} → x ∈ xs → B) →
+ Any P (map-with-∈ xs f) →
+ ∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)
+ map-with-∈⁻ [] f ()
+ map-with-∈⁻ (y ∷ xs) f (here p) = (y , here refl , p)
+ map-with-∈⁻ (y ∷ xs) f (there p) =
+ Prod.map id (Prod.map there id) $ map-with-∈⁻ xs (f ∘ there) p
+
+ map-with-∈↔ : ∀ {xs : List A} {f : ∀ {x} → x ∈ xs → B} →
+ (∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)) ↔ Any P (map-with-∈ xs f)
+ map-with-∈↔ = record
+ { to = P.→-to-⟶ (map-with-∈⁺ _)
+ ; from = P.→-to-⟶ (map-with-∈⁻ _ _)
+ ; inverse-of = record
+ { left-inverse-of = from∘to _
+ ; right-inverse-of = to∘from _ _
+ }
+ }
+ where
+
+ from∘to : ∀ {xs : List A} (f : ∀ {x} → x ∈ xs → B)
+ (p : ∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)) →
+ map-with-∈⁻ xs f (map-with-∈⁺ f p) ≡ p
+ from∘to f (_ , here refl , p) = refl
+ from∘to f (_ , there x∈xs , p)
+ rewrite from∘to (f ∘ there) (_ , x∈xs , p) = refl
+
+ to∘from : ∀ (xs : List A) (f : ∀ {x} → x ∈ xs → B)
+ (p : Any P (map-with-∈ xs f)) →
+ map-with-∈⁺ f (map-with-∈⁻ xs f p) ≡ p
+ to∘from [] f ()
+ to∘from (y ∷ xs) f (here p) = refl
+ to∘from (y ∷ xs) f (there p) =
+ P.cong there $ to∘from xs (f ∘ there) p
+
+------------------------------------------------------------------------
+-- return
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+ return⁺ : ∀ {x} → P x → Any P (return x)
+ return⁺ = here
+
+ return⁻ : ∀ {x} → Any P (return x) → P x
+ return⁻ (here p) = p
+ return⁻ (there ())
+
+ return⁺∘return⁻ : ∀ {x} (p : Any P (return x)) →
+ return⁺ (return⁻ p) ≡ p
+ return⁺∘return⁻ (here p) = refl
+ return⁺∘return⁻ (there ())
+
+ return⁻∘return⁺ : ∀ {x} (p : P x) → return⁻ (return⁺ p) ≡ p
+ return⁻∘return⁺ p = refl
+
+ return↔ : ∀ {x} → P x ↔ Any P (return x)
+ return↔ = record
+ { to = P.→-to-⟶ $ return⁺
+ ; from = P.→-to-⟶ $ return⁻
+ ; inverse-of = record
+ { left-inverse-of = return⁻∘return⁺
+ ; right-inverse-of = return⁺∘return⁻
+ }
+ }
+
+------------------------------------------------------------------------
+-- _∷_.
+
+∷↔ : ∀ {a p} {A : Set a} (P : A → Set p) {x xs} →
+ (P x ⊎ Any P xs) ↔ Any P (x ∷ xs)
+∷↔ P {x} {xs} =
+ (P x ⊎ Any P xs) ↔⟨ return↔ {P = P} ⊎-cong (Any P xs ∎) ⟩
+ (Any P [ x ] ⊎ Any P xs) ↔⟨ ++↔ {P = P} {xs = [ x ]} ⟩
+ Any P (x ∷ xs) ∎
-- _>>=_.
@@ -516,112 +653,3 @@ concat↔ {P = P} {xss = xss} = record
(Any P xs × Any Q ys) ↔⟨ ×↔ ⟩
Any (λ x → Any (λ y → P x × Q y) ys) xs ↔⟨ ⊗↔ ⟩
Any (P ⟨×⟩ Q) (xs ⊗ ys) ∎
-
--- map-with-∈.
-
-map-with-∈↔ :
- ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p} {xs : List A}
- {f : ∀ {x} → x ∈ xs → B} →
- (∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)) ↔ Any P (map-with-∈ xs f)
-map-with-∈↔ {A = A} {B} {P} = record
- { to = P.→-to-⟶ (map-with-∈⁺ _)
- ; from = P.→-to-⟶ (map-with-∈⁻ _ _)
- ; inverse-of = record
- { left-inverse-of = from∘to _
- ; right-inverse-of = to∘from _ _
- }
- }
- where
- map-with-∈⁺ : ∀ {xs : List A}
- (f : ∀ {x} → x ∈ xs → B) →
- (∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)) →
- Any P (map-with-∈ xs f)
- map-with-∈⁺ f (_ , here refl , p) = here p
- map-with-∈⁺ f (_ , there x∈xs , p) =
- there $ map-with-∈⁺ (f ∘ there) (_ , x∈xs , p)
-
- map-with-∈⁻ : ∀ (xs : List A)
- (f : ∀ {x} → x ∈ xs → B) →
- Any P (map-with-∈ xs f) →
- ∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)
- map-with-∈⁻ [] f ()
- map-with-∈⁻ (y ∷ xs) f (here p) = (y , here refl , p)
- map-with-∈⁻ (y ∷ xs) f (there p) =
- Prod.map id (Prod.map there id) $ map-with-∈⁻ xs (f ∘ there) p
-
- from∘to : ∀ {xs : List A} (f : ∀ {x} → x ∈ xs → B)
- (p : ∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)) →
- map-with-∈⁻ xs f (map-with-∈⁺ f p) ≡ p
- from∘to f (_ , here refl , p) = refl
- from∘to f (_ , there x∈xs , p)
- rewrite from∘to (f ∘ there) (_ , x∈xs , p) = refl
-
- to∘from : ∀ (xs : List A) (f : ∀ {x} → x ∈ xs → B)
- (p : Any P (map-with-∈ xs f)) →
- map-with-∈⁺ f (map-with-∈⁻ xs f p) ≡ p
- to∘from [] f ()
- to∘from (y ∷ xs) f (here p) = refl
- to∘from (y ∷ xs) f (there p) =
- P.cong there $ to∘from xs (f ∘ there) p
-
-------------------------------------------------------------------------
--- Any and any are related via T
-
--- These introduction and elimination rules are not inverses, though.
-
-private
-
- any⁺ : ∀ {a} {A : Set a} (p : A → Bool) {xs} →
- Any (T ∘ p) xs → T (any p xs)
- any⁺ p (here px) = Equivalence.from T-∨ ⟨$⟩ inj₁ px
- any⁺ p (there {x = x} pxs) with p x
- ... | true = _
- ... | false = any⁺ p pxs
-
- any⁻ : ∀ {a} {A : Set a} (p : A → Bool) xs →
- T (any p xs) → Any (T ∘ p) xs
- any⁻ p [] ()
- any⁻ p (x ∷ xs) px∷xs with p x | inspect p x
- any⁻ p (x ∷ xs) px∷xs | true | P[ eq ] = here (Equivalence.from T-≡ ⟨$⟩ eq)
- any⁻ p (x ∷ xs) px∷xs | false | _ = there (any⁻ p xs px∷xs)
-
-any⇔ : ∀ {a} {A : Set a} {p : A → Bool} {xs} →
- Any (T ∘ p) xs ⇔ T (any p xs)
-any⇔ = Eq.equivalence (any⁺ _) (any⁻ _ _)
-
-------------------------------------------------------------------------
--- _++_ is commutative
-
-private
-
- ++-comm : ∀ {a p} {A : Set a} {P : A → Set p} xs ys →
- Any P (xs ++ ys) → Any P (ys ++ xs)
- ++-comm xs ys = [ ++⁺ʳ ys , ++⁺ˡ ]′ ∘ ++⁻ xs
-
- ++-comm∘++-comm : ∀ {a p} {A : Set a} {P : A → Set p}
- xs {ys} (p : Any P (xs ++ ys)) →
- ++-comm ys xs (++-comm xs ys p) ≡ p
- ++-comm∘++-comm [] {ys} p
- rewrite ++⁻∘++⁺ ys {ys = []} (inj₁ p) = P.refl
- ++-comm∘++-comm {P = P} (x ∷ xs) {ys} (here p)
- rewrite ++⁻∘++⁺ {P = P} ys {ys = x ∷ xs} (inj₂ (here p)) = P.refl
- ++-comm∘++-comm (x ∷ xs) (there p) with ++⁻ xs p | ++-comm∘++-comm xs p
- ++-comm∘++-comm (x ∷ xs) {ys} (there .([ ++⁺ʳ xs , ++⁺ˡ ]′ (++⁻ ys (++⁺ʳ ys p))))
- | inj₁ p | P.refl
- rewrite ++⁻∘++⁺ ys (inj₂ p)
- | ++⁻∘++⁺ ys (inj₂ $ there {x = x} p) = P.refl
- ++-comm∘++-comm (x ∷ xs) {ys} (there .([ ++⁺ʳ xs , ++⁺ˡ ]′ (++⁻ ys (++⁺ˡ p))))
- | inj₂ p | P.refl
- rewrite ++⁻∘++⁺ ys {ys = xs} (inj₁ p)
- | ++⁻∘++⁺ ys {ys = x ∷ xs} (inj₁ p) = P.refl
-
-++↔++ : ∀ {a p} {A : Set a} {P : A → Set p} xs ys →
- Any P (xs ++ ys) ↔ Any P (ys ++ xs)
-++↔++ {P = P} xs ys = record
- { to = P.→-to-⟶ $ ++-comm {P = P} xs ys
- ; from = P.→-to-⟶ $ ++-comm {P = P} ys xs
- ; inverse-of = record
- { left-inverse-of = ++-comm∘++-comm xs
- ; right-inverse-of = ++-comm∘++-comm ys
- }
- }
diff --git a/src/Data/List/Base.agda b/src/Data/List/Base.agda
index dd082c5..168a93a 100644
--- a/src/Data/List/Base.agda
+++ b/src/Data/List/Base.agda
@@ -7,11 +7,12 @@
module Data.List.Base where
open import Data.Nat.Base using (ℕ; zero; suc; _+_; _*_)
+open import Data.Fin using (Fin) renaming (zero to fzero; suc to fsuc)
open import Data.Sum as Sum using (_⊎_; inj₁; inj₂)
open import Data.Bool.Base using (Bool; false; true; not; _∧_; _∨_; if_then_else_)
open import Data.Maybe.Base using (Maybe; nothing; just)
open import Data.Product as Prod using (_×_; _,_)
-open import Function
+open import Function using (id; _∘_)
------------------------------------------------------------------------
-- Types
@@ -138,16 +139,36 @@ unfold B f {n = suc n} s with f s
... | nothing = []
... | just (x , s') = x ∷ unfold B f s'
+-- applyUpTo 3 = f0 ∷ f1 ∷ f2 ∷ [].
+
+applyUpTo : ∀ {a} {A : Set a} → (ℕ → A) → ℕ → List A
+applyUpTo f zero = []
+applyUpTo f (suc n) = f zero ∷ applyUpTo (f ∘ suc) n
+
+-- upTo 3 = 0 ∷ 1 ∷ 2 ∷ [].
+
+upTo : ℕ → List ℕ
+upTo = applyUpTo id
+
+-- applyDownFrom 3 = f2 ∷ f1 ∷ f0 ∷ [].
+
+applyDownFrom : ∀ {a} {A : Set a} → (ℕ → A) → ℕ → List A
+applyDownFrom f zero = []
+applyDownFrom f (suc n) = f n ∷ applyDownFrom f n
+
-- downFrom 3 = 2 ∷ 1 ∷ 0 ∷ [].
downFrom : ℕ → List ℕ
-downFrom n = unfold Singleton f (wrap n)
- where
- data Singleton : ℕ → Set where
- wrap : (n : ℕ) → Singleton n
+downFrom = applyDownFrom id
+
+-- tabulate f = f 0 ∷ f 1 ∷ ... ∷ f n ∷ []
+
+tabulate : ∀ {a n} {A : Set a} (f : Fin n → A) → List A
+tabulate {_} {zero} f = []
+tabulate {_} {suc n} f = f fzero ∷ tabulate (f ∘ fsuc)
- f : ∀ {n} → Singleton (suc n) → Maybe (ℕ × Singleton n)
- f {n} (wrap .(suc n)) = just (n , wrap n)
+allFin : ∀ n → List (Fin n)
+allFin n = tabulate id
-- ** Conversions
diff --git a/src/Data/List/Countdown.agda b/src/Data/List/Countdown.agda
index 576cd97..4818b07 100644
--- a/src/Data/List/Countdown.agda
+++ b/src/Data/List/Countdown.agda
@@ -11,7 +11,9 @@ open import Relation.Binary
module Data.List.Countdown (D : DecSetoid Level.zero Level.zero) where
open import Data.Empty
-open import Data.Fin using (Fin; zero; suc)
+open import Data.Fin using (Fin; zero; suc; punchOut)
+open import Data.Fin.Properties
+ using (suc-injective; punchOut-injective)
open import Function
open import Function.Equality using (_⟨$⟩_)
open import Function.Injection
@@ -29,17 +31,13 @@ open PropEq.≡-Reasoning
private
open module D = DecSetoid D
hiding (refl) renaming (Carrier to Elem)
- open Any.Membership D.setoid
+ open import Data.List.Any.Membership D.setoid
------------------------------------------------------------------------
-- Helper functions
private
- drop-suc : ∀ {n} {i j : Fin n} →
- _≡_ {A = Fin (suc n)} (suc i) (suc j) → i ≡ j
- drop-suc refl = refl
-
drop-inj₂ : ∀ {A B : Set} {x y} →
inj₂ {A = A} {B = B} x ≡ inj₂ y → x ≡ y
drop-inj₂ refl = refl
@@ -99,34 +97,7 @@ private
helper (there {x = x} x₁∈xs) (there x₂∈xs) () | yes x₁≈x | no x₂≉x
helper (there {x = x} x₁∈xs) (there x₂∈xs) () | no x₁≉x | yes x₂≈x
helper (there {x = x} x₁∈xs) (there x₂∈xs) eq | no x₁≉x | no x₂≉x =
- helper x₁∈xs x₂∈xs (drop-suc eq)
-
- -- If there are at least two elements in Fin (suc n), then Fin n is
- -- inhabited. This is a variant of the thick function from Conor
- -- McBride's "First-order unification by structural recursion".
-
- thick : ∀ {n} (i j : Fin (suc n)) → i ≢ j → Fin n
- thick zero zero i≢j = ⊥-elim (i≢j refl)
- thick zero (suc j) _ = j
- thick {zero} (suc ()) _ _
- thick {suc n} (suc i) zero _ = zero
- thick {suc n} (suc i) (suc j) i≢j =
- suc (thick i j (i≢j ∘ cong suc))
-
- -- thick i is injective in one of its arguments.
-
- thick-injective : ∀ {n} (i j k : Fin (suc n))
- {i≢j : i ≢ j} {i≢k : i ≢ k} →
- thick i j i≢j ≡ thick i k i≢k → j ≡ k
- thick-injective zero zero _ {i≢j = i≢j} _ = ⊥-elim (i≢j refl)
- thick-injective zero _ zero {i≢k = i≢k} _ = ⊥-elim (i≢k refl)
- thick-injective zero (suc j) (suc k) j≡k = cong suc j≡k
- thick-injective {zero} (suc ()) _ _ _
- thick-injective {suc n} (suc i) zero zero _ = refl
- thick-injective {suc n} (suc i) zero (suc k) ()
- thick-injective {suc n} (suc i) (suc j) zero ()
- thick-injective {suc n} (suc i) (suc j) (suc k) {i≢j} {i≢k} eq =
- cong suc $ thick-injective i j k {i≢j ∘ cong suc} {i≢k ∘ cong suc} (drop-suc eq)
+ helper x₁∈xs x₂∈xs (suc-injective eq)
------------------------------------------------------------------------
-- The countdown data structure
@@ -215,7 +186,7 @@ insert {counted} {n} counted⊕1+n x x∉counted =
kind′ y | _ | inj₁ x∈counted | _ | _ = ⊥-elim (x∉counted x∈counted)
kind′ y | _ | _ | inj₁ y∈counted | _ = inj₁ (there y∈counted)
kind′ y | no y≉x | inj₂ i | inj₂ j | hlp =
- inj₂ (thick i j (y≉x ∘ sym ∘ hlp _ refl refl))
+ inj₂ (punchOut (y≉x ∘ sym ∘ hlp _ refl refl))
inj : ∀ {y z i} → kind′ y ≡ inj₂ i → kind′ z ≡ inj₂ i → y ≈ z
inj {y} {z} eq₁ eq₂ with y ≟ x | z ≟ x | kind x | kind y | kind z
@@ -227,8 +198,8 @@ insert {counted} {n} counted⊕1+n x x∉counted =
inj _ () | no _ | no _ | inj₂ _ | _ | inj₁ _ | _ | _ | _
inj eq₁ eq₂ | no _ | no _ | inj₂ i | inj₂ _ | inj₂ _ | _ | _ | hlp =
hlp _ refl refl $
- thick-injective i _ _ $
- PropEq.trans (drop-inj₂ eq₁) (PropEq.sym (drop-inj₂ eq₂))
+ punchOut-injective {i = i} _ _ $
+ (PropEq.trans (drop-inj₂ eq₁) (PropEq.sym (drop-inj₂ eq₂)))
-- Counts an element if it has not already been counted.
diff --git a/src/Data/List/Properties.agda b/src/Data/List/Properties.agda
index 1d8a79d..9fa36a7 100644
--- a/src/Data/List/Properties.agda
+++ b/src/Data/List/Properties.agda
@@ -20,7 +20,7 @@ open import Data.Nat
open import Data.Nat.Properties
open import Data.Product as Prod hiding (map)
open import Function
-import Algebra.FunctionProperties
+open import Algebra.FunctionProperties
import Relation.Binary.EqReasoning as EqR
open import Relation.Binary.PropositionalEquality as P
using (_≡_; _≢_; _≗_; refl)
@@ -29,14 +29,15 @@ open import Relation.Nullary.Decidable using (⌊_⌋)
open import Relation.Unary using (Decidable)
private
- open module FP {a} {A : Set a} =
- Algebra.FunctionProperties (_≡_ {A = A})
open module LMP {ℓ} = RawMonadPlus (List.monadPlus {ℓ = ℓ})
module LM {a} {A : Set a} = Monoid (List.monoid A)
module List-solver {a} {A : Set a} =
Algebra.Monoid-solver (monoid A) renaming (id to nil)
+------------------------------------------------------------------------
+-- Equality
+
∷-injective : ∀ {a} {A : Set a} {x y : A} {xs ys} →
x ∷ xs ≡ y List.∷ ys → x ≡ y × xs ≡ ys
∷-injective refl = (refl , refl)
@@ -45,14 +46,15 @@ module List-solver {a} {A : Set a} =
xs ∷ʳ x ≡ ys ∷ʳ y → xs ≡ ys × x ≡ y
∷ʳ-injective [] [] refl = (refl , refl)
∷ʳ-injective (x ∷ xs) (y ∷ ys) eq with ∷-injective eq
-∷ʳ-injective (x ∷ xs) (.x ∷ ys) eq | (refl , eq′) =
- Prod.map (P.cong (_∷_ x)) id $ ∷ʳ-injective xs ys eq′
-
+... | refl , eq′ = Prod.map (P.cong (x ∷_)) id (∷ʳ-injective xs ys eq′)
∷ʳ-injective [] (_ ∷ []) ()
∷ʳ-injective [] (_ ∷ _ ∷ _) ()
∷ʳ-injective (_ ∷ []) [] ()
∷ʳ-injective (_ ∷ _ ∷ _) [] ()
+------------------------------------------------------------------------
+-- _++_
+
right-identity-unique : ∀ {a} {A : Set a} (xs : List A) {ys} →
xs ≡ xs ++ ys → ys ≡ []
right-identity-unique [] refl = refl
@@ -72,28 +74,58 @@ left-identity-unique {xs = x ∷ xs} (y ∷ ys) eq
left-identity-unique {xs = x ∷ xs} (y ∷ [] ) eq | ()
left-identity-unique {xs = x ∷ xs} (y ∷ _ ∷ _) eq | ()
--- Map, sum, and append.
+length-++ : ∀ {a} {A : Set a} (xs : List A) {ys} →
+ length (xs ++ ys) ≡ length xs + length ys
+length-++ [] = refl
+length-++ (x ∷ xs) = P.cong suc (length-++ xs)
+
+------------------------------------------------------------------------
+-- map
+
+map-id : ∀ {a} {A : Set a} → map id ≗ id {A = List A}
+map-id [] = refl
+map-id (x ∷ xs) = P.cong (x ∷_) (map-id xs)
+
+map-id₂ : ∀ {a} {A : Set a} {f : A → A} {xs} →
+ All (λ x → f x ≡ x) xs → map f xs ≡ xs
+map-id₂ [] = refl
+map-id₂ (fx≡x ∷ pxs) = P.cong₂ _∷_ fx≡x (map-id₂ pxs)
+
+map-compose : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
+ {g : B → C} {f : A → B} → map (g ∘ f) ≗ map g ∘ map f
+map-compose [] = refl
+map-compose (x ∷ xs) = P.cong (_ ∷_) (map-compose xs)
map-++-commute : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) xs ys →
map f (xs ++ ys) ≡ map f xs ++ map f ys
map-++-commute f [] ys = refl
-map-++-commute f (x ∷ xs) ys =
- P.cong (_∷_ (f x)) (map-++-commute f xs ys)
+map-++-commute f (x ∷ xs) ys = P.cong (f x ∷_) (map-++-commute f xs ys)
-sum-++-commute : ∀ xs ys → sum (xs ++ ys) ≡ sum xs + sum ys
-sum-++-commute [] ys = refl
-sum-++-commute (x ∷ xs) ys = begin
- x + sum (xs ++ ys)
- ≡⟨ P.cong (_+_ x) (sum-++-commute xs ys) ⟩
- x + (sum xs + sum ys)
- ≡⟨ P.sym $ +-assoc x _ _ ⟩
- (x + sum xs) + sum ys
- ∎
- where
- open CommutativeSemiring commutativeSemiring hiding (_+_)
- open P.≡-Reasoning
+map-cong : ∀ {a b} {A : Set a} {B : Set b} {f g : A → B} →
+ f ≗ g → map f ≗ map g
+map-cong f≗g [] = refl
+map-cong f≗g (x ∷ xs) = P.cong₂ _∷_ (f≗g x) (map-cong f≗g xs)
+
+map-cong₂ : ∀ {a b} {A : Set a} {B : Set b} {f g : A → B} →
+ ∀ {xs} → All (λ x → f x ≡ g x) xs → map f xs ≡ map g xs
+map-cong₂ [] = refl
+map-cong₂ (fx≡gx ∷ fxs≡gxs) = P.cong₂ _∷_ fx≡gx (map-cong₂ fxs≡gxs)
+
+length-map : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) xs →
+ length (map f xs) ≡ length xs
+length-map f [] = refl
+length-map f (x ∷ xs) = P.cong suc (length-map f xs)
+
+------------------------------------------------------------------------
+-- replicate
+
+length-replicate : ∀ {a} {A : Set a} n {x : A} →
+ length (replicate n x) ≡ n
+length-replicate zero = refl
+length-replicate (suc n) = P.cong suc (length-replicate n)
--- Various properties about folds.
+------------------------------------------------------------------------
+-- foldr
foldr-universal : ∀ {a b} {A : Set a} {B : Set b}
(h : List A → B) f e →
@@ -102,14 +134,21 @@ foldr-universal : ∀ {a b} {A : Set a} {B : Set b}
h ≗ foldr f e
foldr-universal h f e base step [] = base
foldr-universal h f e base step (x ∷ xs) = begin
- h (x ∷ xs)
- ≡⟨ step x xs ⟩
- f x (h xs)
- ≡⟨ P.cong (f x) (foldr-universal h f e base step xs) ⟩
- f x (foldr f e xs)
- ∎
+ h (x ∷ xs)
+ ≡⟨ step x xs ⟩
+ f x (h xs)
+ ≡⟨ P.cong (f x) (foldr-universal h f e base step xs) ⟩
+ f x (foldr f e xs)
+ ∎
where open P.≡-Reasoning
+foldr-cong : ∀ {a b} {A : Set a} {B : Set b}
+ {f g : A → B → B} {d e : B} →
+ (∀ x y → f x y ≡ g x y) → d ≡ e →
+ foldr f d ≗ foldr g e
+foldr-cong f≗g refl [] = refl
+foldr-cong f≗g d≡e (x ∷ xs) rewrite foldr-cong f≗g d≡e xs = f≗g x _
+
foldr-fusion : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
(h : B → C) {f : A → B → B} {g : A → C → C} (e : B) →
(∀ x y → h (f x y) ≡ g x (h y)) →
@@ -126,15 +165,20 @@ idIsFold = foldr-universal id _∷_ [] refl (λ _ _ → refl)
++IsFold xs ys =
begin
xs ++ ys
- ≡⟨ P.cong (λ xs → xs ++ ys) (idIsFold xs) ⟩
+ ≡⟨ P.cong (_++ ys) (idIsFold xs) ⟩
foldr _∷_ [] xs ++ ys
- ≡⟨ foldr-fusion (λ xs → xs ++ ys) [] (λ _ _ → refl) xs ⟩
+ ≡⟨ foldr-fusion (_++ ys) [] (λ _ _ → refl) xs ⟩
foldr _∷_ ([] ++ ys) xs
≡⟨ refl ⟩
foldr _∷_ ys xs
where open P.≡-Reasoning
+foldr-++ : ∀ {a b} {A : Set a} {B : Set b} (f : A → B → B) x ys zs →
+ foldr f x (ys ++ zs) ≡ foldr f (foldr f x zs) ys
+foldr-++ f x [] zs = refl
+foldr-++ f x (y ∷ ys) zs = P.cong (f y) (foldr-++ f x ys zs)
+
mapIsFold : ∀ {a b} {A : Set a} {B : Set b} {f : A → B} →
map f ≗ foldr (λ x ys → f x ∷ ys) []
mapIsFold {f = f} =
@@ -147,6 +191,27 @@ mapIsFold {f = f} =
where open EqR (P._→-setoid_ _ _)
+foldr-∷ʳ : ∀ {a b} {A : Set a} {B : Set b} (f : A → B → B) x y ys →
+ foldr f x (ys ∷ʳ y) ≡ foldr f (f y x) ys
+foldr-∷ʳ f x y [] = refl
+foldr-∷ʳ f x y (z ∷ ys) = P.cong (f z) (foldr-∷ʳ f x y ys)
+
+------------------------------------------------------------------------
+-- foldl
+
+foldl-++ : ∀ {a b} {A : Set a} {B : Set b} (f : A → B → A) x ys zs →
+ foldl f x (ys ++ zs) ≡ foldl f (foldl f x ys) zs
+foldl-++ f x [] zs = refl
+foldl-++ f x (y ∷ ys) zs = foldl-++ f (f x y) ys zs
+
+foldl-∷ʳ : ∀ {a b} {A : Set a} {B : Set b} (f : A → B → A) x y ys →
+ foldl f x (ys ∷ʳ y) ≡ f (foldl f x ys) y
+foldl-∷ʳ f x y [] = refl
+foldl-∷ʳ f x y (z ∷ ys) = foldl-∷ʳ f (f x z) y ys
+
+------------------------------------------------------------------------
+-- concat
+
concat-map : ∀ {a b} {A : Set a} {B : Set b} {f : A → B} →
concat ∘ map (map f) ≗ map f ∘ concat
concat-map {b = b} {f = f} =
@@ -156,133 +221,100 @@ concat-map {b = b} {f = f} =
concat ∘ foldr (λ xs ys → map f xs ∷ ys) []
≈⟨ foldr-fusion {b = b} concat [] (λ _ _ → refl) ⟩
foldr (λ ys zs → map f ys ++ zs) []
- ≈⟨ P.sym ∘
- foldr-fusion (map f) [] (λ ys zs → map-++-commute f ys zs) ⟩
+ ≈⟨ P.sym ∘ foldr-fusion (map f) [] (map-++-commute f) ⟩
map f ∘ concat
where open EqR (P._→-setoid_ _ _)
-map-id : ∀ {a} {A : Set a} → map id ≗ id {A = List A}
-map-id {A = A} = begin
- map id ≈⟨ mapIsFold ⟩
- foldr _∷_ [] ≈⟨ P.sym ∘ idIsFold {A = A} ⟩
- id ∎
- where open EqR (P._→-setoid_ _ _)
-
-map-compose : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
- {g : B → C} {f : A → B} →
- map (g ∘ f) ≗ map g ∘ map f
-map-compose {A = A} {B} {g = g} {f} =
- begin
- map (g ∘ f)
- ≈⟨ P.cong (map (g ∘ f)) ∘ idIsFold ⟩
- map (g ∘ f) ∘ foldr _∷_ []
- ≈⟨ foldr-fusion (map (g ∘ f)) [] (λ _ _ → refl) ⟩
- foldr (λ a y → g (f a) ∷ y) []
- ≈⟨ P.sym ∘ foldr-fusion (map g) [] (λ _ _ → refl) ⟩
- map g ∘ foldr (λ a y → f a ∷ y) []
- ≈⟨ P.cong (map g) ∘ P.sym ∘ mapIsFold {A = A} {B = B} ⟩
- map g ∘ map f
- ∎
- where open EqR (P._→-setoid_ _ _)
-
-foldr-cong : ∀ {a b} {A : Set a} {B : Set b}
- {f₁ f₂ : A → B → B} {e₁ e₂ : B} →
- (∀ x y → f₁ x y ≡ f₂ x y) → e₁ ≡ e₂ →
- foldr f₁ e₁ ≗ foldr f₂ e₂
-foldr-cong {f₁ = f₁} {f₂} {e} f₁≗₂f₂ refl =
- begin
- foldr f₁ e
- ≈⟨ P.cong (foldr f₁ e) ∘ idIsFold ⟩
- foldr f₁ e ∘ foldr _∷_ []
- ≈⟨ foldr-fusion (foldr f₁ e) [] (λ x xs → f₁≗₂f₂ x (foldr f₁ e xs)) ⟩
- foldr f₂ e
- ∎
- where open EqR (P._→-setoid_ _ _)
+------------------------------------------------------------------------
+-- sum
-map-cong : ∀ {a b} {A : Set a} {B : Set b} {f g : A → B} →
- f ≗ g → map f ≗ map g
-map-cong {A = A} {B} {f} {g} f≗g =
- begin
- map f
- ≈⟨ mapIsFold ⟩
- foldr (λ x ys → f x ∷ ys) []
- ≈⟨ foldr-cong (λ x ys → P.cong₂ _∷_ (f≗g x) refl) refl ⟩
- foldr (λ x ys → g x ∷ ys) []
- ≈⟨ P.sym ∘ mapIsFold {A = A} {B = B} ⟩
- map g
- ∎
- where open EqR (P._→-setoid_ _ _)
+sum-++-commute : ∀ xs ys → sum (xs ++ ys) ≡ sum xs + sum ys
+sum-++-commute [] ys = refl
+sum-++-commute (x ∷ xs) ys = begin
+ x + sum (xs ++ ys) ≡⟨ P.cong (x +_) (sum-++-commute xs ys) ⟩
+ x + (sum xs + sum ys) ≡⟨ P.sym $ +-assoc x _ _ ⟩
+ (x + sum xs) + sum ys ∎
+ where open P.≡-Reasoning
--- Take, drop, and splitAt.
+------------------------------------------------------------------------
+-- take, drop, splitAt, takeWhile, dropWhile, and span
take++drop : ∀ {a} {A : Set a}
n (xs : List A) → take n xs ++ drop n xs ≡ xs
take++drop zero xs = refl
take++drop (suc n) [] = refl
-take++drop (suc n) (x ∷ xs) =
- P.cong (λ xs → x ∷ xs) (take++drop n xs)
+take++drop (suc n) (x ∷ xs) = P.cong (x ∷_) (take++drop n xs)
splitAt-defn : ∀ {a} {A : Set a} n →
splitAt {A = A} n ≗ < take n , drop n >
splitAt-defn zero xs = refl
splitAt-defn (suc n) [] = refl
splitAt-defn (suc n) (x ∷ xs) with splitAt n xs | splitAt-defn n xs
-... | (ys , zs) | ih = P.cong (Prod.map (_∷_ x) id) ih
-
--- TakeWhile, dropWhile, and span.
+... | (ys , zs) | ih = P.cong (Prod.map (x ∷_) id) ih
takeWhile++dropWhile : ∀ {a} {A : Set a} (p : A → Bool) (xs : List A) →
takeWhile p xs ++ dropWhile p xs ≡ xs
takeWhile++dropWhile p [] = refl
takeWhile++dropWhile p (x ∷ xs) with p x
-... | true = P.cong (_∷_ x) (takeWhile++dropWhile p xs)
+... | true = P.cong (x ∷_) (takeWhile++dropWhile p xs)
... | false = refl
span-defn : ∀ {a} {A : Set a} (p : A → Bool) →
span p ≗ < takeWhile p , dropWhile p >
span-defn p [] = refl
span-defn p (x ∷ xs) with p x
-... | true = P.cong (Prod.map (_∷_ x) id) (span-defn p xs)
+... | true = P.cong (Prod.map (x ∷_) id) (span-defn p xs)
... | false = refl
--- Partition, filter.
+------------------------------------------------------------------------
+-- Filtering
partition-defn : ∀ {a} {A : Set a} (p : A → Bool) →
partition p ≗ < filter p , filter (not ∘ p) >
partition-defn p [] = refl
-partition-defn p (x ∷ xs)
- with p x | partition p xs | partition-defn p xs
-... | true | (ys , zs) | eq = P.cong (Prod.map (_∷_ x) id) eq
-... | false | (ys , zs) | eq = P.cong (Prod.map id (_∷_ x)) eq
-
-filter-filters : ∀ {a p} {A : Set a} →
- (P : A → Set p) (dec : Decidable P) (xs : List A) →
- All P (filter (⌊_⌋ ∘ dec) xs)
-filter-filters P dec [] = []
-filter-filters P dec (x ∷ xs) with dec x
-filter-filters P dec (x ∷ xs) | yes px = px ∷ filter-filters P dec xs
-filter-filters P dec (x ∷ xs) | no ¬px = filter-filters P dec xs
-
--- Gfilter.
+partition-defn p (x ∷ xs) with p x
+... | true = P.cong (Prod.map (x ∷_) id) (partition-defn p xs)
+... | false = P.cong (Prod.map id (x ∷_)) (partition-defn p xs)
gfilter-just : ∀ {a} {A : Set a} (xs : List A) → gfilter just xs ≡ xs
gfilter-just [] = refl
-gfilter-just (x ∷ xs) = P.cong (_∷_ x) (gfilter-just xs)
+gfilter-just (x ∷ xs) = P.cong (x ∷_) (gfilter-just xs)
gfilter-nothing : ∀ {a} {A : Set a} (xs : List A) →
- gfilter {B = A} (λ _ → nothing) xs ≡ []
+ gfilter {B = A} (λ _ → nothing) xs ≡ []
gfilter-nothing [] = refl
gfilter-nothing (x ∷ xs) = gfilter-nothing xs
gfilter-concatMap : ∀ {a b} {A : Set a} {B : Set b} (f : A → Maybe B) →
- gfilter f ≗ concatMap (fromMaybe ∘ f)
+ gfilter f ≗ concatMap (fromMaybe ∘ f)
gfilter-concatMap f [] = refl
gfilter-concatMap f (x ∷ xs) with f x
-gfilter-concatMap f (x ∷ xs) | just y = P.cong (_∷_ y) (gfilter-concatMap f xs)
-gfilter-concatMap f (x ∷ xs) | nothing = gfilter-concatMap f xs
+... | just y = P.cong (y ∷_) (gfilter-concatMap f xs)
+... | nothing = gfilter-concatMap f xs
--- Inits, tails, and scanr.
+length-gfilter : ∀ {a b} {A : Set a} {B : Set b} (p : A → Maybe B) xs →
+ length (gfilter p xs) ≤ length xs
+length-gfilter p [] = z≤n
+length-gfilter p (x ∷ xs) with p x
+... | just y = s≤s (length-gfilter p xs)
+... | nothing = ≤-step (length-gfilter p xs)
+
+filter-filters : ∀ {a p} {A : Set a} →
+ (P : A → Set p) (dec : Decidable P) (xs : List A) →
+ All P (filter (⌊_⌋ ∘ dec) xs)
+filter-filters P dec [] = []
+filter-filters P dec (x ∷ xs) with dec x
+... | yes px = px ∷ filter-filters P dec xs
+... | no ¬px = filter-filters P dec xs
+
+length-filter : ∀ {a} {A : Set a} (p : A → Bool) xs →
+ length (filter p xs) ≤ length xs
+length-filter p xs =
+ length-gfilter (λ x → if p x then just x else nothing) xs
+
+------------------------------------------------------------------------
+-- Inits, tails, and scanr
scanr-defn : ∀ {a b} {A : Set a} {B : Set b}
(f : A → B → B) (e : B) →
@@ -299,77 +331,41 @@ scanl-defn : ∀ {a b} {A : Set a} {B : Set b}
(f : A → B → A) (e : A) →
scanl f e ≗ map (foldl f e) ∘ inits
scanl-defn f e [] = refl
-scanl-defn f e (x ∷ xs) = P.cong (_∷_ e) (begin
+scanl-defn f e (x ∷ xs) = P.cong (e ∷_) (begin
scanl f (f e x) xs
≡⟨ scanl-defn f (f e x) xs ⟩
map (foldl f (f e x)) (inits xs)
≡⟨ refl ⟩
- map (foldl f e ∘ (_∷_ x)) (inits xs)
+ map (foldl f e ∘ (x ∷_)) (inits xs)
≡⟨ map-compose (inits xs) ⟩
- map (foldl f e) (map (_∷_ x) (inits xs))
+ map (foldl f e) (map (x ∷_) (inits xs))
∎)
where open P.≡-Reasoning
--- Length.
-
-length-map : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) xs →
- length (map f xs) ≡ length xs
-length-map f [] = refl
-length-map f (x ∷ xs) = P.cong suc (length-map f xs)
-
-length-++ : ∀ {a} {A : Set a} (xs : List A) {ys} →
- length (xs ++ ys) ≡ length xs + length ys
-length-++ [] = refl
-length-++ (x ∷ xs) = P.cong suc (length-++ xs)
-
-length-replicate : ∀ {a} {A : Set a} n {x : A} →
- length (replicate n x) ≡ n
-length-replicate zero = refl
-length-replicate (suc n) = P.cong suc (length-replicate n)
-
-length-gfilter : ∀ {a b} {A : Set a} {B : Set b} (p : A → Maybe B) xs →
- length (gfilter p xs) ≤ length xs
-length-gfilter p [] = z≤n
-length-gfilter p (x ∷ xs) with p x
-length-gfilter p (x ∷ xs) | just y = s≤s (length-gfilter p xs)
-length-gfilter p (x ∷ xs) | nothing = ≤-step (length-gfilter p xs)
-
-length-filter : ∀ {a} {A : Set a} (p : A → Bool) xs →
- length (filter p xs) ≤ length xs
-length-filter p xs = length-gfilter (λ x → if p x then just x else nothing) xs
-
--- Reverse.
+------------------------------------------------------------------------
+-- reverse
unfold-reverse : ∀ {a} {A : Set a} (x : A) (xs : List A) →
reverse (x ∷ xs) ≡ reverse xs ∷ʳ x
-unfold-reverse x xs = begin
- foldl (flip _∷_) [ x ] xs ≡⟨ helper [ x ] xs ⟩
- reverse xs ∷ʳ x ∎
+unfold-reverse {A = A} x xs = helper [ x ] xs
where
open P.≡-Reasoning
-
- helper : ∀ {a} {A : Set a} (xs ys : List A) →
- foldl (flip _∷_) xs ys ≡ reverse ys ++ xs
- helper xs [] = P.refl
+ helper : (xs ys : List A) → foldl (flip _∷_) xs ys ≡ reverse ys ++ xs
+ helper xs [] = refl
helper xs (y ∷ ys) = begin
foldl (flip _∷_) (y ∷ xs) ys ≡⟨ helper (y ∷ xs) ys ⟩
reverse ys ++ y ∷ xs ≡⟨ P.sym $ LM.assoc (reverse ys) _ _ ⟩
- (reverse ys ∷ʳ y) ++ xs ≡⟨ P.sym $ P.cong (λ zs → zs ++ xs) (unfold-reverse y ys) ⟩
+ (reverse ys ∷ʳ y) ++ xs ≡⟨ P.sym $ P.cong (_++ xs) (unfold-reverse y ys) ⟩
reverse (y ∷ ys) ++ xs ∎
-reverse-++-commute :
- ∀ {a} {A : Set a} (xs ys : List A) →
- reverse (xs ++ ys) ≡ reverse ys ++ reverse xs
-reverse-++-commute {a} [] ys = begin
- reverse ys ≡⟨ P.sym $ proj₂ {a = a} {b = a} LM.identity _ ⟩
- reverse ys ++ [] ≡⟨ P.refl ⟩
- reverse ys ++ reverse [] ∎
- where open P.≡-Reasoning
+reverse-++-commute : ∀ {a} {A : Set a} (xs ys : List A) →
+ reverse (xs ++ ys) ≡ reverse ys ++ reverse xs
+reverse-++-commute [] ys = P.sym (proj₂ LM.identity _)
reverse-++-commute (x ∷ xs) ys = begin
reverse (x ∷ xs ++ ys) ≡⟨ unfold-reverse x (xs ++ ys) ⟩
- reverse (xs ++ ys) ++ [ x ] ≡⟨ P.cong (λ zs → zs ++ [ x ]) (reverse-++-commute xs ys) ⟩
+ reverse (xs ++ ys) ++ [ x ] ≡⟨ P.cong (_++ [ x ]) (reverse-++-commute xs ys) ⟩
(reverse ys ++ reverse xs) ++ [ x ] ≡⟨ LM.assoc (reverse ys) _ _ ⟩
- reverse ys ++ (reverse xs ++ [ x ]) ≡⟨ P.sym $ P.cong (λ zs → reverse ys ++ zs) (unfold-reverse x xs) ⟩
+ reverse ys ++ (reverse xs ++ [ x ]) ≡⟨ P.sym $ P.cong (reverse ys ++_) (unfold-reverse x xs) ⟩
reverse ys ++ reverse (x ∷ xs) ∎
where open P.≡-Reasoning
@@ -380,20 +376,54 @@ reverse-map-commute f [] = refl
reverse-map-commute f (x ∷ xs) = begin
map f (reverse (x ∷ xs)) ≡⟨ P.cong (map f) $ unfold-reverse x xs ⟩
map f (reverse xs ∷ʳ x) ≡⟨ map-++-commute f (reverse xs) ([ x ]) ⟩
- map f (reverse xs) ∷ʳ f x ≡⟨ P.cong (λ y → y ∷ʳ f x) $ reverse-map-commute f xs ⟩
+ map f (reverse xs) ∷ʳ f x ≡⟨ P.cong (_∷ʳ f x) $ reverse-map-commute f xs ⟩
reverse (map f xs) ∷ʳ f x ≡⟨ P.sym $ unfold-reverse (f x) (map f xs) ⟩
reverse (map f (x ∷ xs)) ∎
where open P.≡-Reasoning
-reverse-involutive : ∀ {a} {A : Set a} → Involutive (reverse {A = A})
+reverse-involutive : ∀ {a} {A : Set a} → Involutive _≡_ (reverse {A = A})
reverse-involutive [] = refl
reverse-involutive (x ∷ xs) = begin
reverse (reverse (x ∷ xs)) ≡⟨ P.cong reverse $ unfold-reverse x xs ⟩
reverse (reverse xs ∷ʳ x) ≡⟨ reverse-++-commute (reverse xs) ([ x ]) ⟩
- x ∷ reverse (reverse (xs)) ≡⟨ P.cong (λ y → x ∷ y) $ reverse-involutive xs ⟩
+ x ∷ reverse (reverse (xs)) ≡⟨ P.cong (x ∷_) $ reverse-involutive xs ⟩
x ∷ xs ∎
where open P.≡-Reasoning
+reverse-foldr : ∀ {a b} {A : Set a} {B : Set b}
+ (f : A → B → B) x ys →
+ foldr f x (reverse ys) ≡ foldl (flip f) x ys
+reverse-foldr f x [] = refl
+reverse-foldr f x (y ∷ ys) = begin
+ foldr f x (reverse (y ∷ ys)) ≡⟨ P.cong (foldr f x) (unfold-reverse y ys) ⟩
+ foldr f x ((reverse ys) ∷ʳ y) ≡⟨ foldr-∷ʳ f x y (reverse ys) ⟩
+ foldr f (f y x) (reverse ys) ≡⟨ reverse-foldr f (f y x) ys ⟩
+ foldl (flip f) (f y x) ys ∎
+ where open P.≡-Reasoning
+
+reverse-foldl : ∀ {a b} {A : Set a} {B : Set b}
+ (f : A → B → A) x ys →
+ foldl f x (reverse ys) ≡ foldr (flip f) x ys
+reverse-foldl f x [] = refl
+reverse-foldl f x (y ∷ ys) = begin
+ foldl f x (reverse (y ∷ ys)) ≡⟨ P.cong (foldl f x) (unfold-reverse y ys) ⟩
+ foldl f x ((reverse ys) ∷ʳ y) ≡⟨ foldl-∷ʳ f x y (reverse ys) ⟩
+ f (foldl f x (reverse ys)) y ≡⟨ P.cong (flip f y) (reverse-foldl f x ys) ⟩
+ f (foldr (flip f) x ys) y ∎
+ where open P.≡-Reasoning
+
+length-reverse : ∀ {a} {A : Set a} (xs : List A) →
+ length (reverse xs) ≡ length xs
+length-reverse [] = refl
+length-reverse (x ∷ xs) = begin
+ length (reverse (x ∷ xs)) ≡⟨ P.cong length $ unfold-reverse x xs ⟩
+ length (reverse xs ∷ʳ x) ≡⟨ length-++ (reverse xs) ⟩
+ length (reverse xs) + 1 ≡⟨ P.cong (_+ 1) (length-reverse xs) ⟩
+ length xs + 1 ≡⟨ +-comm _ 1 ⟩
+ suc (length xs) ∎
+ where open P.≡-Reasoning
+
+------------------------------------------------------------------------
-- The list monad.
module Monad where
diff --git a/src/Data/List/Reverse.agda b/src/Data/List/Reverse.agda
index a1347b6..40c5296 100644
--- a/src/Data/List/Reverse.agda
+++ b/src/Data/List/Reverse.agda
@@ -9,7 +9,7 @@ module Data.List.Reverse where
open import Data.List.Base
open import Data.Nat
import Data.Nat.Properties as Nat
-open import Induction.Nat
+open import Induction.Nat using (<′-Rec; <′-rec)
open import Relation.Binary.PropositionalEquality
-- If you want to traverse a list from the end, then you can use the
@@ -22,7 +22,7 @@ data Reverse {A : Set} : List A → Set where
_∶_∶ʳ_ : ∀ xs (rs : Reverse xs) (x : A) → Reverse (xs ∷ʳ x)
reverseView : ∀ {A} (xs : List A) → Reverse xs
-reverseView {A} xs = <-rec Pred rev (length xs) xs refl
+reverseView {A} xs = <′-rec Pred rev (length xs) xs refl
where
Pred : ℕ → Set
Pred n = (xs : List A) → length xs ≡ n → Reverse xs
@@ -31,7 +31,7 @@ reverseView {A} xs = <-rec Pred rev (length xs) xs refl
lem [] = ≤′-refl
lem (x ∷ xs) = Nat.s≤′s (lem xs)
- rev : (n : ℕ) → <-Rec Pred n → Pred n
+ rev : (n : ℕ) → <′-Rec Pred n → Pred n
rev n rec xs eq with initLast xs
rev n rec .[] eq | [] = []
rev .(length (xs ∷ʳ x)) rec .(xs ∷ʳ x) refl | xs ∷ʳ' x
diff --git a/src/Data/Maybe.agda b/src/Data/Maybe.agda
index 2b6d837..df29181 100644
--- a/src/Data/Maybe.agda
+++ b/src/Data/Maybe.agda
@@ -6,22 +6,24 @@
module Data.Maybe where
-open import Level
+open import Category.Functor
+open import Category.Monad
+open import Category.Monad.Identity
open import Function
+open import Level
open import Relation.Nullary
+import Relation.Nullary.Decidable as Dec
+open import Relation.Binary as B
+open import Relation.Unary as U
------------------------------------------------------------------------
--- The type and some operations
+-- The base type and some operations
open import Data.Maybe.Base public
------------------------------------------------------------------------
-- Maybe monad
-open import Category.Functor
-open import Category.Monad
-open import Category.Monad.Identity
-
functor : ∀ {f} → RawFunctor (Maybe {a = f})
functor = record
{ _<$>_ = map
@@ -57,8 +59,6 @@ monadPlus {f} = record
------------------------------------------------------------------------
-- Equality
-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
@@ -67,54 +67,60 @@ 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
- }
+Eq-refl : ∀ {a ℓ} {A : Set a} {_≈_ : Rel A ℓ} →
+ Reflexive _≈_ → Reflexive (Eq _≈_)
+Eq-refl refl {just x} = just refl
+Eq-refl refl {nothing} = nothing
+
+Eq-sym : ∀ {a ℓ} {A : Set a} {_≈_ : Rel A ℓ} →
+ Symmetric _≈_ → Symmetric (Eq _≈_)
+Eq-sym sym (just x≈y) = just (sym x≈y)
+Eq-sym sym nothing = nothing
+
+Eq-trans : ∀ {a ℓ} {A : Set a} {_≈_ : Rel A ℓ} →
+ Transitive _≈_ → Transitive (Eq _≈_)
+Eq-trans trans (just x≈y) (just y≈z) = just (trans x≈y y≈z)
+Eq-trans trans nothing nothing = nothing
+
+Eq-dec : ∀ {a ℓ} {A : Set a} {_≈_ : Rel A ℓ} →
+ B.Decidable _≈_ → B.Decidable (Eq _≈_)
+Eq-dec dec (just x) (just y) with dec x y
+... | yes x≈y = yes (just x≈y)
+... | no x≉y = no (x≉y ∘ drop-just)
+Eq-dec dec (just x) nothing = no λ()
+Eq-dec dec nothing (just y) = no λ()
+Eq-dec dec nothing nothing = yes nothing
+
+Eq-isEquivalence : ∀ {a ℓ} {A : Set a} {_≈_ : Rel A ℓ}
+ → IsEquivalence _≈_ → IsEquivalence (Eq _≈_)
+Eq-isEquivalence isEq = record
+ { refl = Eq-refl (IsEquivalence.refl isEq)
+ ; sym = Eq-sym (IsEquivalence.sym isEq)
+ ; trans = Eq-trans (IsEquivalence.trans isEq)
}
- 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
+Eq-isDecEquivalence : ∀ {a ℓ} {A : Set a} {_≈_ : Rel A ℓ}
+ → IsDecEquivalence _≈_ → IsDecEquivalence (Eq _≈_)
+Eq-isDecEquivalence isDecEq = record
+ { isEquivalence = Eq-isEquivalence
+ (IsDecEquivalence.isEquivalence isDecEq)
+ ; _≟_ = Eq-dec (IsDecEquivalence._≟_ isDecEq)
+ }
- 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
+setoid : ∀ {ℓ₁ ℓ₂} → Setoid ℓ₁ ℓ₂ → Setoid _ _
+setoid S = record {
+ isEquivalence = Eq-isEquivalence (Setoid.isEquivalence S)
+ }
decSetoid : ∀ {ℓ₁ ℓ₂} → DecSetoid ℓ₁ ℓ₂ → DecSetoid _ _
-decSetoid D = record
- { isDecEquivalence = record
- { isEquivalence = Setoid.isEquivalence (setoid (DecSetoid.setoid D))
- ; _≟_ = _≟_
- }
+decSetoid D = record {
+ isDecEquivalence = Eq-isDecEquivalence
+ (DecSetoid.isDecEquivalence 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 are preserving decidability
-import Relation.Nullary.Decidable as Dec
-open import Relation.Unary as U
-
anyDec : ∀ {a p} {A : Set a} {P : A → Set p} →
U.Decidable P → U.Decidable (Any P)
anyDec p nothing = no λ()
diff --git a/src/Data/Maybe/Base.agda b/src/Data/Maybe/Base.agda
index 7f8e78a..33317e5 100644
--- a/src/Data/Maybe/Base.agda
+++ b/src/Data/Maybe/Base.agda
@@ -14,8 +14,8 @@ data Maybe {a} (A : Set a) : Set a where
just : (x : A) → Maybe A
nothing : Maybe A
-{-# HASKELL type AgdaMaybe a b = Maybe b #-}
-{-# COMPILED_DATA Maybe MAlonzo.Code.Data.Maybe.Base.AgdaMaybe Just Nothing #-}
+{-# FOREIGN GHC type AgdaMaybe a b = Maybe b #-}
+{-# COMPILE GHC Maybe = data MAlonzo.Code.Data.Maybe.Base.AgdaMaybe (Just | Nothing) #-}
------------------------------------------------------------------------
-- Some operations
diff --git a/src/Data/Nat.agda b/src/Data/Nat.agda
index ea13190..522478e 100644
--- a/src/Data/Nat.agda
+++ b/src/Data/Nat.agda
@@ -26,58 +26,3 @@ open import Data.Nat.Base public
eq? : ∀ {a} {A : Set a} → A ↣ ℕ → Decidable {A = A} _≡_
eq? inj = Dec.via-injection inj _≟_
-
-------------------------------------------------------------------------
--- Some properties
-
-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
- refl′ : _≡_ ⇒ _≤_
- refl′ {zero} refl = z≤n
- refl′ {suc m} refl = s≤s (refl′ refl)
-
- antisym : Antisymmetric _≡_ _≤_
- antisym z≤n z≤n = refl
- antisym (s≤s m≤n) (s≤s n≤m) with antisym m≤n n≤m
- ... | refl = refl
-
- trans : Transitive _≤_
- trans z≤n _ = z≤n
- trans (s≤s m≤n) (s≤s n≤o) = s≤s (trans m≤n n≤o)
-
- total : Total _≤_
- total zero _ = inj₁ z≤n
- total _ zero = inj₂ z≤n
- total (suc m) (suc n) with total m n
- ... | inj₁ m≤n = inj₁ (s≤s m≤n)
- ... | inj₂ n≤m = inj₂ (s≤s n≤m)
-
-import Relation.Binary.PartialOrderReasoning as POR
-module ≤-Reasoning where
- open POR (DecTotalOrder.poset decTotalOrder) public
- renaming (_≈⟨_⟩_ to _≡⟨_⟩_)
-
- infixr 2 _<⟨_⟩_
-
- _<⟨_⟩_ : ∀ x {y z} → x < y → y IsRelatedTo z → suc x IsRelatedTo z
- x <⟨ x<y ⟩ y≤z = suc x ≤⟨ x<y ⟩ y≤z
diff --git a/src/Data/Nat/Base.agda b/src/Data/Nat/Base.agda
index 8c24abc..52367e1 100644
--- a/src/Data/Nat/Base.agda
+++ b/src/Data/Nat/Base.agda
@@ -90,21 +90,6 @@ erase : ∀ {m n} → m ≤″ n → m ≤″ n
erase (less-than-or-equal eq) = less-than-or-equal (TrustMe.erase eq)
------------------------------------------------------------------------
--- A generalisation of the arithmetic operations
-
-fold : ∀{ℓ} {a : Set ℓ} → a → (a → a) → ℕ → a
-fold z s zero = z
-fold z s (suc n) = s (fold z s n)
-
-module GeneralisedArithmetic {ℓ} {a : Set ℓ} (0# : a) (1+ : a → a) where
-
- add : ℕ → a → a
- add n z = fold z 1+ n
-
- mul : (+ : a → a → a) → (ℕ → a → a)
- mul _+_ n x = fold 0# (λ s → x + s) n
-
-------------------------------------------------------------------------
-- Arithmetic
pred : ℕ → ℕ
@@ -146,6 +131,12 @@ suc m ⊓ suc n = suc (m ⊓ n)
⌈_/2⌉ : ℕ → ℕ
⌈ n /2⌉ = ⌊ suc n /2⌋
+-- Naïve exponentiation
+
+_^_ : ℕ → ℕ → ℕ
+x ^ zero = 1
+x ^ suc n = x * x ^ n
+
------------------------------------------------------------------------
-- Queries
diff --git a/src/Data/Nat/Coprimality.agda b/src/Data/Nat/Coprimality.agda
index 7734a59..85ed06c 100644
--- a/src/Data/Nat/Coprimality.agda
+++ b/src/Data/Nat/Coprimality.agda
@@ -24,7 +24,6 @@ open import Relation.Binary
open import Algebra
private
module P = Poset Div.poset
- module CS = CommutativeSemiring NatProp.commutativeSemiring
-- Coprime m n is inhabited iff m and n are coprime (relatively
-- prime), i.e. if their only common divisor is 1.
@@ -126,8 +125,8 @@ data GCD′ : ℕ → ℕ → ℕ → Set where
gcd-gcd′ : ∀ {d m n} → GCD m n d → GCD′ m n d
gcd-gcd′ g with GCD.commonDivisor g
gcd-gcd′ {zero} g | (divides q₁ refl , divides q₂ refl)
- with q₁ * 0 | CS.*-comm 0 q₁
- | q₂ * 0 | CS.*-comm 0 q₂
+ with q₁ * 0 | NatProp.*-comm 0 q₁
+ | q₂ * 0 | NatProp.*-comm 0 q₂
... | .0 | refl | .0 | refl = gcd-* 1 1 (1-coprimeTo 1)
gcd-gcd′ {suc d} g | (divides q₁ refl , divides q₂ refl) =
gcd-* q₁ q₂ (Bézout-coprime (Bézout.identity g))
@@ -152,7 +151,7 @@ prime⇒coprime (suc (suc m)) _ _ _ _ {1} _ = refl
prime⇒coprime (suc (suc m)) p _ _ _ {0} (divides q 2+m≡q*0 , _) =
⊥-elim $ NatProp.i+1+j≢i 0 (begin
2 + m ≡⟨ 2+m≡q*0 ⟩
- q * 0 ≡⟨ proj₂ CS.zero q ⟩
+ q * 0 ≡⟨ proj₂ NatProp.*-zero q ⟩
0 ∎)
where open PropEq.≡-Reasoning
prime⇒coprime (suc (suc m)) p (suc n) _ 1+n<2+m {suc (suc i)}
@@ -164,7 +163,7 @@ prime⇒coprime (suc (suc m)) p (suc n) _ 1+n<2+m {suc (suc i)}
3 + i ≤⟨ s≤s (∣⇒≤ 2+i∣1+n) ⟩
2 + n ≤⟨ 1+n<2+m ⟩
2 + m ∎)
- where open ≤-Reasoning
+ where open NatProp.≤-Reasoning
2+i′∣2+m : 2 + toℕ (fromℕ≤ i<m) ∣ 2 + m
2+i′∣2+m = PropEq.subst
diff --git a/src/Data/Nat/DivMod.agda b/src/Data/Nat/DivMod.agda
index b76f60f..a07726f 100644
--- a/src/Data/Nat/DivMod.agda
+++ b/src/Data/Nat/DivMod.agda
@@ -9,8 +9,7 @@ module Data.Nat.DivMod where
open import Data.Fin as Fin using (Fin; toℕ)
import Data.Fin.Properties as FinP
open import Data.Nat as Nat
-open import Data.Nat.Properties as NatP
-open import Data.Nat.Properties.Simple
+open import Data.Nat.Properties
open import Relation.Nullary.Decidable
open import Relation.Binary.PropositionalEquality as P using (_≡_)
import Relation.Binary.PropositionalEquality.TrustMe as TrustMe
@@ -18,9 +17,9 @@ import Relation.Binary.PropositionalEquality.TrustMe as TrustMe
open import Agda.Builtin.Nat using ( div-helper; mod-helper )
-open NatP.SemiringSolver
+open SemiringSolver
open P.≡-Reasoning
-open Nat.≤-Reasoning
+open ≤-Reasoning
renaming (begin_ to start_; _∎ to _□; _≡⟨_⟩_ to _≡⟨_⟩′_)
infixl 7 _div_ _mod_ _divMod_
@@ -84,7 +83,7 @@ private
division-lemma mod-acc div-acc zero n = begin
- mod-acc + div-acc * suc s + zero ≡⟨ +-right-identity _ ⟩
+ mod-acc + div-acc * suc s + zero ≡⟨ +-identityʳ _ ⟩
mod-acc + div-acc * suc s ∎
diff --git a/src/Data/Nat/Divisibility.agda b/src/Data/Nat/Divisibility.agda
index 3cee228..8faaa60 100644
--- a/src/Data/Nat/Divisibility.agda
+++ b/src/Data/Nat/Divisibility.agda
@@ -8,205 +8,201 @@ module Data.Nat.Divisibility where
open import Data.Nat as Nat
open import Data.Nat.DivMod
-import Data.Nat.Properties as NatProp
-open import Data.Fin as Fin using (Fin; zero; suc)
+open import Data.Nat.Properties
+open import Data.Fin using (Fin; zero; suc; toℕ)
import Data.Fin.Properties as FP
-open NatProp.SemiringSolver
+open SemiringSolver
open import Algebra
-private
- module CS = CommutativeSemiring NatProp.commutativeSemiring
open import Data.Product
open import Relation.Nullary
open import Relation.Binary
import Relation.Binary.PartialOrderReasoning as PartialOrderReasoning
open import Relation.Binary.PropositionalEquality as PropEq
- using (_≡_; _≢_; refl; sym; cong; subst)
+ using (_≡_; _≢_; refl; sym; trans; cong; cong₂; subst)
open import Function
+------------------------------------------------------------------------
-- m ∣ n is inhabited iff m divides n. Some sources, like Hardy and
-- Wright's "An Introduction to the Theory of Numbers", require m to
-- be non-zero. However, some things become a bit nicer if m is
-- allowed to be zero. For instance, _∣_ becomes a partial order, and
-- the gcd of 0 and 0 becomes defined.
-infix 4 _∣_
+infix 4 _∣_ _∤_
data _∣_ : ℕ → ℕ → Set where
divides : {m n : ℕ} (q : ℕ) (eq : n ≡ q * m) → m ∣ n
--- Extracts the quotient.
+_∤_ : Rel ℕ _
+m ∤ n = ¬ (m ∣ n)
quotient : ∀ {m n} → m ∣ n → ℕ
quotient (divides q _) = q
--- If m divides n, and n is positive, then m ≤ n.
+------------------------------------------------------------------------
+-- _∣_ is a partial order
∣⇒≤ : ∀ {m n} → m ∣ suc n → m ≤ suc n
∣⇒≤ (divides zero ())
∣⇒≤ {m} {n} (divides (suc q) eq) = begin
- m ≤⟨ NatProp.m≤m+n m (q * m) ⟩
+ m ≤⟨ m≤m+n m (q * m) ⟩
suc q * m ≡⟨ sym eq ⟩
suc n ∎
where open ≤-Reasoning
--- _∣_ is a partial order.
+∣-reflexive : _≡_ ⇒ _∣_
+∣-reflexive {n} refl = divides 1 (sym (*-identityˡ n))
+
+∣-refl : Reflexive _∣_
+∣-refl = ∣-reflexive refl
+
+∣-trans : Transitive _∣_
+∣-trans (divides p refl) (divides q refl) =
+ divides (q * p) (sym (*-assoc q p _))
+
+∣-antisym : Antisymmetric _≡_ _∣_
+∣-antisym (divides {n = zero} _ _) (divides q refl) = *-comm q 0
+∣-antisym (divides p eq) (divides {n = zero} _ _) =
+ trans (*-comm 0 p) (sym eq)
+∣-antisym (divides {n = suc _} p eq₁) (divides {n = suc _} q eq₂) =
+ ≤-antisym (∣⇒≤ (divides p eq₁)) (∣⇒≤ (divides q eq₂))
+
+∣-isPreorder : IsPreorder _≡_ _∣_
+∣-isPreorder = record
+ { isEquivalence = PropEq.isEquivalence
+ ; reflexive = ∣-reflexive
+ ; trans = ∣-trans
+ }
+
+∣-isPartialOrder : IsPartialOrder _≡_ _∣_
+∣-isPartialOrder = record
+ { isPreorder = ∣-isPreorder
+ ; antisym = ∣-antisym
+ }
poset : Poset _ _ _
poset = record
{ Carrier = ℕ
; _≈_ = _≡_
; _≤_ = _∣_
- ; isPartialOrder = record
- { isPreorder = record
- { isEquivalence = PropEq.isEquivalence
- ; reflexive = reflexive
- ; trans = trans
- }
- ; antisym = antisym
- }
+ ; isPartialOrder = ∣-isPartialOrder
}
- where
- module DTO = DecTotalOrder Nat.decTotalOrder
- open PropEq.≡-Reasoning
-
- reflexive : _≡_ ⇒ _∣_
- reflexive {n} refl = divides 1 (sym $ proj₁ CS.*-identity n)
-
- antisym : Antisymmetric _≡_ _∣_
- antisym (divides {n = zero} q₁ eq₁) (divides {n = n₂} q₂ eq₂) = begin
- n₂ ≡⟨ eq₂ ⟩
- q₂ * 0 ≡⟨ CS.*-comm q₂ 0 ⟩
- 0 ∎
- antisym (divides {n = n₁} q₁ eq₁) (divides {n = zero} q₂ eq₂) = begin
- 0 ≡⟨ CS.*-comm 0 q₁ ⟩
- q₁ * 0 ≡⟨ sym eq₁ ⟩
- n₁ ∎
- antisym (divides {n = suc n₁} q₁ eq₁) (divides {n = suc n₂} q₂ eq₂) =
- DTO.antisym (∣⇒≤ (divides q₁ eq₁)) (∣⇒≤ (divides q₂ eq₂))
-
- trans : Transitive _∣_
- trans (divides q₁ refl) (divides q₂ refl) =
- divides (q₂ * q₁) (sym (CS.*-assoc q₂ q₁ _))
module ∣-Reasoning = PartialOrderReasoning poset
renaming (_≤⟨_⟩_ to _∣⟨_⟩_; _≈⟨_⟩_ to _≡⟨_⟩_)
-private module P = Poset poset
-
--- 1 divides everything.
+------------------------------------------------------------------------
+-- Simple properties of _∣_
-infix 10 1∣_
+infix 10 1∣_ _∣0
1∣_ : ∀ n → 1 ∣ n
-1∣ n = divides n (sym $ proj₂ CS.*-identity n)
-
--- Everything divides 0.
-
-infix 10 _∣0
+1∣ n = divides n (sym (*-identityʳ n))
_∣0 : ∀ n → n ∣ 0
n ∣0 = divides 0 refl
--- 0 only divides 0.
+n∣n : ∀ {n} → n ∣ n
+n∣n = ∣-refl
-0∣⇒≡0 : ∀ {n} → 0 ∣ n → n ≡ 0
-0∣⇒≡0 {n} 0∣n = P.antisym (n ∣0) 0∣n
+n|m*n : ∀ m {n} → n ∣ m * n
+n|m*n m = divides m refl
--- Only 1 divides 1.
+0∣⇒≡0 : ∀ {n} → 0 ∣ n → n ≡ 0
+0∣⇒≡0 {n} 0∣n = ∣-antisym (n ∣0) 0∣n
∣1⇒≡1 : ∀ {n} → n ∣ 1 → n ≡ 1
-∣1⇒≡1 {n} n∣1 = P.antisym n∣1 (1∣ n)
+∣1⇒≡1 {n} n∣1 = ∣-antisym n∣1 (1∣ n)
--- If i divides m and n, then i divides their sum.
-
-∣-+ : ∀ {i m n} → i ∣ m → i ∣ n → i ∣ m + n
-∣-+ (divides {m = i} q refl) (divides q' refl) =
- divides (q + q') (sym $ proj₂ CS.distrib i q q')
-
--- If i divides m and n, then i divides their difference.
-
-∣-∸ : ∀ {i m n} → i ∣ m + n → i ∣ m → i ∣ n
-∣-∸ (divides {m = i} q' eq) (divides q refl) =
- divides (q' ∸ q)
- (sym $ NatProp.im≡jm+n⇒[i∸j]m≡n q' q i _ $ sym eq)
-
--- A simple lemma: n divides kn.
-
-∣-* : ∀ k {n} → n ∣ k * n
-∣-* k = divides k refl
+------------------------------------------------------------------------
+-- Operators and divisibility
+
+∣m∣n⇒∣m+n : ∀ {i m n} → i ∣ m → i ∣ n → i ∣ m + n
+∣m∣n⇒∣m+n (divides p refl) (divides q refl) =
+ divides (p + q) (sym (*-distribʳ-+ _ p q))
+
+∣m+n|m⇒|n : ∀ {i m n} → i ∣ m + n → i ∣ m → i ∣ n
+∣m+n|m⇒|n {i} {m} {n} (divides p m+n≡p*i) (divides q m≡q*i) =
+ divides (p ∸ q) (begin
+ n ≡⟨ sym (m+n∸n≡m n m) ⟩
+ n + m ∸ m ≡⟨ cong (_∸ m) (+-comm n m) ⟩
+ m + n ∸ m ≡⟨ cong₂ _∸_ m+n≡p*i m≡q*i ⟩
+ p * i ∸ q * i ≡⟨ sym (*-distribʳ-∸ i p q) ⟩
+ (p ∸ q) * i ∎)
+ where open PropEq.≡-Reasoning
--- If i divides j, then ki divides kj.
+∣m∸n∣n⇒∣m : ∀ i {m n} → n ≤ m → i ∣ m ∸ n → i ∣ n → i ∣ m
+∣m∸n∣n⇒∣m i {m} {n} n≤m (divides p m∸n≡p*i) (divides q n≡q*o) =
+ divides (p + q) (begin
+ m ≡⟨ sym (m+n∸m≡n n≤m) ⟩
+ n + (m ∸ n) ≡⟨ +-comm n (m ∸ n) ⟩
+ m ∸ n + n ≡⟨ cong₂ _+_ m∸n≡p*i n≡q*o ⟩
+ p * i + q * i ≡⟨ sym (*-distribʳ-+ i p q) ⟩
+ (p + q) * i ∎)
+ where open PropEq.≡-Reasoning
*-cong : ∀ {i j} k → i ∣ j → k * i ∣ k * j
-*-cong {i} {j} k (divides q eq) = divides q lemma
- where
- open PropEq.≡-Reasoning
- lemma = begin
- k * j ≡⟨ cong (_*_ k) eq ⟩
- k * (q * i) ≡⟨ solve 3 (λ k q i → k :* (q :* i)
- := q :* (k :* i))
- refl k q i ⟩
- q * (k * i) ∎
-
--- If ki divides kj, and k is positive, then i divides j.
+*-cong {i} {j} k (divides q j≡q*i) = divides q (begin
+ k * j ≡⟨ cong (_*_ k) j≡q*i ⟩
+ k * (q * i) ≡⟨ sym (*-assoc k q i) ⟩
+ (k * q) * i ≡⟨ cong (_* i) (*-comm k q) ⟩
+ (q * k) * i ≡⟨ *-assoc q k i ⟩
+ q * (k * i) ∎)
+ where open PropEq.≡-Reasoning
/-cong : ∀ {i j} k → suc k * i ∣ suc k * j → i ∣ j
-/-cong {i} {j} k (divides q eq) = divides q lemma
- where
- open PropEq.≡-Reasoning
- k′ = suc k
- lemma = NatProp.cancel-*-right j (q * i) (begin
- j * k′ ≡⟨ CS.*-comm j k′ ⟩
- k′ * j ≡⟨ eq ⟩
- q * (k′ * i) ≡⟨ solve 3 (λ q k i → q :* (k :* i)
- := q :* i :* k)
- refl q k′ i ⟩
- q * i * k′ ∎)
+/-cong {i} {j} k (divides q eq) =
+ divides q (*-cancelʳ-≡ j (q * i) (begin
+ j * (suc k) ≡⟨ *-comm j (suc k) ⟩
+ (suc k) * j ≡⟨ eq ⟩
+ q * ((suc k) * i) ≡⟨ cong (q *_) (*-comm (suc k) i) ⟩
+ q * (i * (suc k)) ≡⟨ sym (*-assoc q i (suc k)) ⟩
+ (q * i) * (suc k) ∎))
+ where open PropEq.≡-Reasoning
-- If the remainder after division is non-zero, then the divisor does
-- not divide the dividend.
nonZeroDivisor-lemma
- : ∀ m q (r : Fin (1 + m)) → Fin.toℕ r ≢ 0 →
- ¬ (1 + m) ∣ (Fin.toℕ r + q * (1 + m))
+ : ∀ m q (r : Fin (1 + m)) → toℕ r ≢ 0 →
+ 1 + m ∤ toℕ r + q * (1 + m)
nonZeroDivisor-lemma m zero r r≢zero (divides zero eq) = r≢zero $ begin
- Fin.toℕ r
- ≡⟨ sym $ proj₁ CS.*-identity (Fin.toℕ r) ⟩
- 1 * Fin.toℕ r
- ≡⟨ eq ⟩
- 0
- ∎
+ toℕ r ≡⟨ sym (*-identityˡ (toℕ r)) ⟩
+ 1 * toℕ r ≡⟨ eq ⟩
+ 0 ∎
where open PropEq.≡-Reasoning
nonZeroDivisor-lemma m zero r r≢zero (divides (suc q) eq) =
- NatProp.¬i+1+j≤i m $ begin
- m + suc (q * suc m)
- ≡⟨ solve 2 (λ m q → m :+ (con 1 :+ q) := con 1 :+ m :+ q)
- refl m (q * suc m) ⟩
- suc (m + q * suc m)
- ≡⟨ sym eq ⟩
- 1 * Fin.toℕ r
- ≡⟨ proj₁ CS.*-identity (Fin.toℕ r) ⟩
- Fin.toℕ r
- ≤⟨ ≤-pred $ FP.bounded r ⟩
- m
- ∎
+ ¬i+1+j≤i m $ begin
+ m + suc (q * suc m) ≡⟨ +-suc m (q * suc m) ⟩
+ suc (m + q * suc m) ≡⟨ sym eq ⟩
+ 1 * toℕ r ≡⟨ *-identityˡ (toℕ r) ⟩
+ toℕ r ≤⟨ ≤-pred $ FP.bounded r ⟩
+ m ∎
where open ≤-Reasoning
nonZeroDivisor-lemma m (suc q) r r≢zero d =
- nonZeroDivisor-lemma m q r r≢zero (∣-∸ d' P.refl)
+ nonZeroDivisor-lemma m q r r≢zero (∣m+n|m⇒|n d' ∣-refl)
where
lem = solve 3 (λ m r q → r :+ (m :+ q) := m :+ (r :+ q))
- refl (suc m) (Fin.toℕ r) (q * suc m)
- d' = subst (λ x → (1 + m) ∣ x) lem d
+ refl (suc m) (toℕ r) (q * suc m)
+ d' = subst (1 + m ∣_) lem d
-- Divisibility is decidable.
infix 4 _∣?_
_∣?_ : Decidable _∣_
-zero ∣? zero = yes (0 ∣0)
-zero ∣? suc n = no ((λ ()) ∘′ 0∣⇒≡0)
-suc m ∣? n with n divMod suc m
-suc m ∣? .(q * suc m) | result q zero refl =
+zero ∣? zero = yes (0 ∣0)
+zero ∣? suc n = no ((λ()) ∘′ 0∣⇒≡0)
+suc m ∣? n with n divMod suc m
+suc m ∣? .(q * suc m) | result q zero refl =
yes $ divides q refl
-suc m ∣? .(1 + Fin.toℕ r + q * suc m) | result q (suc r) refl =
+suc m ∣? .(1 + toℕ r + q * suc m) | result q (suc r) refl =
no $ nonZeroDivisor-lemma m q (suc r) (λ())
+
+------------------------------------------------------------------------
+-- DEPRECATED - please use new names as continuing support for the old
+-- names is not guaranteed.
+
+∣-+ = ∣m∣n⇒∣m+n
+∣-∸ = ∣m+n|m⇒|n
+∣-* = n|m*n
diff --git a/src/Data/Nat/GCD.agda b/src/Data/Nat/GCD.agda
index d78b0e5..de37813 100644
--- a/src/Data/Nat/GCD.agda
+++ b/src/Data/Nat/GCD.agda
@@ -11,9 +11,10 @@ open import Data.Nat.Divisibility as Div
open import Relation.Binary
private module P = Poset Div.poset
open import Data.Product
-open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
+open import Relation.Binary.PropositionalEquality as PropEq using (_≡_; subst)
+open import Relation.Nullary using (Dec; yes; no)
open import Induction
-open import Induction.Nat
+open import Induction.Nat using (<′-Rec; <′-rec-builder)
open import Induction.Lexicographic
open import Function
open import Data.Nat.GCD.Lemmas
@@ -64,10 +65,10 @@ module GCD where
step : ∀ {n k d} → GCD n k d → GCD n (n + k) d
step g with GCD.commonDivisor g
- step {n} {k} {d} g | (d₁ , d₂) = is (d₁ , ∣-+ d₁ d₂) greatest′
+ step {n} {k} {d} g | (d₁ , d₂) = is (d₁ , ∣m∣n⇒∣m+n d₁ d₂) greatest′
where
greatest′ : ∀ {d′} → d′ ∣ n × d′ ∣ n + k → d′ ∣ d
- greatest′ (d₁ , d₂) = GCD.greatest g (d₁ , ∣-∸ d₂ d₁)
+ greatest′ (d₁ , d₂) = GCD.greatest g (d₁ , ∣m+n|m⇒|n d₂ d₁)
open GCD public using (GCD) hiding (module GCD)
@@ -154,12 +155,12 @@ module Bézout where
-- Euclidean algorithm.
lemma : (m n : ℕ) → Lemma m n
- lemma m n = build [ <-rec-builder ⊗ <-rec-builder ] P gcd (m , n)
+ lemma m n = build [ <′-rec-builder ⊗ <′-rec-builder ] P gcd (m , n)
where
P : ℕ × ℕ → Set
P (m , n) = Lemma m n
- gcd : ∀ p → (<-Rec ⊗ <-Rec) P p → P p
+ gcd : ∀ p → (<′-Rec ⊗ <′-Rec) P p → P p
gcd (zero , n ) rec = Lemma.base n
gcd (suc m , zero ) rec = Lemma.sym (Lemma.base (suc m))
gcd (suc m , suc n ) rec with compare m n
@@ -183,3 +184,11 @@ module Bézout where
gcd : (m n : ℕ) → ∃ λ d → GCD m n d
gcd m n with Bézout.lemma m n
gcd m n | Bézout.result d g _ = (d , g)
+
+-- gcd as a proposition is decidable
+
+gcd? : (m n d : ℕ) → Dec (GCD m n d)
+gcd? m n d with gcd m n
+... | d′ , p with d′ ≟ d
+... | no ¬g = no (λ p′ → ¬g (GCD.unique p p′))
+... | yes g = yes (subst (GCD m n) g p)
diff --git a/src/Data/Nat/GCD/Lemmas.agda b/src/Data/Nat/GCD/Lemmas.agda
index 8b5a6c0..188e7d4 100644
--- a/src/Data/Nat/GCD/Lemmas.agda
+++ b/src/Data/Nat/GCD/Lemmas.agda
@@ -138,7 +138,7 @@ lem₁₀ : ∀ {a′} b c {d} e f → let a = suc a′ in
lem₁₀ {a′} b c {d} e f eq =
NatProp.i*j≡1⇒j≡1 (e * f ∸ b * c) d
(NatProp.im≡jm+n⇒[i∸j]m≡n (e * f) (b * c) d 1
- (NatProp.cancel-*-right (e * f * d) (b * c * d + 1) (begin
+ (NatProp.*-cancelʳ-≡ (e * f * d) (b * c * d + 1) (begin
e * f * d * a ≡⟨ solve 4 (λ e f d a → e :* f :* d :* a
:= e :* (f :* d :* a))
refl e f d a ⟩
diff --git a/src/Data/Nat/GeneralisedArithmetic.agda b/src/Data/Nat/GeneralisedArithmetic.agda
new file mode 100644
index 0000000..9220636
--- /dev/null
+++ b/src/Data/Nat/GeneralisedArithmetic.agda
@@ -0,0 +1,88 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- A generalisation of the arithmetic operations
+------------------------------------------------------------------------
+
+module Data.Nat.GeneralisedArithmetic where
+
+open import Data.Nat
+open import Data.Nat.Properties
+open import Function using (_∘′_; _∘_; id)
+open import Relation.Binary.PropositionalEquality
+open ≡-Reasoning
+
+module _ {a} {A : Set a} where
+
+ fold : A → (A → A) → ℕ → A
+ fold z s zero = z
+ fold z s (suc n) = s (fold z s n)
+
+ add : (0# : A) (1+ : A → A) → ℕ → A → A
+ add 0# 1+ n z = fold z 1+ n
+
+ mul : (0# : A) (1+ : A → A) → (+ : A → A → A) → (ℕ → A → A)
+ mul 0# 1+ _+_ n x = fold 0# (λ s → x + s) n
+
+-- Properties
+
+module _ {a} {A : Set a} where
+
+ fold-+ : ∀ {s : A → A} {z : A} →
+ ∀ m {n} → fold z s (m + n) ≡ fold (fold z s n) s m
+ fold-+ zero = refl
+ fold-+ {s = s} (suc m) = cong s (fold-+ m)
+
+ fold-k : ∀ {s : A → A} {z : A} {k} m →
+ fold k (s ∘′_) m z ≡ fold (k z) s m
+ fold-k zero = refl
+ fold-k {s = s} (suc m) = cong s (fold-k m)
+
+ fold-* : ∀ {s : A → A} {z : A} m {n} →
+ fold z s (m * n) ≡ fold z (fold id (s ∘_) n) m
+ fold-* zero = refl
+ fold-* {s = s} {z} (suc m) {n} = let +n = fold id (s ∘′_) n in begin
+ fold z s (n + m * n) ≡⟨ fold-+ n ⟩
+ fold (fold z s (m * n)) s n ≡⟨ cong (λ z → fold z s n) (fold-* m) ⟩
+ fold (fold z +n m) s n ≡⟨ sym (fold-k n) ⟩
+ fold z +n (suc m) ∎
+
+ fold-pull : ∀ {s : A → A} {z : A} (g : A → A → A) (p : A)
+ (eqz : g z p ≡ p)
+ (eqs : ∀ l → s (g l p) ≡ g (s l) p) →
+ ∀ m → fold p s m ≡ g (fold z s m) p
+ fold-pull _ _ eqz _ zero = sym eqz
+ fold-pull {s = s} {z} g p eqz eqs (suc m) = begin
+ s (fold p s m) ≡⟨ cong s (fold-pull g p eqz eqs m) ⟩
+ s (g (fold z s m) p) ≡⟨ eqs (fold z s m) ⟩
+ g (s (fold z s m)) p ∎
+
+id-is-fold : ∀ m → fold zero suc m ≡ m
+id-is-fold zero = refl
+id-is-fold (suc m) = cong suc (id-is-fold m)
+
++-is-fold : ∀ m {n} → fold n suc m ≡ m + n
++-is-fold zero = refl
++-is-fold (suc m) = cong suc (+-is-fold m)
+
+*-is-fold : ∀ m {n} → fold zero (n +_) m ≡ m * n
+*-is-fold zero = refl
+*-is-fold (suc m) {n} = cong (n +_) (*-is-fold m)
+
+^-is-fold : ∀ {m} n → fold 1 (m *_) n ≡ m ^ n
+^-is-fold zero = refl
+^-is-fold {m} (suc n) = cong (m *_) (^-is-fold n)
+
+*+-is-fold : ∀ m n {p} → fold p (n +_) m ≡ m * n + p
+*+-is-fold m n {p} = begin
+ fold p (n +_) m ≡⟨ fold-pull _+_ p refl
+ (λ l → sym (+-assoc n l p)) m ⟩
+ fold 0 (n +_) m + p ≡⟨ cong (_+ p) (*-is-fold m) ⟩
+ m * n + p ∎
+
+^*-is-fold : ∀ m n {p} → fold p (m *_) n ≡ m ^ n * p
+^*-is-fold m n {p} = begin
+ fold p (m *_) n ≡⟨ fold-pull _*_ p (*-identityˡ p)
+ (λ l → sym (*-assoc m l p)) n ⟩
+ fold 1 (m *_) n * p ≡⟨ cong (_* p) (^-is-fold n) ⟩
+ m ^ n * p ∎
diff --git a/src/Data/Nat/InfinitelyOften.agda b/src/Data/Nat/InfinitelyOften.agda
index f5c552a..dfeb6c7 100644
--- a/src/Data/Nat/InfinitelyOften.agda
+++ b/src/Data/Nat/InfinitelyOften.agda
@@ -12,7 +12,7 @@ open import Category.Monad
open import Data.Empty
open import Function
open import Data.Nat
-import Data.Nat.Properties as NatProp
+open import Data.Nat.Properties
open import Data.Product as Prod hiding (map)
open import Data.Sum hiding (map)
open import Relation.Binary.PropositionalEquality
@@ -20,8 +20,6 @@ open import Relation.Nullary
open import Relation.Nullary.Negation
open import Relation.Unary using (_∪_; _⊆_)
open RawMonad (¬¬-Monad {p = Level.zero})
-private
- module NatLattice = DistributiveLattice NatProp.distributiveLattice
-- Only true finitely often.
@@ -37,12 +35,12 @@ _∪-Fin_ {P} {Q} (i , ¬p) (j , ¬q) = (i ⊔ j , helper)
helper : ∀ k → i ⊔ j ≤ k → ¬ (P ∪ Q) k
helper k i⊔j≤k (inj₁ p) = ¬p k (begin
- i ≤⟨ NatProp.m≤m⊔n i j ⟩
+ i ≤⟨ m≤m⊔n i j ⟩
i ⊔ j ≤⟨ i⊔j≤k ⟩
k ∎) p
helper k i⊔j≤k (inj₂ q) = ¬q k (begin
- j ≤⟨ NatProp.m≤m⊔n j i ⟩
- j ⊔ i ≡⟨ NatLattice.∧-comm j i ⟩
+ j ≤⟨ m≤m⊔n j i ⟩
+ j ⊔ i ≡⟨ ⊔-comm j i ⟩
i ⊔ j ≤⟨ i⊔j≤k ⟩
k ∎) q
@@ -88,4 +86,4 @@ twoDifferentWitnesses
twoDifferentWitnesses inf =
witness inf >>= λ w₁ →
witness (up (1 + proj₁ w₁) inf) >>= λ w₂ →
- return (_ , _ , NatProp.m≢1+m+n (proj₁ w₁) , proj₂ w₁ , proj₂ w₂)
+ return (_ , _ , m≢1+m+n (proj₁ w₁) , proj₂ w₁ , proj₂ w₂)
diff --git a/src/Data/Nat/LCM.agda b/src/Data/Nat/LCM.agda
index 7292123..e044d33 100644
--- a/src/Data/Nat/LCM.agda
+++ b/src/Data/Nat/LCM.agda
@@ -20,7 +20,6 @@ open import Algebra
open import Relation.Binary
private
module P = Poset Div.poset
- module CS = CommutativeSemiring NatProp.commutativeSemiring
------------------------------------------------------------------------
-- Least common multiple (lcm).
@@ -91,7 +90,7 @@ lcm .(q₁ * d) .(q₂ * d) | (d , gcd-* q₁ q₂ q₁-q₂-coprime) =
q₂∣q₃ : q₂ ∣ q₃
q₂∣q₃ = coprime-divisor (Coprime.sym q₁-q₂-coprime)
- (divides q₄ $ NatProp.cancel-*-right _ _ (begin
+ (divides q₄ $ NatProp.*-cancelʳ-≡ _ _ (begin
q₁ * q₃ * d′ ≡⟨ lem₁ q₁ q₃ d′ ⟩
q₃ * (q₁ * d′) ≡⟨ PropEq.sym eq₃ ⟩
m ≡⟨ eq₄ ⟩
@@ -116,10 +115,10 @@ lcm .(q₁ * d) .(q₂ * d) | (d , gcd-* q₁ q₂ q₁-q₂-coprime) =
gcd*lcm : ∀ {i j d m} → GCD i j d → LCM i j m → i * j ≡ d * m
gcd*lcm {i} {j} {d} {m} g l with LCM.unique l (proj₂ (lcm i j))
gcd*lcm {i} {j} {d} .{proj₁ (lcm i j)} g l | refl with gcd′ i j
-gcd*lcm .{q₁ * d′} .{q₂ * d′} {d} .{q₁ * q₂ * d′} g l | refl | (d′ , gcd-* q₁ q₂ q₁-q₂-coprime)
+gcd*lcm .{q₁ * d′} .{q₂ * d′} {d} g l | refl | (d′ , gcd-* q₁ q₂ q₁-q₂-coprime)
with GCD.unique g
(gcd′-gcd (gcd-* q₁ q₂ q₁-q₂-coprime))
-gcd*lcm .{q₁ * d} .{q₂ * d} {d} .{q₁ * q₂ * d} g l | refl | (.d , gcd-* q₁ q₂ q₁-q₂-coprime) | refl =
+gcd*lcm .{q₁ * d} .{q₂ * d} {d} g l | refl | (.d , gcd-* q₁ q₂ q₁-q₂-coprime) | refl =
solve 3 (λ q₁ q₂ d → q₁ :* d :* (q₂ :* d)
:= d :* (q₁ :* q₂ :* d))
refl q₁ q₂ d
diff --git a/src/Data/Nat/Properties.agda b/src/Data/Nat/Properties.agda
index a28023c..f3df5b3 100644
--- a/src/Data/Nat/Properties.agda
+++ b/src/Data/Nat/Properties.agda
@@ -9,251 +9,219 @@
module Data.Nat.Properties where
-open import Data.Nat as Nat
-open ≤-Reasoning
- renaming (begin_ to start_; _∎ to _□; _≡⟨_⟩_ to _≡⟨_⟩'_)
open import Relation.Binary
-open DecTotalOrder Nat.decTotalOrder using () renaming (refl to ≤-refl)
open import Function
open import Algebra
+import Algebra.RingSolver.Simple as Solver
+import Algebra.RingSolver.AlmostCommutativeRing as ACR
open import Algebra.Structures
-open import Relation.Nullary
-open import Relation.Binary.PropositionalEquality as PropEq
- using (_≡_; _≢_; refl; sym; cong; cong₂)
-open PropEq.≡-Reasoning
-import Algebra.FunctionProperties as P; open P (_≡_ {A = ℕ})
+open import Data.Nat as Nat
open import Data.Product
open import Data.Sum
+open import Relation.Nullary
+open import Relation.Nullary.Negation using (contradiction)
+open import Relation.Binary.PropositionalEquality
+open import Algebra.FunctionProperties (_≡_ {A = ℕ})
+ hiding (LeftCancellative; RightCancellative; Cancellative)
+open import Algebra.FunctionProperties
+ using (LeftCancellative; RightCancellative; Cancellative)
+open import Algebra.FunctionProperties.Consequences (setoid ℕ)
+
+open ≡-Reasoning
------------------------------------------------------------------------
+-- Properties of _≡_
--- basic lemmas about (ℕ, +, *, 0, 1):
-open import Data.Nat.Properties.Simple
-
--- (ℕ, +, *, 0, 1) is a commutative semiring
-isCommutativeSemiring : IsCommutativeSemiring _≡_ _+_ _*_ 0 1
-isCommutativeSemiring = record
- { +-isCommutativeMonoid = record
- { isSemigroup = record
- { isEquivalence = PropEq.isEquivalence
- ; assoc = +-assoc
- ; ∙-cong = cong₂ _+_
- }
- ; identityˡ = λ _ → refl
- ; comm = +-comm
- }
- ; *-isCommutativeMonoid = record
- { isSemigroup = record
- { isEquivalence = PropEq.isEquivalence
- ; assoc = *-assoc
- ; ∙-cong = cong₂ _*_
- }
- ; identityˡ = +-right-identity
- ; comm = *-comm
- }
- ; distribʳ = distribʳ-*-+
- ; zeroˡ = λ _ → refl
- }
+suc-injective : ∀ {m n} → suc m ≡ suc n → m ≡ n
+suc-injective refl = refl
-commutativeSemiring : CommutativeSemiring _ _
-commutativeSemiring = record
- { _+_ = _+_
- ; _*_ = _*_
- ; 0# = 0
- ; 1# = 1
- ; isCommutativeSemiring = isCommutativeSemiring
+≡-isDecEquivalence : IsDecEquivalence (_≡_ {A = ℕ})
+≡-isDecEquivalence = record
+ { isEquivalence = isEquivalence
+ ; _≟_ = _≟_
}
-import Algebra.RingSolver.Simple as Solver
-import Algebra.RingSolver.AlmostCommutativeRing as ACR
-module SemiringSolver =
- Solver (ACR.fromCommutativeSemiring commutativeSemiring) _≟_
+≡-decSetoid : DecSetoid _ _
+≡-decSetoid = record
+ { Carrier = ℕ
+ ; _≈_ = _≡_
+ ; isDecEquivalence = ≡-isDecEquivalence
+ }
------------------------------------------------------------------------
--- (ℕ, ⊔, ⊓, 0) is a commutative semiring without one
-
-private
-
- ⊔-assoc : Associative _⊔_
- ⊔-assoc zero _ _ = refl
- ⊔-assoc (suc m) zero o = refl
- ⊔-assoc (suc m) (suc n) zero = refl
- ⊔-assoc (suc m) (suc n) (suc o) = cong suc $ ⊔-assoc m n o
-
- ⊔-identity : Identity 0 _⊔_
- ⊔-identity = (λ _ → refl) , n⊔0≡n
- where
- n⊔0≡n : RightIdentity 0 _⊔_
- n⊔0≡n zero = refl
- n⊔0≡n (suc n) = refl
-
- ⊔-comm : Commutative _⊔_
- ⊔-comm zero n = sym $ proj₂ ⊔-identity n
- ⊔-comm (suc m) zero = refl
- ⊔-comm (suc m) (suc n) =
- begin
- suc m ⊔ suc n
- ≡⟨ refl ⟩
- suc (m ⊔ n)
- ≡⟨ cong suc (⊔-comm m n) ⟩
- suc (n ⊔ m)
- ≡⟨ refl ⟩
- suc n ⊔ suc m
- ∎
-
- ⊓-assoc : Associative _⊓_
- ⊓-assoc zero _ _ = refl
- ⊓-assoc (suc m) zero o = refl
- ⊓-assoc (suc m) (suc n) zero = refl
- ⊓-assoc (suc m) (suc n) (suc o) = cong suc $ ⊓-assoc m n o
-
- ⊓-zero : Zero 0 _⊓_
- ⊓-zero = (λ _ → refl) , n⊓0≡0
- where
- n⊓0≡0 : RightZero 0 _⊓_
- n⊓0≡0 zero = refl
- n⊓0≡0 (suc n) = refl
-
- ⊓-comm : Commutative _⊓_
- ⊓-comm zero n = sym $ proj₂ ⊓-zero n
- ⊓-comm (suc m) zero = refl
- ⊓-comm (suc m) (suc n) =
- begin
- suc m ⊓ suc n
- ≡⟨ refl ⟩
- suc (m ⊓ n)
- ≡⟨ cong suc (⊓-comm m n) ⟩
- suc (n ⊓ m)
- ≡⟨ refl ⟩
- suc n ⊓ suc m
- ∎
-
- distrib-⊓-⊔ : _⊓_ DistributesOver _⊔_
- distrib-⊓-⊔ = (distribˡ-⊓-⊔ , distribʳ-⊓-⊔)
- where
- distribʳ-⊓-⊔ : _⊓_ DistributesOverʳ _⊔_
- distribʳ-⊓-⊔ (suc m) (suc n) (suc o) = cong suc $ distribʳ-⊓-⊔ m n o
- distribʳ-⊓-⊔ (suc m) (suc n) zero = cong suc $ refl
- distribʳ-⊓-⊔ (suc m) zero o = refl
- distribʳ-⊓-⊔ zero n o = begin
- (n ⊔ o) ⊓ 0 ≡⟨ ⊓-comm (n ⊔ o) 0 ⟩
- 0 ⊓ (n ⊔ o) ≡⟨ refl ⟩
- 0 ⊓ n ⊔ 0 ⊓ o ≡⟨ ⊓-comm 0 n ⟨ cong₂ _⊔_ ⟩ ⊓-comm 0 o ⟩
- n ⊓ 0 ⊔ o ⊓ 0 ∎
-
- distribˡ-⊓-⊔ : _⊓_ DistributesOverˡ _⊔_
- distribˡ-⊓-⊔ m n o = begin
- m ⊓ (n ⊔ o) ≡⟨ ⊓-comm m _ ⟩
- (n ⊔ o) ⊓ m ≡⟨ distribʳ-⊓-⊔ m n o ⟩
- n ⊓ m ⊔ o ⊓ m ≡⟨ ⊓-comm n m ⟨ cong₂ _⊔_ ⟩ ⊓-comm o m ⟩
- m ⊓ n ⊔ m ⊓ o ∎
-
-⊔-⊓-0-isCommutativeSemiringWithoutOne
- : IsCommutativeSemiringWithoutOne _≡_ _⊔_ _⊓_ 0
-⊔-⊓-0-isCommutativeSemiringWithoutOne = record
- { isSemiringWithoutOne = record
- { +-isCommutativeMonoid = record
- { isSemigroup = record
- { isEquivalence = PropEq.isEquivalence
- ; assoc = ⊔-assoc
- ; ∙-cong = cong₂ _⊔_
- }
- ; identityˡ = proj₁ ⊔-identity
- ; comm = ⊔-comm
- }
- ; *-isSemigroup = record
- { isEquivalence = PropEq.isEquivalence
- ; assoc = ⊓-assoc
- ; ∙-cong = cong₂ _⊓_
- }
- ; distrib = distrib-⊓-⊔
- ; zero = ⊓-zero
- }
- ; *-comm = ⊓-comm
+-- Properties of _≤_
+
+-- Relation-theoretic properties of _≤_
+≤-reflexive : _≡_ ⇒ _≤_
+≤-reflexive {zero} refl = z≤n
+≤-reflexive {suc m} refl = s≤s (≤-reflexive refl)
+
+≤-refl : Reflexive _≤_
+≤-refl = ≤-reflexive refl
+
+≤-antisym : Antisymmetric _≡_ _≤_
+≤-antisym z≤n z≤n = refl
+≤-antisym (s≤s m≤n) (s≤s n≤m) with ≤-antisym m≤n n≤m
+... | refl = refl
+
+≤-trans : Transitive _≤_
+≤-trans z≤n _ = z≤n
+≤-trans (s≤s m≤n) (s≤s n≤o) = s≤s (≤-trans m≤n n≤o)
+
+≤-total : Total _≤_
+≤-total zero _ = inj₁ z≤n
+≤-total _ zero = inj₂ z≤n
+≤-total (suc m) (suc n) with ≤-total m n
+... | inj₁ m≤n = inj₁ (s≤s m≤n)
+... | inj₂ n≤m = inj₂ (s≤s n≤m)
+
+≤-isPreorder : IsPreorder _≡_ _≤_
+≤-isPreorder = record
+ { isEquivalence = isEquivalence
+ ; reflexive = ≤-reflexive
+ ; trans = ≤-trans
}
-⊔-⊓-0-commutativeSemiringWithoutOne : CommutativeSemiringWithoutOne _ _
-⊔-⊓-0-commutativeSemiringWithoutOne = record
- { _+_ = _⊔_
- ; _*_ = _⊓_
- ; 0# = 0
- ; isCommutativeSemiringWithoutOne =
- ⊔-⊓-0-isCommutativeSemiringWithoutOne
+≤-preorder : Preorder _ _ _
+≤-preorder = record
+ { isPreorder = ≤-isPreorder
}
-------------------------------------------------------------------------
--- (ℕ, ⊓, ⊔) is a lattice
-
-private
-
- absorptive-⊓-⊔ : Absorptive _⊓_ _⊔_
- absorptive-⊓-⊔ = abs-⊓-⊔ , abs-⊔-⊓
- where
- abs-⊔-⊓ : _⊔_ Absorbs _⊓_
- abs-⊔-⊓ zero n = refl
- abs-⊔-⊓ (suc m) zero = refl
- abs-⊔-⊓ (suc m) (suc n) = cong suc $ abs-⊔-⊓ m n
-
- abs-⊓-⊔ : _⊓_ Absorbs _⊔_
- abs-⊓-⊔ zero n = refl
- abs-⊓-⊔ (suc m) (suc n) = cong suc $ abs-⊓-⊔ m n
- abs-⊓-⊔ (suc m) zero = cong suc $
- begin
- m ⊓ m
- ≡⟨ cong (_⊓_ m) $ sym $ proj₂ ⊔-identity m ⟩
- m ⊓ (m ⊔ 0)
- ≡⟨ abs-⊓-⊔ m zero ⟩
- m
- ∎
-
-isDistributiveLattice : IsDistributiveLattice _≡_ _⊓_ _⊔_
-isDistributiveLattice = record
- { isLattice = record
- { isEquivalence = PropEq.isEquivalence
- ; ∨-comm = ⊓-comm
- ; ∨-assoc = ⊓-assoc
- ; ∨-cong = cong₂ _⊓_
- ; ∧-comm = ⊔-comm
- ; ∧-assoc = ⊔-assoc
- ; ∧-cong = cong₂ _⊔_
- ; absorptive = absorptive-⊓-⊔
- }
- ; ∨-∧-distribʳ = proj₂ distrib-⊓-⊔
+≤-isPartialOrder : IsPartialOrder _≡_ _≤_
+≤-isPartialOrder = record
+ { isPreorder = ≤-isPreorder
+ ; antisym = ≤-antisym
}
-distributiveLattice : DistributiveLattice _ _
-distributiveLattice = record
- { _∨_ = _⊓_
- ; _∧_ = _⊔_
- ; isDistributiveLattice = isDistributiveLattice
+≤-isTotalOrder : IsTotalOrder _≡_ _≤_
+≤-isTotalOrder = record
+ { isPartialOrder = ≤-isPartialOrder
+ ; total = ≤-total
}
--- Selectivity of ⊓ and ⊔
-
-⊓-sel : Selective _⊓_
-⊓-sel zero _ = inj₁ refl
-⊓-sel (suc m) zero = inj₂ refl
-⊓-sel (suc m) (suc n) with ⊓-sel m n
-... | inj₁ m⊓n≡m = inj₁ (cong suc m⊓n≡m)
-... | inj₂ m⊓n≡n = inj₂ (cong suc m⊓n≡n)
+≤-totalOrder : TotalOrder _ _ _
+≤-totalOrder = record
+ { isTotalOrder = ≤-isTotalOrder
+ }
-⊔-sel : Selective _⊔_
-⊔-sel zero _ = inj₂ refl
-⊔-sel (suc m) zero = inj₁ refl
-⊔-sel (suc m) (suc n) with ⊔-sel m n
-... | inj₁ m⊔n≡m = inj₁ (cong suc m⊔n≡m)
-... | inj₂ m⊔n≡n = inj₂ (cong suc m⊔n≡n)
+≤-isDecTotalOrder : IsDecTotalOrder _≡_ _≤_
+≤-isDecTotalOrder = record
+ { isTotalOrder = ≤-isTotalOrder
+ ; _≟_ = _≟_
+ ; _≤?_ = _≤?_
+ }
-------------------------------------------------------------------------
--- Converting between ≤ and ≤′
+≤-decTotalOrder : DecTotalOrder _ _ _
+≤-decTotalOrder = record
+ { isDecTotalOrder = ≤-isDecTotalOrder
+ }
+-- Other properties of _≤_
≤-step : ∀ {m n} → m ≤ n → m ≤ 1 + n
≤-step z≤n = z≤n
≤-step (s≤s m≤n) = s≤s (≤-step m≤n)
-≤′⇒≤ : _≤′_ ⇒ _≤_
-≤′⇒≤ ≤′-refl = ≤-refl
-≤′⇒≤ (≤′-step m≤′n) = ≤-step (≤′⇒≤ m≤′n)
+n≤1+n : ∀ n → n ≤ 1 + n
+n≤1+n _ = ≤-step ≤-refl
+
+1+n≰n : ∀ {n} → ¬ 1 + n ≤ n
+1+n≰n (s≤s le) = 1+n≰n le
+
+pred-mono : pred Preserves _≤_ ⟶ _≤_
+pred-mono z≤n = z≤n
+pred-mono (s≤s le) = le
+
+≤pred⇒≤ : ∀ {m n} → m ≤ pred n → m ≤ n
+≤pred⇒≤ {m} {zero} le = le
+≤pred⇒≤ {m} {suc n} le = ≤-step le
+
+≤⇒pred≤ : ∀ {m n} → m ≤ n → pred m ≤ n
+≤⇒pred≤ {zero} le = le
+≤⇒pred≤ {suc m} le = ≤-trans (n≤1+n m) le
+
+------------------------------------------------------------------------
+-- Properties of _<_
+
+-- Relation theoretic properties of _<_
+_<?_ : Decidable _<_
+x <? y = suc x ≤? y
+
+<-irrefl : Irreflexive _≡_ _<_
+<-irrefl refl (s≤s n<n) = <-irrefl refl n<n
+
+<-asym : Asymmetric _<_
+<-asym (s≤s n<m) (s≤s m<n) = <-asym n<m m<n
+
+<-trans : Transitive _<_
+<-trans (s≤s i≤j) (s≤s j<k) = s≤s (≤-trans i≤j (≤⇒pred≤ j<k))
+
+<-transʳ : Trans _≤_ _<_ _<_
+<-transʳ m≤n (s≤s n≤o) = s≤s (≤-trans m≤n n≤o)
+
+<-transˡ : Trans _<_ _≤_ _<_
+<-transˡ (s≤s m≤n) (s≤s n≤o) = s≤s (≤-trans m≤n n≤o)
+
+<-cmp : Trichotomous _≡_ _<_
+<-cmp zero zero = tri≈ (λ()) refl (λ())
+<-cmp zero (suc n) = tri< (s≤s z≤n) (λ()) (λ())
+<-cmp (suc m) zero = tri> (λ()) (λ()) (s≤s z≤n)
+<-cmp (suc m) (suc n) with <-cmp m n
+... | tri< ≤ ≢ ≱ = tri< (s≤s ≤) (≢ ∘ suc-injective) (≱ ∘ ≤-pred)
+... | tri≈ ≰ ≡ ≱ = tri≈ (≰ ∘ ≤-pred) (cong suc ≡) (≱ ∘ ≤-pred)
+... | tri> ≰ ≢ ≥ = tri> (≰ ∘ ≤-pred) (≢ ∘ suc-injective) (s≤s ≥)
+
+<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
+<-isStrictTotalOrder = record
+ { isEquivalence = isEquivalence
+ ; trans = <-trans
+ ; compare = <-cmp
+ }
+
+<-strictTotalOrder : StrictTotalOrder _ _ _
+<-strictTotalOrder = record
+ { isStrictTotalOrder = <-isStrictTotalOrder
+ }
+
+-- Other properties of _<_
+<⇒≤pred : ∀ {m n} → m < n → m ≤ pred n
+<⇒≤pred (s≤s le) = le
+
+<⇒≤ : _<_ ⇒ _≤_
+<⇒≤ (s≤s m≤n) = ≤-trans m≤n (≤-step ≤-refl)
+
+<⇒≢ : _<_ ⇒ _≢_
+<⇒≢ m<n refl = 1+n≰n m<n
+
+<⇒≱ : _<_ ⇒ _≱_
+<⇒≱ (s≤s m+1≤n) (s≤s n≤m) = <⇒≱ m+1≤n n≤m
+
+<⇒≯ : _<_ ⇒ _≯_
+<⇒≯ (s≤s m<n) (s≤s n<m) = <⇒≯ m<n n<m
+
+≰⇒≮ : _≰_ ⇒ _≮_
+≰⇒≮ m≰n 1+m≤n = m≰n (<⇒≤ 1+m≤n)
+
+≰⇒> : _≰_ ⇒ _>_
+≰⇒> {zero} z≰n = contradiction z≤n z≰n
+≰⇒> {suc m} {zero} _ = s≤s z≤n
+≰⇒> {suc m} {suc n} m≰n = s≤s (≰⇒> (m≰n ∘ s≤s))
+
+≰⇒≥ : _≰_ ⇒ _≥_
+≰⇒≥ = <⇒≤ ∘ ≰⇒>
+
+≮⇒≥ : _≮_ ⇒ _≥_
+≮⇒≥ {_} {zero} _ = z≤n
+≮⇒≥ {zero} {suc j} 1≮j+1 = contradiction (s≤s z≤n) 1≮j+1
+≮⇒≥ {suc i} {suc j} i+1≮j+1 = s≤s (≮⇒≥ (i+1≮j+1 ∘ s≤s))
+
+≤+≢⇒< : ∀ {m n} → m ≤ n → m ≢ n → m < n
+≤+≢⇒< {_} {zero} z≤n m≢n = contradiction refl m≢n
+≤+≢⇒< {_} {suc n} z≤n m≢n = s≤s z≤n
+≤+≢⇒< {_} {suc n} (s≤s m≤n) 1+m≢1+n =
+ s≤s (≤+≢⇒< m≤n (1+m≢1+n ∘ cong suc))
+
+------------------------------------------------------------------------
+-- Properties of _≤′_
z≤′n : ∀ {n} → zero ≤′ n
z≤′n {zero} = ≤′-refl
@@ -263,21 +231,161 @@ s≤′s : ∀ {m n} → m ≤′ n → suc m ≤′ suc n
s≤′s ≤′-refl = ≤′-refl
s≤′s (≤′-step m≤′n) = ≤′-step (s≤′s m≤′n)
+≤′⇒≤ : _≤′_ ⇒ _≤_
+≤′⇒≤ ≤′-refl = ≤-refl
+≤′⇒≤ (≤′-step m≤′n) = ≤-step (≤′⇒≤ m≤′n)
+
≤⇒≤′ : _≤_ ⇒ _≤′_
≤⇒≤′ z≤n = z≤′n
≤⇒≤′ (s≤s m≤n) = s≤′s (≤⇒≤′ m≤n)
------------------------------------------------------------------------
--- Various order-related properties
+-- Properties of _≤″_
-≤-steps : ∀ {m n} k → m ≤ n → m ≤ k + n
-≤-steps zero m≤n = m≤n
-≤-steps (suc k) m≤n = ≤-step (≤-steps k m≤n)
+≤″⇒≤ : _≤″_ ⇒ _≤_
+≤″⇒≤ {zero} (less-than-or-equal refl) = z≤n
+≤″⇒≤ {suc m} (less-than-or-equal refl) =
+ s≤s (≤″⇒≤ (less-than-or-equal refl))
+
+≤⇒≤″ : _≤_ ⇒ _≤″_
+≤⇒≤″ m≤n = less-than-or-equal (proof m≤n)
+ where
+ k : ∀ m n → m ≤ n → ℕ
+ k zero n _ = n
+ k (suc m) zero ()
+ k (suc m) (suc n) m≤n = k m n (≤-pred m≤n)
+
+ proof : ∀ {m n} (m≤n : m ≤ n) → m + k m n m≤n ≡ n
+ proof z≤n = refl
+ proof (s≤s m≤n) = cong suc (proof m≤n)
+
+------------------------------------------------------------------------
+-- Properties of _+_
+
+-- Algebraic properties of _+_
++-suc : ∀ m n → m + suc n ≡ suc (m + n)
++-suc zero n = refl
++-suc (suc m) n = cong suc (+-suc m n)
+
++-assoc : Associative _+_
++-assoc zero _ _ = refl
++-assoc (suc m) n o = cong suc (+-assoc m n o)
+
++-identityˡ : LeftIdentity 0 _+_
++-identityˡ _ = refl
+
++-identityʳ : RightIdentity 0 _+_
++-identityʳ zero = refl
++-identityʳ (suc n) = cong suc (+-identityʳ n)
+
++-identity : Identity 0 _+_
++-identity = +-identityˡ , +-identityʳ
+
++-comm : Commutative _+_
++-comm zero n = sym (+-identityʳ n)
++-comm (suc m) n = begin
+ suc m + n ≡⟨⟩
+ suc (m + n) ≡⟨ cong suc (+-comm m n) ⟩
+ suc (n + m) ≡⟨ sym (+-suc n m) ⟩
+ n + suc m ∎
+
++-isSemigroup : IsSemigroup _≡_ _+_
++-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = +-assoc
+ ; ∙-cong = cong₂ _+_
+ }
+
++-0-isCommutativeMonoid : IsCommutativeMonoid _≡_ _+_ 0
++-0-isCommutativeMonoid = record
+ { isSemigroup = +-isSemigroup
+ ; identityˡ = +-identityˡ
+ ; comm = +-comm
+ }
+
+-- Other properties of _+_ and _≡_
+
++-cancelˡ-≡ : LeftCancellative _≡_ _+_
++-cancelˡ-≡ zero eq = eq
++-cancelˡ-≡ (suc m) eq = +-cancelˡ-≡ m (cong pred eq)
+
++-cancelʳ-≡ : RightCancellative _≡_ _+_
++-cancelʳ-≡ = comm+cancelˡ⇒cancelʳ +-comm +-cancelˡ-≡
+
++-cancel-≡ : Cancellative _≡_ _+_
++-cancel-≡ = +-cancelˡ-≡ , +-cancelʳ-≡
+
+m≢1+m+n : ∀ m {n} → m ≢ suc (m + n)
+m≢1+m+n zero ()
+m≢1+m+n (suc m) eq = m≢1+m+n m (cong pred eq)
+
+i+1+j≢i : ∀ i {j} → i + suc j ≢ i
+i+1+j≢i zero ()
+i+1+j≢i (suc i) = (i+1+j≢i i) ∘ suc-injective
+
+i+j≡0⇒i≡0 : ∀ i {j} → i + j ≡ 0 → i ≡ 0
+i+j≡0⇒i≡0 zero eq = refl
+i+j≡0⇒i≡0 (suc i) ()
+
+i+j≡0⇒j≡0 : ∀ i {j} → i + j ≡ 0 → j ≡ 0
+i+j≡0⇒j≡0 i {j} i+j≡0 = i+j≡0⇒i≡0 j (trans (+-comm j i) (i+j≡0))
+
+-- Properties of _+_ and orderings
+
++-cancelˡ-≤ : LeftCancellative _≤_ _+_
++-cancelˡ-≤ zero le = le
++-cancelˡ-≤ (suc m) (s≤s le) = +-cancelˡ-≤ m le
+
++-cancelʳ-≤ : RightCancellative _≤_ _+_
++-cancelʳ-≤ {m} n o le =
+ +-cancelˡ-≤ m (subst₂ _≤_ (+-comm n m) (+-comm o m) le)
+
++-cancel-≤ : Cancellative _≤_ _+_
++-cancel-≤ = +-cancelˡ-≤ , +-cancelʳ-≤
m≤m+n : ∀ m n → m ≤ m + n
m≤m+n zero n = z≤n
m≤m+n (suc m) n = s≤s (m≤m+n m n)
+n≤m+n : ∀ m n → n ≤ m + n
+n≤m+n m zero = z≤n
+n≤m+n m (suc n) = subst (suc n ≤_) (sym (+-suc m n)) (s≤s (n≤m+n m n))
+
+≤-steps : ∀ {m n} k → m ≤ n → m ≤ k + n
+≤-steps zero m≤n = m≤n
+≤-steps (suc k) m≤n = ≤-step (≤-steps k m≤n)
+
+m+n≤o⇒m≤o : ∀ m {n o} → m + n ≤ o → m ≤ o
+m+n≤o⇒m≤o zero m+n≤o = z≤n
+m+n≤o⇒m≤o (suc m) (s≤s m+n≤o) = s≤s (m+n≤o⇒m≤o m m+n≤o)
+
+m+n≤o⇒n≤o : ∀ m {n o} → m + n ≤ o → n ≤ o
+m+n≤o⇒n≤o zero n≤o = n≤o
+m+n≤o⇒n≤o (suc m) m+n<o = m+n≤o⇒n≤o m (<⇒≤ m+n<o)
+
++-mono-≤ : _+_ Preserves₂ _≤_ ⟶ _≤_ ⟶ _≤_
++-mono-≤ {_} {m} z≤n o≤p = ≤-trans o≤p (n≤m+n m _)
++-mono-≤ {_} {_} (s≤s m≤n) o≤p = s≤s (+-mono-≤ m≤n o≤p)
+
++-monoˡ-< : _+_ Preserves₂ _<_ ⟶ _≤_ ⟶ _<_
++-monoˡ-< {_} {suc y} (s≤s z≤n) u≤v = s≤s (≤-steps y u≤v)
++-monoˡ-< {_} {_} (s≤s (s≤s x<y)) u≤v = s≤s (+-monoˡ-< (s≤s x<y) u≤v)
+
++-monoʳ-< : _+_ Preserves₂ _≤_ ⟶ _<_ ⟶ _<_
++-monoʳ-< {_} {y} z≤n u<v = ≤-trans u<v (n≤m+n y _)
++-monoʳ-< {_} {_} (s≤s x≤y) u<v = s≤s (+-monoʳ-< x≤y u<v)
+
++-mono-< : _+_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_
++-mono-< x≤y = +-monoʳ-< (<⇒≤ x≤y)
+
+¬i+1+j≤i : ∀ i {j} → i + suc j ≰ i
+¬i+1+j≤i zero ()
+¬i+1+j≤i (suc i) le = ¬i+1+j≤i i (≤-pred le)
+
+m+n≮n : ∀ m n → m + n ≮ n
+m+n≮n zero n = <-irrefl refl
+m+n≮n (suc m) (suc n) (s≤s m+n<n) = m+n≮n m (suc n) (≤-step m+n<n)
+
m≤′m+n : ∀ m n → m ≤′ m + n
m≤′m+n m n = ≤⇒≤′ (m≤m+n m n)
@@ -285,126 +393,405 @@ n≤′m+n : ∀ m n → n ≤′ m + n
n≤′m+n zero n = ≤′-refl
n≤′m+n (suc m) n = ≤′-step (n≤′m+n m n)
-n≤m+n : ∀ m n → n ≤ m + n
-n≤m+n m n = ≤′⇒≤ (n≤′m+n m n)
+------------------------------------------------------------------------
+-- Properties of _*_
+
++-*-suc : ∀ m n → m * suc n ≡ m + m * n
++-*-suc zero n = refl
++-*-suc (suc m) n = begin
+ suc m * suc n ≡⟨⟩
+ suc n + m * suc n ≡⟨ cong (suc n +_) (+-*-suc m n) ⟩
+ suc n + (m + m * n) ≡⟨⟩
+ suc (n + (m + m * n)) ≡⟨ cong suc (sym (+-assoc n m (m * n))) ⟩
+ suc (n + m + m * n) ≡⟨ cong (λ x → suc (x + m * n)) (+-comm n m) ⟩
+ suc (m + n + m * n) ≡⟨ cong suc (+-assoc m n (m * n)) ⟩
+ suc (m + (n + m * n)) ≡⟨⟩
+ suc m + suc m * n ∎
+
+*-identityˡ : LeftIdentity 1 _*_
+*-identityˡ x = +-identityʳ x
+
+*-identityʳ : RightIdentity 1 _*_
+*-identityʳ zero = refl
+*-identityʳ (suc x) = cong suc (*-identityʳ x)
+
+*-identity : Identity 1 _*_
+*-identity = *-identityˡ , *-identityʳ
+
+*-zeroˡ : LeftZero 0 _*_
+*-zeroˡ _ = refl
+
+*-zeroʳ : RightZero 0 _*_
+*-zeroʳ zero = refl
+*-zeroʳ (suc n) = *-zeroʳ n
+
+*-zero : Zero 0 _*_
+*-zero = *-zeroˡ , *-zeroʳ
+
+*-comm : Commutative _*_
+*-comm zero n = sym (*-zeroʳ n)
+*-comm (suc m) n = begin
+ suc m * n ≡⟨⟩
+ n + m * n ≡⟨ cong (n +_) (*-comm m n) ⟩
+ n + n * m ≡⟨ sym (+-*-suc n m) ⟩
+ n * suc m ∎
+
+*-distribʳ-+ : _*_ DistributesOverʳ _+_
+*-distribʳ-+ m zero o = refl
+*-distribʳ-+ m (suc n) o = begin
+ (suc n + o) * m ≡⟨⟩
+ m + (n + o) * m ≡⟨ cong (m +_) (*-distribʳ-+ m n o) ⟩
+ m + (n * m + o * m) ≡⟨ sym (+-assoc m (n * m) (o * m)) ⟩
+ m + n * m + o * m ≡⟨⟩
+ suc n * m + o * m ∎
+
+*-distribˡ-+ : _*_ DistributesOverˡ _+_
+*-distribˡ-+ = comm+distrʳ⇒distrˡ (cong₂ _+_) *-comm *-distribʳ-+
+
+*-distrib-+ : _*_ DistributesOver _+_
+*-distrib-+ = *-distribˡ-+ , *-distribʳ-+
+
+*-assoc : Associative _*_
+*-assoc zero n o = refl
+*-assoc (suc m) n o = begin
+ (suc m * n) * o ≡⟨⟩
+ (n + m * n) * o ≡⟨ *-distribʳ-+ o n (m * n) ⟩
+ n * o + (m * n) * o ≡⟨ cong (n * o +_) (*-assoc m n o) ⟩
+ n * o + m * (n * o) ≡⟨⟩
+ suc m * (n * o) ∎
+
+*-isSemigroup : IsSemigroup _≡_ _*_
+*-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = *-assoc
+ ; ∙-cong = cong₂ _*_
+ }
-n≤1+n : ∀ n → n ≤ 1 + n
-n≤1+n _ = ≤-step ≤-refl
+*-1-isCommutativeMonoid : IsCommutativeMonoid _≡_ _*_ 1
+*-1-isCommutativeMonoid = record
+ { isSemigroup = *-isSemigroup
+ ; identityˡ = *-identityˡ
+ ; comm = *-comm
+ }
-1+n≰n : ∀ {n} → ¬ 1 + n ≤ n
-1+n≰n (s≤s le) = 1+n≰n le
+*-+-isCommutativeSemiring : IsCommutativeSemiring _≡_ _+_ _*_ 0 1
+*-+-isCommutativeSemiring = record
+ { +-isCommutativeMonoid = +-0-isCommutativeMonoid
+ ; *-isCommutativeMonoid = *-1-isCommutativeMonoid
+ ; distribʳ = *-distribʳ-+
+ ; zeroˡ = *-zeroˡ
+ }
-≤pred⇒≤ : ∀ m n → m ≤ pred n → m ≤ n
-≤pred⇒≤ m zero le = le
-≤pred⇒≤ m (suc n) le = ≤-step le
+*-+-commutativeSemiring : CommutativeSemiring _ _
+*-+-commutativeSemiring = record
+ { isCommutativeSemiring = *-+-isCommutativeSemiring
+ }
-≤⇒pred≤ : ∀ m n → m ≤ n → pred m ≤ n
-≤⇒pred≤ zero n le = le
-≤⇒pred≤ (suc m) n le = start
- m ≤⟨ n≤1+n m ⟩
- suc m ≤⟨ le ⟩
- n □
+-- Other properties of _*_
-<⇒≤pred : ∀ {m n} → m < n → m ≤ pred n
-<⇒≤pred (s≤s le) = le
+*-cancelʳ-≡ : ∀ i j {k} → i * suc k ≡ j * suc k → i ≡ j
+*-cancelʳ-≡ zero zero eq = refl
+*-cancelʳ-≡ zero (suc j) ()
+*-cancelʳ-≡ (suc i) zero ()
+*-cancelʳ-≡ (suc i) (suc j) {k} eq =
+ cong suc (*-cancelʳ-≡ i j (+-cancelˡ-≡ (suc k) eq))
-¬i+1+j≤i : ∀ i {j} → ¬ i + suc j ≤ i
-¬i+1+j≤i zero ()
-¬i+1+j≤i (suc i) le = ¬i+1+j≤i i (≤-pred le)
+*-cancelˡ-≡ : ∀ {i j} k → suc k * i ≡ suc k * j → i ≡ j
+*-cancelˡ-≡ {i} {j} k eq = *-cancelʳ-≡ i j
+ (subst₂ _≡_ (*-comm (suc k) i) (*-comm (suc k) j) eq)
-n∸m≤n : ∀ m n → n ∸ m ≤ n
-n∸m≤n zero n = ≤-refl
-n∸m≤n (suc m) zero = ≤-refl
-n∸m≤n (suc m) (suc n) = start
- n ∸ m ≤⟨ n∸m≤n m n ⟩
- n ≤⟨ n≤1+n n ⟩
- suc n □
+i*j≡0⇒i≡0∨j≡0 : ∀ i {j} → i * j ≡ 0 → i ≡ 0 ⊎ j ≡ 0
+i*j≡0⇒i≡0∨j≡0 zero {j} eq = inj₁ refl
+i*j≡0⇒i≡0∨j≡0 (suc i) {zero} eq = inj₂ refl
+i*j≡0⇒i≡0∨j≡0 (suc i) {suc j} ()
-n≤m+n∸m : ∀ m n → n ≤ m + (n ∸ m)
-n≤m+n∸m m zero = z≤n
-n≤m+n∸m zero (suc n) = ≤-refl
-n≤m+n∸m (suc m) (suc n) = s≤s (n≤m+n∸m m n)
+i*j≡1⇒i≡1 : ∀ i j → i * j ≡ 1 → i ≡ 1
+i*j≡1⇒i≡1 (suc zero) j _ = refl
+i*j≡1⇒i≡1 zero j ()
+i*j≡1⇒i≡1 (suc (suc i)) (suc (suc j)) ()
+i*j≡1⇒i≡1 (suc (suc i)) (suc zero) ()
+i*j≡1⇒i≡1 (suc (suc i)) zero eq =
+ contradiction (trans (*-comm 0 i) eq) λ()
+
+i*j≡1⇒j≡1 : ∀ i j → i * j ≡ 1 → j ≡ 1
+i*j≡1⇒j≡1 i j eq = i*j≡1⇒i≡1 j i (trans (*-comm j i) eq)
+
+*-cancelʳ-≤ : ∀ i j k → i * suc k ≤ j * suc k → i ≤ j
+*-cancelʳ-≤ zero _ _ _ = z≤n
+*-cancelʳ-≤ (suc i) zero _ ()
+*-cancelʳ-≤ (suc i) (suc j) k le =
+ s≤s (*-cancelʳ-≤ i j k (+-cancelˡ-≤ (suc k) le))
+
+*-mono-≤ : _*_ Preserves₂ _≤_ ⟶ _≤_ ⟶ _≤_
+*-mono-≤ z≤n _ = z≤n
+*-mono-≤ (s≤s m≤n) u≤v = +-mono-≤ u≤v (*-mono-≤ m≤n u≤v)
+
+*-mono-< : _*_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_
+*-mono-< (s≤s z≤n) (s≤s u≤v) = s≤s z≤n
+*-mono-< (s≤s (s≤s m≤n)) (s≤s u≤v) =
+ +-mono-< (s≤s u≤v) (*-mono-< (s≤s m≤n) (s≤s u≤v))
+
+*-monoˡ-< : ∀ n → (_* suc n) Preserves _<_ ⟶ _<_
+*-monoˡ-< n (s≤s z≤n) = s≤s z≤n
+*-monoˡ-< n (s≤s (s≤s m≤o)) =
+ +-monoʳ-< (≤-refl {suc n}) (*-monoˡ-< n (s≤s m≤o))
+
+*-monoʳ-< : ∀ n → (suc n *_) Preserves _<_ ⟶ _<_
+*-monoʳ-< zero (s≤s m≤o) = +-mono-≤ (s≤s m≤o) z≤n
+*-monoʳ-< (suc n) (s≤s m≤o) =
+ +-mono-≤ (s≤s m≤o) (<⇒≤ (*-monoʳ-< n (s≤s m≤o)))
+
+------------------------------------------------------------------------
+-- Properties of _^_
+
+^-distribˡ-+-* : ∀ m n p → m ^ (n + p) ≡ m ^ n * m ^ p
+^-distribˡ-+-* m zero p = sym (+-identityʳ (m ^ p))
+^-distribˡ-+-* m (suc n) p = begin
+ m * (m ^ (n + p)) ≡⟨ cong (m *_) (^-distribˡ-+-* m n p) ⟩
+ m * ((m ^ n) * (m ^ p)) ≡⟨ sym (*-assoc m _ _) ⟩
+ (m * (m ^ n)) * (m ^ p) ∎
+
+i^j≡0⇒i≡0 : ∀ i j → i ^ j ≡ 0 → i ≡ 0
+i^j≡0⇒i≡0 i zero ()
+i^j≡0⇒i≡0 i (suc j) eq = [ id , i^j≡0⇒i≡0 i j ]′ (i*j≡0⇒i≡0∨j≡0 i eq)
+
+i^j≡1⇒j≡0∨i≡1 : ∀ i j → i ^ j ≡ 1 → j ≡ 0 ⊎ i ≡ 1
+i^j≡1⇒j≡0∨i≡1 i zero _ = inj₁ refl
+i^j≡1⇒j≡0∨i≡1 i (suc j) eq = inj₂ (i*j≡1⇒i≡1 i (i ^ j) eq)
+
+------------------------------------------------------------------------
+-- Properties of _⊔_ and _⊓_
+
+⊔-assoc : Associative _⊔_
+⊔-assoc zero _ _ = refl
+⊔-assoc (suc m) zero o = refl
+⊔-assoc (suc m) (suc n) zero = refl
+⊔-assoc (suc m) (suc n) (suc o) = cong suc $ ⊔-assoc m n o
+
+⊔-identityˡ : LeftIdentity 0 _⊔_
+⊔-identityˡ _ = refl
+
+⊔-identityʳ : RightIdentity 0 _⊔_
+⊔-identityʳ zero = refl
+⊔-identityʳ (suc n) = refl
+
+⊔-identity : Identity 0 _⊔_
+⊔-identity = ⊔-identityˡ , ⊔-identityʳ
+
+⊔-comm : Commutative _⊔_
+⊔-comm zero n = sym $ ⊔-identityʳ n
+⊔-comm (suc m) zero = refl
+⊔-comm (suc m) (suc n) = cong suc (⊔-comm m n)
+
+⊔-sel : Selective _⊔_
+⊔-sel zero _ = inj₂ refl
+⊔-sel (suc m) zero = inj₁ refl
+⊔-sel (suc m) (suc n) with ⊔-sel m n
+... | inj₁ m⊔n≡m = inj₁ (cong suc m⊔n≡m)
+... | inj₂ m⊔n≡n = inj₂ (cong suc m⊔n≡n)
+
+⊔-idem : Idempotent _⊔_
+⊔-idem = sel⇒idem ⊔-sel
+
+⊓-assoc : Associative _⊓_
+⊓-assoc zero _ _ = refl
+⊓-assoc (suc m) zero o = refl
+⊓-assoc (suc m) (suc n) zero = refl
+⊓-assoc (suc m) (suc n) (suc o) = cong suc $ ⊓-assoc m n o
+
+⊓-zeroˡ : LeftZero 0 _⊓_
+⊓-zeroˡ _ = refl
+
+⊓-zeroʳ : RightZero 0 _⊓_
+⊓-zeroʳ zero = refl
+⊓-zeroʳ (suc n) = refl
+⊓-zero : Zero 0 _⊓_
+⊓-zero = ⊓-zeroˡ , ⊓-zeroʳ
+
+⊓-comm : Commutative _⊓_
+⊓-comm zero n = sym $ ⊓-zeroʳ n
+⊓-comm (suc m) zero = refl
+⊓-comm (suc m) (suc n) = cong suc (⊓-comm m n)
+
+⊓-sel : Selective _⊓_
+⊓-sel zero _ = inj₁ refl
+⊓-sel (suc m) zero = inj₂ refl
+⊓-sel (suc m) (suc n) with ⊓-sel m n
+... | inj₁ m⊓n≡m = inj₁ (cong suc m⊓n≡m)
+... | inj₂ m⊓n≡n = inj₂ (cong suc m⊓n≡n)
+
+⊓-idem : Idempotent _⊓_
+⊓-idem = sel⇒idem ⊓-sel
+
+⊓-distribʳ-⊔ : _⊓_ DistributesOverʳ _⊔_
+⊓-distribʳ-⊔ (suc m) (suc n) (suc o) = cong suc $ ⊓-distribʳ-⊔ m n o
+⊓-distribʳ-⊔ (suc m) (suc n) zero = cong suc $ refl
+⊓-distribʳ-⊔ (suc m) zero o = refl
+⊓-distribʳ-⊔ zero n o = begin
+ (n ⊔ o) ⊓ 0 ≡⟨ ⊓-comm (n ⊔ o) 0 ⟩
+ 0 ⊓ (n ⊔ o) ≡⟨⟩
+ 0 ⊓ n ⊔ 0 ⊓ o ≡⟨ ⊓-comm 0 n ⟨ cong₂ _⊔_ ⟩ ⊓-comm 0 o ⟩
+ n ⊓ 0 ⊔ o ⊓ 0 ∎
+
+⊓-distribˡ-⊔ : _⊓_ DistributesOverˡ _⊔_
+⊓-distribˡ-⊔ = comm+distrʳ⇒distrˡ (cong₂ _⊔_) ⊓-comm ⊓-distribʳ-⊔
+
+⊓-distrib-⊔ : _⊓_ DistributesOver _⊔_
+⊓-distrib-⊔ = ⊓-distribˡ-⊔ , ⊓-distribʳ-⊔
+
+⊔-abs-⊓ : _⊔_ Absorbs _⊓_
+⊔-abs-⊓ zero n = refl
+⊔-abs-⊓ (suc m) zero = refl
+⊔-abs-⊓ (suc m) (suc n) = cong suc $ ⊔-abs-⊓ m n
+
+⊓-abs-⊔ : _⊓_ Absorbs _⊔_
+⊓-abs-⊔ zero n = refl
+⊓-abs-⊔ (suc m) (suc n) = cong suc $ ⊓-abs-⊔ m n
+⊓-abs-⊔ (suc m) zero = cong suc $ begin
+ m ⊓ m ≡⟨ cong (m ⊓_) $ sym $ ⊔-identityʳ m ⟩
+ m ⊓ (m ⊔ 0) ≡⟨ ⊓-abs-⊔ m zero ⟩
+ m ∎
+
+⊓-⊔-absorptive : Absorptive _⊓_ _⊔_
+⊓-⊔-absorptive = ⊓-abs-⊔ , ⊔-abs-⊓
+
+⊔-isSemigroup : IsSemigroup _≡_ _⊔_
+⊔-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ⊔-assoc
+ ; ∙-cong = cong₂ _⊔_
+ }
+
+⊔-0-isCommutativeMonoid : IsCommutativeMonoid _≡_ _⊔_ 0
+⊔-0-isCommutativeMonoid = record
+ { isSemigroup = ⊔-isSemigroup
+ ; identityˡ = ⊔-identityˡ
+ ; comm = ⊔-comm
+ }
+
+⊓-isSemigroup : IsSemigroup _≡_ _⊓_
+⊓-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ⊓-assoc
+ ; ∙-cong = cong₂ _⊓_
+ }
+
+⊔-⊓-isSemiringWithoutOne : IsSemiringWithoutOne _≡_ _⊔_ _⊓_ 0
+⊔-⊓-isSemiringWithoutOne = record
+ { +-isCommutativeMonoid = ⊔-0-isCommutativeMonoid
+ ; *-isSemigroup = ⊓-isSemigroup
+ ; distrib = ⊓-distrib-⊔
+ ; zero = ⊓-zero
+ }
+
+⊔-⊓-isCommutativeSemiringWithoutOne
+ : IsCommutativeSemiringWithoutOne _≡_ _⊔_ _⊓_ 0
+⊔-⊓-isCommutativeSemiringWithoutOne = record
+ { isSemiringWithoutOne = ⊔-⊓-isSemiringWithoutOne
+ ; *-comm = ⊓-comm
+ }
+
+⊔-⊓-commutativeSemiringWithoutOne : CommutativeSemiringWithoutOne _ _
+⊔-⊓-commutativeSemiringWithoutOne = record
+ { isCommutativeSemiringWithoutOne =
+ ⊔-⊓-isCommutativeSemiringWithoutOne
+ }
+
+⊓-⊔-isLattice : IsLattice _≡_ _⊓_ _⊔_
+⊓-⊔-isLattice = record
+ { isEquivalence = isEquivalence
+ ; ∨-comm = ⊓-comm
+ ; ∨-assoc = ⊓-assoc
+ ; ∨-cong = cong₂ _⊓_
+ ; ∧-comm = ⊔-comm
+ ; ∧-assoc = ⊔-assoc
+ ; ∧-cong = cong₂ _⊔_
+ ; absorptive = ⊓-⊔-absorptive
+ }
+
+⊓-⊔-isDistributiveLattice : IsDistributiveLattice _≡_ _⊓_ _⊔_
+⊓-⊔-isDistributiveLattice = record
+ { isLattice = ⊓-⊔-isLattice
+ ; ∨-∧-distribʳ = ⊓-distribʳ-⊔
+ }
+
+⊓-⊔-distributiveLattice : DistributiveLattice _ _
+⊓-⊔-distributiveLattice = record
+ { isDistributiveLattice = ⊓-⊔-isDistributiveLattice
+ }
+
+-- Ordering properties of _⊔_ and _⊓_
m⊓n≤m : ∀ m n → m ⊓ n ≤ m
m⊓n≤m zero _ = z≤n
m⊓n≤m (suc m) zero = z≤n
m⊓n≤m (suc m) (suc n) = s≤s $ m⊓n≤m m n
+m⊓n≤n : ∀ m n → m ⊓ n ≤ n
+m⊓n≤n m n = subst (_≤ n) (⊓-comm n m) (m⊓n≤m n m)
+
m≤m⊔n : ∀ m n → m ≤ m ⊔ n
m≤m⊔n zero _ = z≤n
m≤m⊔n (suc m) zero = ≤-refl
m≤m⊔n (suc m) (suc n) = s≤s $ m≤m⊔n m n
-⌈n/2⌉≤′n : ∀ n → ⌈ n /2⌉ ≤′ n
-⌈n/2⌉≤′n zero = ≤′-refl
-⌈n/2⌉≤′n (suc zero) = ≤′-refl
-⌈n/2⌉≤′n (suc (suc n)) = s≤′s (≤′-step (⌈n/2⌉≤′n n))
+n≤m⊔n : ∀ m n → n ≤ m ⊔ n
+n≤m⊔n m n = subst (n ≤_) (⊔-comm n m) (m≤m⊔n n m)
-⌊n/2⌋≤′n : ∀ n → ⌊ n /2⌋ ≤′ n
-⌊n/2⌋≤′n zero = ≤′-refl
-⌊n/2⌋≤′n (suc n) = ≤′-step (⌈n/2⌉≤′n n)
+m⊓n≤m⊔n : ∀ m n → m ⊔ n ≤ m ⊔ n
+m⊓n≤m⊔n zero n = ≤-refl
+m⊓n≤m⊔n (suc m) zero = ≤-refl
+m⊓n≤m⊔n (suc m) (suc n) = s≤s (m⊓n≤m⊔n m n)
-<-trans : Transitive _<_
-<-trans {i} {j} {k} i<j j<k = start
- 1 + i ≤⟨ i<j ⟩
- j ≤⟨ n≤1+n j ⟩
- 1 + j ≤⟨ j<k ⟩
- k □
+⊔-mono-≤ : _⊔_ Preserves₂ _≤_ ⟶ _≤_ ⟶ _≤_
+⊔-mono-≤ {x} {y} {u} {v} x≤y u≤v with ⊔-sel x u
+... | inj₁ x⊔u≡x rewrite x⊔u≡x = ≤-trans x≤y (m≤m⊔n y v)
+... | inj₂ x⊔u≡u rewrite x⊔u≡u = ≤-trans u≤v (n≤m⊔n y v)
-≰⇒> : _≰_ ⇒ _>_
-≰⇒> {zero} z≰n with z≰n z≤n
-... | ()
-≰⇒> {suc m} {zero} _ = s≤s z≤n
-≰⇒> {suc m} {suc n} m≰n = s≤s (≰⇒> (m≰n ∘ s≤s))
+⊔-mono-< : _⊔_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_
+⊔-mono-< = ⊔-mono-≤
-------------------------------------------------------------------------
--- Converting between ≤ and ≤″
+⊓-mono-≤ : _⊓_ Preserves₂ _≤_ ⟶ _≤_ ⟶ _≤_
+⊓-mono-≤ {x} {y} {u} {v} x≤y u≤v with ⊓-sel y v
+... | inj₁ y⊓v≡y rewrite y⊓v≡y = ≤-trans (m⊓n≤m x u) x≤y
+... | inj₂ y⊓v≡v rewrite y⊓v≡v = ≤-trans (m⊓n≤n x u) u≤v
-≤″⇒≤ : _≤″_ ⇒ _≤_
-≤″⇒≤ (less-than-or-equal refl) = m≤m+n _ _
+⊓-mono-< : _⊓_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_
+⊓-mono-< = ⊓-mono-≤
-≤⇒≤″ : _≤_ ⇒ _≤″_
-≤⇒≤″ m≤n = less-than-or-equal (proof m≤n)
- where
- k : ∀ m n → m ≤ n → ℕ
- k zero n _ = n
- k (suc m) zero ()
- k (suc m) (suc n) m≤n = k m n (≤-pred m≤n)
+-- Properties of _⊔_ and _⊓_ and _+_
+m⊔n≤m+n : ∀ m n → m ⊔ n ≤ m + n
+m⊔n≤m+n m n with ⊔-sel m n
+... | inj₁ m⊔n≡m rewrite m⊔n≡m = m≤m+n m n
+... | inj₂ m⊔n≡n rewrite m⊔n≡n = n≤m+n m n
- proof : ∀ {m n} (m≤n : m ≤ n) → m + k m n m≤n ≡ n
- proof z≤n = refl
- proof (s≤s m≤n) = cong suc (proof m≤n)
+m⊓n≤m+n : ∀ m n → m ⊓ n ≤ m + n
+m⊓n≤m+n m n with ⊓-sel m n
+... | inj₁ m⊓n≡m rewrite m⊓n≡m = m≤m+n m n
+... | inj₂ m⊓n≡n rewrite m⊓n≡n = n≤m+n m n
-------------------------------------------------------------------------
--- (ℕ, _≡_, _<_) is a strict total order
++-distribˡ-⊔ : _+_ DistributesOverˡ _⊔_
++-distribˡ-⊔ zero y z = refl
++-distribˡ-⊔ (suc x) y z = cong suc (+-distribˡ-⊔ x y z)
-m≢1+m+n : ∀ m {n} → m ≢ suc (m + n)
-m≢1+m+n zero ()
-m≢1+m+n (suc m) eq = m≢1+m+n m (cong pred eq)
++-distribʳ-⊔ : _+_ DistributesOverʳ _⊔_
++-distribʳ-⊔ = comm+distrˡ⇒distrʳ (cong₂ _⊔_) +-comm +-distribˡ-⊔
-strictTotalOrder : StrictTotalOrder _ _ _
-strictTotalOrder = record
- { Carrier = ℕ
- ; _≈_ = _≡_
- ; _<_ = _<_
- ; isStrictTotalOrder = record
- { isEquivalence = PropEq.isEquivalence
- ; trans = <-trans
- ; compare = cmp
- }
- }
- where
- 2+m+n≰m : ∀ {m n} → ¬ 2 + (m + n) ≤ m
- 2+m+n≰m (s≤s le) = 2+m+n≰m le
++-distrib-⊔ : _+_ DistributesOver _⊔_
++-distrib-⊔ = +-distribˡ-⊔ , +-distribʳ-⊔
+
++-distribˡ-⊓ : _+_ DistributesOverˡ _⊓_
++-distribˡ-⊓ zero y z = refl
++-distribˡ-⊓ (suc x) y z = cong suc (+-distribˡ-⊓ x y z)
+
++-distribʳ-⊓ : _+_ DistributesOverʳ _⊓_
++-distribʳ-⊓ = comm+distrˡ⇒distrʳ (cong₂ _⊓_) +-comm +-distribˡ-⊓
- cmp : Trichotomous _≡_ _<_
- cmp m n with compare m n
- cmp .m .(suc (m + k)) | less m k = tri< (m≤m+n (suc m) k) (m≢1+m+n _) 2+m+n≰m
- cmp .n .n | equal n = tri≈ 1+n≰n refl 1+n≰n
- cmp .(suc (n + k)) .n | greater n k = tri> 2+m+n≰m (m≢1+m+n _ ∘ sym) (m≤m+n (suc n) k)
++-distrib-⊓ : _+_ DistributesOver _⊓_
++-distrib-⊓ = +-distribˡ-⊓ , +-distribʳ-⊓
------------------------------------------------------------------------
--- Miscellaneous other properties
+-- Properties of _∸_
0∸n≡0 : LeftZero zero _∸_
0∸n≡0 zero = refl
@@ -414,8 +801,14 @@ n∸n≡0 : ∀ n → n ∸ n ≡ 0
n∸n≡0 zero = refl
n∸n≡0 (suc n) = n∸n≡0 n
++-∸-comm : ∀ {m} n {o} → o ≤ m → (m + n) ∸ o ≡ (m ∸ o) + n
++-∸-comm {zero} _ {suc o} ()
++-∸-comm {zero} _ {zero} _ = refl
++-∸-comm {suc m} _ {zero} _ = refl
++-∸-comm {suc m} n {suc o} (s≤s o≤m) = +-∸-comm n o≤m
+
∸-+-assoc : ∀ m n o → (m ∸ n) ∸ o ≡ m ∸ (n + o)
-∸-+-assoc m n zero = cong (_∸_ m) (sym $ +-right-identity n)
+∸-+-assoc m n zero = cong (m ∸_) (sym $ +-identityʳ n)
∸-+-assoc zero zero (suc o) = refl
∸-+-assoc zero (suc n) (suc o) = refl
∸-+-assoc (suc m) zero (suc o) = refl
@@ -424,22 +817,32 @@ n∸n≡0 (suc n) = n∸n≡0 n
+-∸-assoc : ∀ m {n o} → o ≤ n → (m + n) ∸ o ≡ m + (n ∸ o)
+-∸-assoc m (z≤n {n = n}) = begin m + n ∎
+-∸-assoc m (s≤s {m = o} {n = n} o≤n) = begin
- (m + suc n) ∸ suc o ≡⟨ cong (λ n → n ∸ suc o) (+-suc m n) ⟩
- suc (m + n) ∸ suc o ≡⟨ refl ⟩
+ (m + suc n) ∸ suc o ≡⟨ cong (_∸ suc o) (+-suc m n) ⟩
+ suc (m + n) ∸ suc o ≡⟨⟩
(m + n) ∸ o ≡⟨ +-∸-assoc m o≤n ⟩
m + (n ∸ o) ∎
+n∸m≤n : ∀ m n → n ∸ m ≤ n
+n∸m≤n zero n = ≤-refl
+n∸m≤n (suc m) zero = ≤-refl
+n∸m≤n (suc m) (suc n) = ≤-trans (n∸m≤n m n) (n≤1+n n)
+
+n≤m+n∸m : ∀ m n → n ≤ m + (n ∸ m)
+n≤m+n∸m m zero = z≤n
+n≤m+n∸m zero (suc n) = ≤-refl
+n≤m+n∸m (suc m) (suc n) = s≤s (n≤m+n∸m m n)
+
m+n∸n≡m : ∀ m n → (m + n) ∸ n ≡ m
m+n∸n≡m m n = begin
(m + n) ∸ n ≡⟨ +-∸-assoc m (≤-refl {x = n}) ⟩
- m + (n ∸ n) ≡⟨ cong (_+_ m) (n∸n≡0 n) ⟩
- m + 0 ≡⟨ +-right-identity m ⟩
+ m + (n ∸ n) ≡⟨ cong (m +_) (n∸n≡0 n) ⟩
+ m + 0 ≡⟨ +-identityʳ m ⟩
m ∎
m+n∸m≡n : ∀ {m n} → m ≤ n → m + (n ∸ m) ≡ n
m+n∸m≡n {m} {n} m≤n = begin
m + (n ∸ m) ≡⟨ sym $ +-∸-assoc m m≤n ⟩
- (m + n) ∸ m ≡⟨ cong (λ n → n ∸ m) (+-comm m n) ⟩
+ (m + n) ∸ m ≡⟨ cong (_∸ m) (+-comm m n) ⟩
(n + m) ∸ m ≡⟨ m+n∸n≡m n m ⟩
n ∎
@@ -460,118 +863,62 @@ m⊓n+n∸m≡n (suc m) (suc n) = cong suc $ m⊓n+n∸m≡n m n
-- TODO: Can this proof be simplified? An automatic solver which can
-- handle ∸ would be nice...
-
i∸k∸j+j∸k≡i+j∸k : ∀ i j k → i ∸ (k ∸ j) + (j ∸ k) ≡ i + j ∸ k
i∸k∸j+j∸k≡i+j∸k zero j k = begin
- 0 ∸ (k ∸ j) + (j ∸ k)
- ≡⟨ cong (λ x → x + (j ∸ k)) (0∸n≡0 (k ∸ j)) ⟩
- 0 + (j ∸ k)
- ≡⟨ refl ⟩
- j ∸ k
- ∎
+ 0 ∸ (k ∸ j) + (j ∸ k) ≡⟨ cong (_+ (j ∸ k)) (0∸n≡0 (k ∸ j)) ⟩
+ 0 + (j ∸ k) ≡⟨⟩
+ j ∸ k ∎
i∸k∸j+j∸k≡i+j∸k (suc i) j zero = begin
- suc i ∸ (0 ∸ j) + j
- ≡⟨ cong (λ x → suc i ∸ x + j) (0∸n≡0 j) ⟩
- suc i ∸ 0 + j
- ≡⟨ refl ⟩
- suc (i + j)
- ∎
+ suc i ∸ (0 ∸ j) + j ≡⟨ cong (λ x → suc i ∸ x + j) (0∸n≡0 j) ⟩
+ suc i ∸ 0 + j ≡⟨⟩
+ suc (i + j) ∎
i∸k∸j+j∸k≡i+j∸k (suc i) zero (suc k) = begin
- i ∸ k + 0
- ≡⟨ +-right-identity _ ⟩
- i ∸ k
- ≡⟨ cong (λ x → x ∸ k) (sym (+-right-identity _)) ⟩
- i + 0 ∸ k
- ∎
+ i ∸ k + 0 ≡⟨ +-identityʳ _ ⟩
+ i ∸ k ≡⟨ cong (_∸ k) (sym (+-identityʳ _)) ⟩
+ i + 0 ∸ k ∎
i∸k∸j+j∸k≡i+j∸k (suc i) (suc j) (suc k) = begin
- suc i ∸ (k ∸ j) + (j ∸ k)
- ≡⟨ i∸k∸j+j∸k≡i+j∸k (suc i) j k ⟩
- suc i + j ∸ k
- ≡⟨ cong (λ x → x ∸ k)
- (sym (+-suc i j)) ⟩
- i + suc j ∸ k
- ∎
-
-i+j≡0⇒i≡0 : ∀ i {j} → i + j ≡ 0 → i ≡ 0
-i+j≡0⇒i≡0 zero eq = refl
-i+j≡0⇒i≡0 (suc i) ()
-
-i+j≡0⇒j≡0 : ∀ i {j} → i + j ≡ 0 → j ≡ 0
-i+j≡0⇒j≡0 i {j} i+j≡0 = i+j≡0⇒i≡0 j $ begin
- j + i
- ≡⟨ +-comm j i ⟩
- i + j
- ≡⟨ i+j≡0 ⟩
- 0
- ∎
-
-i*j≡0⇒i≡0∨j≡0 : ∀ i {j} → i * j ≡ 0 → i ≡ 0 ⊎ j ≡ 0
-i*j≡0⇒i≡0∨j≡0 zero {j} eq = inj₁ refl
-i*j≡0⇒i≡0∨j≡0 (suc i) {zero} eq = inj₂ refl
-i*j≡0⇒i≡0∨j≡0 (suc i) {suc j} ()
+ suc i ∸ (k ∸ j) + (j ∸ k) ≡⟨ i∸k∸j+j∸k≡i+j∸k (suc i) j k ⟩
+ suc i + j ∸ k ≡⟨ cong (_∸ k) (sym (+-suc i j)) ⟩
+ i + suc j ∸ k ∎
-i*j≡1⇒i≡1 : ∀ i j → i * j ≡ 1 → i ≡ 1
-i*j≡1⇒i≡1 (suc zero) j _ = refl
-i*j≡1⇒i≡1 zero j ()
-i*j≡1⇒i≡1 (suc (suc i)) (suc (suc j)) ()
-i*j≡1⇒i≡1 (suc (suc i)) (suc zero) ()
-i*j≡1⇒i≡1 (suc (suc i)) zero eq with begin
- 0 ≡⟨ *-comm 0 i ⟩
- i * 0 ≡⟨ eq ⟩
- 1 ∎
-... | ()
-
-i*j≡1⇒j≡1 : ∀ i j → i * j ≡ 1 → j ≡ 1
-i*j≡1⇒j≡1 i j eq = i*j≡1⇒i≡1 j i (begin
- j * i ≡⟨ *-comm j i ⟩
- i * j ≡⟨ eq ⟩
- 1 ∎)
-
-cancel-+-left : ∀ i {j k} → i + j ≡ i + k → j ≡ k
-cancel-+-left zero eq = eq
-cancel-+-left (suc i) eq = cancel-+-left i (cong pred eq)
-
-cancel-+-left-≤ : ∀ i {j k} → i + j ≤ i + k → j ≤ k
-cancel-+-left-≤ zero le = le
-cancel-+-left-≤ (suc i) (s≤s le) = cancel-+-left-≤ i le
-
-cancel-*-right : ∀ i j {k} → i * suc k ≡ j * suc k → i ≡ j
-cancel-*-right zero zero eq = refl
-cancel-*-right zero (suc j) ()
-cancel-*-right (suc i) zero ()
-cancel-*-right (suc i) (suc j) {k} eq =
- cong suc (cancel-*-right i j (cancel-+-left (suc k) eq))
-
-cancel-*-right-≤ : ∀ i j k → i * suc k ≤ j * suc k → i ≤ j
-cancel-*-right-≤ zero _ _ _ = z≤n
-cancel-*-right-≤ (suc i) zero _ ()
-cancel-*-right-≤ (suc i) (suc j) k le =
- s≤s (cancel-*-right-≤ i j k (cancel-+-left-≤ (suc k) le))
-
-*-distrib-∸ʳ : _*_ DistributesOverʳ _∸_
-*-distrib-∸ʳ i zero k = begin
- (0 ∸ k) * i ≡⟨ cong₂ _*_ (0∸n≡0 k) refl ⟩
+*-distribʳ-∸ : _*_ DistributesOverʳ _∸_
+*-distribʳ-∸ i zero k = begin
+ (0 ∸ k) * i ≡⟨ cong (_* i) (0∸n≡0 k) ⟩
0 ≡⟨ sym $ 0∸n≡0 (k * i) ⟩
- 0 ∸ k * i ∎
-*-distrib-∸ʳ i (suc j) zero = begin i + j * i ∎
-*-distrib-∸ʳ i (suc j) (suc k) = begin
- (j ∸ k) * i ≡⟨ *-distrib-∸ʳ i j k ⟩
+ 0 ∸ (k * i) ∎
+*-distribʳ-∸ i (suc j) zero = refl
+*-distribʳ-∸ i (suc j) (suc k) = begin
+ (j ∸ k) * i ≡⟨ *-distribʳ-∸ i j k ⟩
j * i ∸ k * i ≡⟨ sym $ [i+j]∸[i+k]≡j∸k i _ _ ⟩
i + j * i ∸ (i + k * i) ∎
-im≡jm+n⇒[i∸j]m≡n
- : ∀ i j m n →
- i * m ≡ j * m + n → (i ∸ j) * m ≡ n
+∸-distribʳ-⊓ : _∸_ DistributesOverʳ _⊓_
+∸-distribʳ-⊓ zero y z = refl
+∸-distribʳ-⊓ (suc x) zero z = refl
+∸-distribʳ-⊓ (suc x) (suc y) zero = sym (⊓-zeroʳ (y ∸ x))
+∸-distribʳ-⊓ (suc x) (suc y) (suc z) = ∸-distribʳ-⊓ x y z
+
+∸-distribʳ-⊔ : _∸_ DistributesOverʳ _⊔_
+∸-distribʳ-⊔ zero y z = refl
+∸-distribʳ-⊔ (suc x) zero z = refl
+∸-distribʳ-⊔ (suc x) (suc y) zero = sym (⊔-identityʳ (y ∸ x))
+∸-distribʳ-⊔ (suc x) (suc y) (suc z) = ∸-distribʳ-⊔ x y z
+
+im≡jm+n⇒[i∸j]m≡n : ∀ i j m n → i * m ≡ j * m + n → (i ∸ j) * m ≡ n
im≡jm+n⇒[i∸j]m≡n i j m n eq = begin
- (i ∸ j) * m ≡⟨ *-distrib-∸ʳ m i j ⟩
- (i * m) ∸ (j * m) ≡⟨ cong₂ _∸_ eq (refl {x = j * m}) ⟩
- (j * m + n) ∸ (j * m) ≡⟨ cong₂ _∸_ (+-comm (j * m) n) (refl {x = j * m}) ⟩
+ (i ∸ j) * m ≡⟨ *-distribʳ-∸ m i j ⟩
+ (i * m) ∸ (j * m) ≡⟨ cong (_∸ j * m) eq ⟩
+ (j * m + n) ∸ (j * m) ≡⟨ cong (_∸ j * m) (+-comm (j * m) n) ⟩
(n + j * m) ∸ (j * m) ≡⟨ m+n∸n≡m n (j * m) ⟩
n ∎
-i+1+j≢i : ∀ i {j} → i + suc j ≢ i
-i+1+j≢i i eq = ¬i+1+j≤i i (reflexive eq)
- where open DecTotalOrder decTotalOrder
+∸-mono : _∸_ Preserves₂ _≤_ ⟶ _≥_ ⟶ _≤_
+∸-mono z≤n (s≤s n₁≥n₂) = z≤n
+∸-mono (s≤s m₁≤m₂) (s≤s n₁≥n₂) = ∸-mono m₁≤m₂ n₁≥n₂
+∸-mono m₁≤m₂ (z≤n {n = n₁}) = ≤-trans (n∸m≤n n₁ _) m₁≤m₂
+
+------------------------------------------------------------------------
+-- Properties of ⌊_/2⌋
⌊n/2⌋-mono : ⌊_/2⌋ Preserves _≤_ ⟶ _≤_
⌊n/2⌋-mono z≤n = z≤n
@@ -581,25 +928,56 @@ i+1+j≢i i eq = ¬i+1+j≤i i (reflexive eq)
⌈n/2⌉-mono : ⌈_/2⌉ Preserves _≤_ ⟶ _≤_
⌈n/2⌉-mono m≤n = ⌊n/2⌋-mono (s≤s m≤n)
-pred-mono : pred Preserves _≤_ ⟶ _≤_
-pred-mono z≤n = z≤n
-pred-mono (s≤s le) = le
+⌈n/2⌉≤′n : ∀ n → ⌈ n /2⌉ ≤′ n
+⌈n/2⌉≤′n zero = ≤′-refl
+⌈n/2⌉≤′n (suc zero) = ≤′-refl
+⌈n/2⌉≤′n (suc (suc n)) = s≤′s (≤′-step (⌈n/2⌉≤′n n))
-_+-mono_ : _+_ Preserves₂ _≤_ ⟶ _≤_ ⟶ _≤_
-_+-mono_ {zero} {m₂} {n₁} {n₂} z≤n n₁≤n₂ = start
- n₁ ≤⟨ n₁≤n₂ ⟩
- n₂ ≤⟨ n≤m+n m₂ n₂ ⟩
- m₂ + n₂ □
-s≤s m₁≤m₂ +-mono n₁≤n₂ = s≤s (m₁≤m₂ +-mono n₁≤n₂)
+⌊n/2⌋≤′n : ∀ n → ⌊ n /2⌋ ≤′ n
+⌊n/2⌋≤′n zero = ≤′-refl
+⌊n/2⌋≤′n (suc n) = ≤′-step (⌈n/2⌉≤′n n)
-_*-mono_ : _*_ Preserves₂ _≤_ ⟶ _≤_ ⟶ _≤_
-z≤n *-mono n₁≤n₂ = z≤n
-s≤s m₁≤m₂ *-mono n₁≤n₂ = n₁≤n₂ +-mono (m₁≤m₂ *-mono n₁≤n₂)
+------------------------------------------------------------------------
+-- Modules for reasoning about natural number relations
-∸-mono : _∸_ Preserves₂ _≤_ ⟶ _≥_ ⟶ _≤_
-∸-mono z≤n (s≤s n₁≥n₂) = z≤n
-∸-mono (s≤s m₁≤m₂) (s≤s n₁≥n₂) = ∸-mono m₁≤m₂ n₁≥n₂
-∸-mono {m₁} {m₂} m₁≤m₂ (z≤n {n = n₁}) = start
- m₁ ∸ n₁ ≤⟨ n∸m≤n n₁ m₁ ⟩
- m₁ ≤⟨ m₁≤m₂ ⟩
- m₂ □
+-- A module for automatically solving propositional equivalences
+module SemiringSolver =
+ Solver (ACR.fromCommutativeSemiring *-+-commutativeSemiring) _≟_
+
+-- A module for reasoning about the _≤_ relation
+module ≤-Reasoning where
+ open import Relation.Binary.PartialOrderReasoning
+ (DecTotalOrder.poset ≤-decTotalOrder) public
+ renaming (_≈⟨_⟩_ to _≡⟨_⟩_)
+
+ infixr 2 _<⟨_⟩_
+
+ _<⟨_⟩_ : ∀ x {y z} → x < y → y IsRelatedTo z → suc x IsRelatedTo z
+ x <⟨ x<y ⟩ y≤z = suc x ≤⟨ x<y ⟩ y≤z
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+_*-mono_ = *-mono-≤
+_+-mono_ = +-mono-≤
+
++-right-identity = +-identityʳ
+*-right-zero = *-zeroʳ
+distribʳ-*-+ = *-distribʳ-+
+*-distrib-∸ʳ = *-distribʳ-∸
+cancel-+-left = +-cancelˡ-≡
+cancel-+-left-≤ = +-cancelˡ-≤
+cancel-*-right = *-cancelʳ-≡
+cancel-*-right-≤ = *-cancelʳ-≤
+
+strictTotalOrder = <-strictTotalOrder
+isCommutativeSemiring = *-+-isCommutativeSemiring
+commutativeSemiring = *-+-commutativeSemiring
+isDistributiveLattice = ⊓-⊔-isDistributiveLattice
+distributiveLattice = ⊓-⊔-distributiveLattice
+⊔-⊓-0-isSemiringWithoutOne = ⊔-⊓-isSemiringWithoutOne
+⊔-⊓-0-isCommutativeSemiringWithoutOne = ⊔-⊓-isCommutativeSemiringWithoutOne
+⊔-⊓-0-commutativeSemiringWithoutOne = ⊔-⊓-commutativeSemiringWithoutOne
diff --git a/src/Data/Nat/Properties/Simple.agda b/src/Data/Nat/Properties/Simple.agda
index fe94081..c2e0992 100644
--- a/src/Data/Nat/Properties/Simple.agda
+++ b/src/Data/Nat/Properties/Simple.agda
@@ -2,109 +2,23 @@
-- The Agda standard library
--
-- A bunch of properties about natural number operations
+--
+-- This module is DEPRECATED. Please use Data.Nat.Properties directly.
------------------------------------------------------------------------
module Data.Nat.Properties.Simple where
-open import Data.Nat.Base as Nat
-open import Function
-open import Relation.Binary.PropositionalEquality as PropEq
- using (_≡_; _≢_; refl; sym; cong; cong₂)
-open PropEq.≡-Reasoning
-open import Data.Product
-open import Data.Sum
-
-------------------------------------------------------------------------
-
-+-assoc : ∀ m n o → (m + n) + o ≡ m + (n + o)
-+-assoc zero _ _ = refl
-+-assoc (suc m) n o = cong suc $ +-assoc m n o
-
-+-right-identity : ∀ n → n + 0 ≡ n
-+-right-identity zero = refl
-+-right-identity (suc n) = cong suc $ +-right-identity n
-
-+-suc : ∀ m n → m + suc n ≡ suc (m + n)
-+-suc zero n = refl
-+-suc (suc m) n = cong suc (+-suc m n)
-
-+-comm : ∀ m n → m + n ≡ n + m
-+-comm zero n = sym $ +-right-identity n
-+-comm (suc m) n =
- begin
- suc m + n
- ≡⟨ refl ⟩
- suc (m + n)
- ≡⟨ cong suc (+-comm m n) ⟩
- suc (n + m)
- ≡⟨ sym (+-suc n m) ⟩
- n + suc m
- ∎
-
-+-*-suc : ∀ m n → m * suc n ≡ m + m * n
-+-*-suc zero n = refl
-+-*-suc (suc m) n =
- begin
- suc m * suc n
- ≡⟨ refl ⟩
- suc n + m * suc n
- ≡⟨ cong (λ x → suc n + x) (+-*-suc m n) ⟩
- suc n + (m + m * n)
- ≡⟨ refl ⟩
- suc (n + (m + m * n))
- ≡⟨ cong suc (sym $ +-assoc n m (m * n)) ⟩
- suc (n + m + m * n)
- ≡⟨ cong (λ x → suc (x + m * n)) (+-comm n m) ⟩
- suc (m + n + m * n)
- ≡⟨ cong suc (+-assoc m n (m * n)) ⟩
- suc (m + (n + m * n))
- ≡⟨ refl ⟩
- suc m + suc m * n
- ∎
-
-*-right-zero : ∀ n → n * 0 ≡ 0
-*-right-zero zero = refl
-*-right-zero (suc n) = *-right-zero n
-
-*-comm : ∀ m n → m * n ≡ n * m
-*-comm zero n = sym $ *-right-zero n
-*-comm (suc m) n =
- begin
- suc m * n
- ≡⟨ refl ⟩
- n + m * n
- ≡⟨ cong (λ x → n + x) (*-comm m n) ⟩
- n + n * m
- ≡⟨ sym (+-*-suc n m) ⟩
- n * suc m
- ∎
-
-distribʳ-*-+ : ∀ m n o → (n + o) * m ≡ n * m + o * m
-distribʳ-*-+ m zero o = refl
-distribʳ-*-+ m (suc n) o =
- begin
- (suc n + o) * m
- ≡⟨ refl ⟩
- m + (n + o) * m
- ≡⟨ cong (_+_ m) $ distribʳ-*-+ m n o ⟩
- m + (n * m + o * m)
- ≡⟨ sym $ +-assoc m (n * m) (o * m) ⟩
- m + n * m + o * m
- ≡⟨ refl ⟩
- suc n * m + o * m
- ∎
-
-*-assoc : ∀ m n o → (m * n) * o ≡ m * (n * o)
-*-assoc zero n o = refl
-*-assoc (suc m) n o =
- begin
- (suc m * n) * o
- ≡⟨ refl ⟩
- (n + m * n) * o
- ≡⟨ distribʳ-*-+ o n (m * n) ⟩
- n * o + (m * n) * o
- ≡⟨ cong (λ x → n * o + x) $ *-assoc m n o ⟩
- n * o + m * (n * o)
- ≡⟨ refl ⟩
- suc m * (n * o)
- ∎
+open import Data.Nat.Properties public using
+ ( +-assoc
+ ; +-suc
+ ; +-comm
+ ; +-*-suc
+ ; *-comm
+ ; *-assoc
+ ) renaming
+ ( +-identityʳ to +-right-identity
+ ; +-identityˡ to +-left-identity
+ ; *-zeroʳ to *-right-zero
+ ; *-zeroˡ to *-left-zero
+ ; *-distribʳ-+ to distribʳ-*-+
+ )
diff --git a/src/Data/Product.agda b/src/Data/Product.agda
index 365f2c6..bfc8939 100644
--- a/src/Data/Product.agda
+++ b/src/Data/Product.agda
@@ -38,9 +38,19 @@ syntax Σ-syntax A (λ x → B) = Σ[ x ∈ A ] B
∃ : ∀ {a b} {A : Set a} → (A → Set b) → Set (a ⊔ b)
∃ = Σ _
+∃-syntax : ∀ {a b} {A : Set a} → (A → Set b) → Set (a ⊔ b)
+∃-syntax = ∃
+
+syntax ∃-syntax (λ x → B) = ∃[ x ] B
+
∄ : ∀ {a b} {A : Set a} → (A → Set b) → Set (a ⊔ b)
∄ P = ¬ ∃ P
+∄-syntax : ∀ {a b} {A : Set a} → (A → Set b) → Set (a ⊔ b)
+∄-syntax = ∄
+
+syntax ∄-syntax (λ x → B) = ∄[ x ] B
+
∃₂ : ∀ {a b c} {A : Set a} {B : A → Set b}
(C : (x : A) → B x → Set c) → Set (a ⊔ b ⊔ c)
∃₂ C = ∃ λ a → ∃ λ b → C a b
diff --git a/src/Data/Product/N-ary.agda b/src/Data/Product/N-ary.agda
index b8b4e28..1bec530 100644
--- a/src/Data/Product/N-ary.agda
+++ b/src/Data/Product/N-ary.agda
@@ -12,7 +12,7 @@
module Data.Product.N-ary where
-open import Data.Nat
+open import Data.Nat hiding (_^_)
open import Data.Product
open import Data.Unit
open import Data.Vec
diff --git a/src/Data/Rational.agda b/src/Data/Rational.agda
index 66be57e..535d6a0 100644
--- a/src/Data/Rational.agda
+++ b/src/Data/Rational.agda
@@ -54,8 +54,11 @@ _÷_ : (numerator : ℤ) (denominator : ℕ)
{≢0 : False (ℕ._≟_ denominator 0)} →
(n ÷ zero) {≢0 = ()}
-(n ÷ suc d) {c} =
- record { numerator = n; denominator-1 = d; isCoprime = c }
+(n ÷ suc d) {c} = record
+ { numerator = n
+ ; denominator-1 = d
+ ; isCoprime = c
+ }
private
@@ -151,67 +154,3 @@ p ≤? q with ℚ.numerator p ℤ.* ℚ.denominator q ℤ.≤?
ℚ.numerator q ℤ.* ℚ.denominator p
p ≤? q | yes pq≤qp = yes (*≤* pq≤qp)
p ≤? q | no ¬pq≤qp = no (λ { (*≤* pq≤qp) → ¬pq≤qp pq≤qp })
-
-decTotalOrder : DecTotalOrder _ _ _
-decTotalOrder = record
- { Carrier = ℚ
- ; _≈_ = _≡_
- ; _≤_ = _≤_
- ; isDecTotalOrder = record
- { isTotalOrder = record
- { isPartialOrder = record
- { isPreorder = record
- { isEquivalence = P.isEquivalence
- ; reflexive = refl′
- ; trans = trans
- }
- ; antisym = antisym
- }
- ; total = total
- }
- ; _≟_ = _≟_
- ; _≤?_ = _≤?_
- }
- }
- where
- module ℤO = DecTotalOrder ℤ.decTotalOrder
-
- refl′ : _≡_ ⇒ _≤_
- refl′ refl = *≤* ℤO.refl
-
- trans : Transitive _≤_
- trans {i = p} {j = q} {k = r} (*≤* le₁) (*≤* le₂)
- = *≤* (ℤ.cancel-*-+-right-≤ _ _ _
- (lemma
- (ℚ.numerator p) (ℚ.denominator p)
- (ℚ.numerator q) (ℚ.denominator q)
- (ℚ.numerator r) (ℚ.denominator r)
- (ℤ.*-+-right-mono (ℚ.denominator-1 r) le₁)
- (ℤ.*-+-right-mono (ℚ.denominator-1 p) le₂)))
- where
- open Algebra.CommutativeRing ℤ.commutativeRing
-
- lemma : ∀ n₁ d₁ n₂ d₂ n₃ d₃ →
- n₁ ℤ.* d₂ ℤ.* d₃ ℤ.≤ n₂ ℤ.* d₁ ℤ.* d₃ →
- n₂ ℤ.* d₃ ℤ.* d₁ ℤ.≤ n₃ ℤ.* d₂ ℤ.* d₁ →
- n₁ ℤ.* d₃ ℤ.* d₂ ℤ.≤ n₃ ℤ.* d₁ ℤ.* d₂
- lemma n₁ d₁ n₂ d₂ n₃ d₃
- rewrite *-assoc n₁ d₂ d₃
- | *-comm d₂ d₃
- | sym (*-assoc n₁ d₃ d₂)
- | *-assoc n₃ d₂ d₁
- | *-comm d₂ d₁
- | sym (*-assoc n₃ d₁ d₂)
- | *-assoc n₂ d₁ d₃
- | *-comm d₁ d₃
- | sym (*-assoc n₂ d₃ d₁)
- = ℤO.trans
-
- antisym : Antisymmetric _≡_ _≤_
- antisym (*≤* le₁) (*≤* le₂) = ≃⇒≡ (ℤO.antisym le₁ le₂)
-
- total : Total _≤_
- total p q =
- [ inj₁ ∘′ *≤* , inj₂ ∘′ *≤* ]′
- (ℤO.total (ℚ.numerator p ℤ.* ℚ.denominator q)
- (ℚ.numerator q ℤ.* ℚ.denominator p))
diff --git a/src/Data/Rational/Properties.agda b/src/Data/Rational/Properties.agda
new file mode 100644
index 0000000..1f377e9
--- /dev/null
+++ b/src/Data/Rational/Properties.agda
@@ -0,0 +1,97 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties of Rational numbers
+------------------------------------------------------------------------
+
+module Data.Rational.Properties where
+
+import Algebra
+open import Function using (_∘_)
+import Data.Integer as ℤ
+import Data.Integer.Properties as ℤₚ
+open import Data.Rational
+open import Data.Nat as ℕ using (ℕ; zero; suc)
+open import Data.Sum
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; refl; sym; cong; cong₂)
+
+------------------------------------------------------------------------
+-- _≤_
+
+≤-reflexive : _≡_ ⇒ _≤_
+≤-reflexive refl = *≤* ℤₚ.≤-refl
+
+≤-refl : Reflexive _≤_
+≤-refl = ≤-reflexive refl
+
+≤-trans : Transitive _≤_
+≤-trans {i = p} {j = q} {k = r} (*≤* le₁) (*≤* le₂)
+ = *≤* (ℤₚ.cancel-*-+-right-≤ _ _ _
+ (lemma
+ (ℚ.numerator p) (ℚ.denominator p)
+ (ℚ.numerator q) (ℚ.denominator q)
+ (ℚ.numerator r) (ℚ.denominator r)
+ (ℤₚ.*-+-right-mono (ℚ.denominator-1 r) le₁)
+ (ℤₚ.*-+-right-mono (ℚ.denominator-1 p) le₂)))
+ where
+ lemma : ∀ n₁ d₁ n₂ d₂ n₃ d₃ →
+ n₁ ℤ.* d₂ ℤ.* d₃ ℤ.≤ n₂ ℤ.* d₁ ℤ.* d₃ →
+ n₂ ℤ.* d₃ ℤ.* d₁ ℤ.≤ n₃ ℤ.* d₂ ℤ.* d₁ →
+ n₁ ℤ.* d₃ ℤ.* d₂ ℤ.≤ n₃ ℤ.* d₁ ℤ.* d₂
+ lemma n₁ d₁ n₂ d₂ n₃ d₃
+ rewrite ℤₚ.*-assoc n₁ d₂ d₃
+ | ℤₚ.*-comm d₂ d₃
+ | sym (ℤₚ.*-assoc n₁ d₃ d₂)
+ | ℤₚ.*-assoc n₃ d₂ d₁
+ | ℤₚ.*-comm d₂ d₁
+ | sym (ℤₚ.*-assoc n₃ d₁ d₂)
+ | ℤₚ.*-assoc n₂ d₁ d₃
+ | ℤₚ.*-comm d₁ d₃
+ | sym (ℤₚ.*-assoc n₂ d₃ d₁)
+ = ℤₚ.≤-trans
+
+≤-antisym : Antisymmetric _≡_ _≤_
+≤-antisym (*≤* le₁) (*≤* le₂) = ≃⇒≡ (ℤₚ.≤-antisym le₁ le₂)
+
+≤-total : Total _≤_
+≤-total p q =
+ [ inj₁ ∘ *≤* , inj₂ ∘ *≤* ]′
+ (ℤₚ.≤-total (ℚ.numerator p ℤ.* ℚ.denominator q)
+ (ℚ.numerator q ℤ.* ℚ.denominator p))
+
+≤-isPreorder : IsPreorder _≡_ _≤_
+≤-isPreorder = record
+ { isEquivalence = P.isEquivalence
+ ; reflexive = ≤-reflexive
+ ; trans = ≤-trans
+ }
+
+≤-isPartialOrder : IsPartialOrder _≡_ _≤_
+≤-isPartialOrder = record
+ { isPreorder = ≤-isPreorder
+ ; antisym = ≤-antisym
+ }
+
+≤-isTotalOrder : IsTotalOrder _≡_ _≤_
+≤-isTotalOrder = record
+ { isPartialOrder = ≤-isPartialOrder
+ ; total = ≤-total
+ }
+
+≤-isDecTotalOrder : IsDecTotalOrder _≡_ _≤_
+≤-isDecTotalOrder = record
+ { isTotalOrder = ≤-isTotalOrder
+ ; _≟_ = _≟_
+ ; _≤?_ = _≤?_
+ }
+
+≤-decTotalOrder : DecTotalOrder _ _ _
+≤-decTotalOrder = record
+ { Carrier = ℚ
+ ; _≈_ = _≡_
+ ; _≤_ = _≤_
+ ; isDecTotalOrder = ≤-isDecTotalOrder
+ }
+
diff --git a/src/Data/Sign/Properties.agda b/src/Data/Sign/Properties.agda
index 55f6278..8ac8bc4 100644
--- a/src/Data/Sign/Properties.agda
+++ b/src/Data/Sign/Properties.agda
@@ -9,7 +9,9 @@ module Data.Sign.Properties where
open import Data.Empty
open import Function
open import Data.Sign
+open import Data.Product using (_,_)
open import Relation.Binary.PropositionalEquality
+open import Algebra.FunctionProperties (_≡_ {A = Sign})
-- The opposite of a sign is not equal to the sign.
@@ -17,10 +19,51 @@ opposite-not-equal : ∀ s → s ≢ opposite s
opposite-not-equal - ()
opposite-not-equal + ()
--- Sign multiplication is right cancellative.
+opposite-cong : ∀ {s t} → opposite s ≡ opposite t → s ≡ t
+opposite-cong { - } { - } refl = refl
+opposite-cong { - } { + } ()
+opposite-cong { + } { - } ()
+opposite-cong { + } { + } refl = refl
-cancel-*-right : ∀ s₁ s₂ {s} → s₁ * s ≡ s₂ * s → s₁ ≡ s₂
+------------------------------------------------------------------------
+-- _*_
+
+*-identityˡ : LeftIdentity + _*_
+*-identityˡ _ = refl
+
+*-identityʳ : RightIdentity + _*_
+*-identityʳ - = refl
+*-identityʳ + = refl
+
+*-identity : Identity + _*_
+*-identity = *-identityˡ , *-identityʳ
+
+*-comm : Commutative _*_
+*-comm + + = refl
+*-comm + - = refl
+*-comm - + = refl
+*-comm - - = refl
+
+*-assoc : Associative _*_
+*-assoc + + _ = refl
+*-assoc + - _ = refl
+*-assoc - + _ = refl
+*-assoc - - + = refl
+*-assoc - - - = refl
+
+cancel-*-right : RightCancellative _*_
cancel-*-right - - _ = refl
cancel-*-right - + eq = ⊥-elim (opposite-not-equal _ $ sym eq)
cancel-*-right + - eq = ⊥-elim (opposite-not-equal _ eq)
cancel-*-right + + _ = refl
+
+cancel-*-left : LeftCancellative _*_
+cancel-*-left - eq = opposite-cong eq
+cancel-*-left + eq = eq
+
+*-cancellative : Cancellative _*_
+*-cancellative = cancel-*-left , cancel-*-right
+
+s*s≡+ : ∀ s → s * s ≡ +
+s*s≡+ + = refl
+s*s≡+ - = refl
diff --git a/src/Data/Stream.agda b/src/Data/Stream.agda
index 3231139..3ae4e28 100644
--- a/src/Data/Stream.agda
+++ b/src/Data/Stream.agda
@@ -21,8 +21,8 @@ infixr 5 _∷_
data Stream {a} (A : Set a) : Set a where
_∷_ : (x : A) (xs : ∞ (Stream A)) → Stream A
-{-# HASKELL data AgdaStream a = Cons a (AgdaStream a) #-}
-{-# COMPILED_DATA Stream MAlonzo.Code.Data.Stream.AgdaStream MAlonzo.Code.Data.Stream.Cons #-}
+{-# FOREIGN GHC data AgdaStream a = Cons a (AgdaStream a) #-}
+{-# COMPILE GHC Stream = data MAlonzo.Code.Data.Stream.AgdaStream (MAlonzo.Code.Data.Stream.Cons) #-}
------------------------------------------------------------------------
-- Some operations
diff --git a/src/Data/Sum.agda b/src/Data/Sum.agda
index 3942c7f..2ea1c07 100644
--- a/src/Data/Sum.agda
+++ b/src/Data/Sum.agda
@@ -7,6 +7,7 @@
module Data.Sum where
open import Function
+open import Data.Unit.Base using (⊤; tt)
open import Data.Maybe.Base using (Maybe; just; nothing)
open import Level
@@ -19,8 +20,8 @@ data _⊎_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where
inj₁ : (x : A) → A ⊎ B
inj₂ : (y : B) → A ⊎ B
-{-# HASKELL type AgdaEither a b c d = Either c d #-}
-{-# COMPILED_DATA _⊎_ MAlonzo.Code.Data.Sum.AgdaEither Left Right #-}
+{-# FOREIGN GHC type AgdaEither a b c d = Either c d #-}
+{-# COMPILE GHC _⊎_ = data MAlonzo.Code.Data.Sum.AgdaEither (Left | Right) #-}
------------------------------------------------------------------------
-- Functions
@@ -52,3 +53,19 @@ isInj₁ (inj₂ y) = nothing
isInj₂ : ∀ {a b} {A : Set a} {B : Set b} → A ⊎ B → Maybe B
isInj₂ (inj₁ x) = nothing
isInj₂ (inj₂ y) = just y
+
+From-inj₁ : ∀ {a b} {A : Set a} {B : Set b} → A ⊎ B → Set a
+From-inj₁ {A = A} (inj₁ _) = A
+From-inj₁ (inj₂ _) = Lift ⊤
+
+from-inj₁ : ∀ {a b} {A : Set a} {B : Set b} (x : A ⊎ B) → From-inj₁ x
+from-inj₁ (inj₁ x) = x
+from-inj₁ (inj₂ _) = lift tt
+
+From-inj₂ : ∀ {a b} {A : Set a} {B : Set b} → A ⊎ B → Set b
+From-inj₂ (inj₁ _) = Lift ⊤
+From-inj₂ {B = B} (inj₂ _) = B
+
+from-inj₂ : ∀ {a b} {A : Set a} {B : Set b} (x : A ⊎ B) → From-inj₂ x
+from-inj₂ (inj₁ _) = lift tt
+from-inj₂ (inj₂ x) = x
diff --git a/src/Data/Vec.agda b/src/Data/Vec.agda
index e86e113..7be843a 100644
--- a/src/Data/Vec.agda
+++ b/src/Data/Vec.agda
@@ -6,6 +6,7 @@
module Data.Vec where
+open import Category.Functor
open import Category.Applicative
open import Data.Nat
open import Data.Fin using (Fin; zero; suc)
@@ -74,11 +75,18 @@ applicative = record
map : ∀ {a b n} {A : Set a} {B : Set b} →
(A → B) → Vec A n → Vec B n
-map f xs = replicate f ⊛ xs
+map f [] = []
+map f (x ∷ xs) = f x ∷ map f xs
+
+functor : ∀ {a n} → RawFunctor (λ (A : Set a) → Vec A n)
+functor = record
+ { _<$>_ = map
+ }
zipWith : ∀ {a b c n} {A : Set a} {B : Set b} {C : Set c} →
(A → B → C) → Vec A n → Vec B n → Vec C n
-zipWith _⊕_ xs ys = replicate _⊕_ ⊛ xs ⊛ ys
+zipWith f [] [] = []
+zipWith f (x ∷ xs) (y ∷ ys) = f x y ∷ zipWith f xs ys
zip : ∀ {a b n} {A : Set a} {B : Set b} →
Vec A n → Vec B n → Vec (A × B) n
@@ -87,7 +95,7 @@ zip = zipWith _,_
unzip : ∀ {a b n} {A : Set a} {B : Set b} →
Vec (A × B) n → Vec A n × Vec B n
unzip [] = [] , []
-unzip ((x , y) ∷ xys) = Prod.map (_∷_ x) (_∷_ y) (unzip xys)
+unzip ((x , y) ∷ xys) = Prod.map (x ∷_) (y ∷_) (unzip xys)
foldr : ∀ {a b} {A : Set a} (B : ℕ → Set b) {m} →
(∀ {n} → A → B n → B (suc n)) →
diff --git a/src/Data/Vec/All.agda b/src/Data/Vec/All.agda
index 64f75ba..0dfe1b1 100644
--- a/src/Data/Vec/All.agda
+++ b/src/Data/Vec/All.agda
@@ -17,6 +17,7 @@ import Relation.Nullary.Decidable as Dec
open import Relation.Unary using (Decidable) renaming (_⊆_ to _⋐_)
open import Relation.Binary.PropositionalEquality using (subst)
+------------------------------------------------------------------------
-- All P xs means that all elements in xs satisfy P.
infixr 5 _∷_
@@ -67,23 +68,24 @@ zipWith _⊕_ {xs = x ∷ xs} {y ∷ ys} (px ∷ pxs) (qy ∷ qys) =
px ⊕ qy ∷ zipWith _⊕_ pxs qys
--- A shorthand for a pair of vectors being point-wise related.
-
-All₂ : ∀ {a p} {A B : Set a} (P : A → B → Set p) {k} →
- Vec A k → Vec B k → Set _
-All₂ P xs ys = All (uncurry P) (zip xs ys)
+------------------------------------------------------------------------
+-- All₂ P xs ys means that every pointwise pair in xs ys satisfy P.
--- A variant of lookup tailored to All₂.
+data All₂ {a b p} {A : Set a} {B : Set b} (P : A → B → Set p) :
+ ∀ {n} → Vec A n → Vec B n → Set (a ⊔ b ⊔ p) where
+ [] : All₂ P [] []
+ _∷_ : ∀ {n x y} {xs : Vec A n} {ys : Vec B n} →
+ P x y → All₂ P xs ys → All₂ P (x ∷ xs) (y ∷ ys)
-lookup₂ : ∀ {a p} {A B : Set a} {P : A → B → Set p} {k}
+lookup₂ : ∀ {a b p} {A : Set a} {B : Set b} {P : A → B → Set p} {k}
{xs : Vec A k} {ys : Vec B k} →
- (i : Fin k) → All₂ P xs ys → P (Vec.lookup i xs) (Vec.lookup i ys)
-lookup₂ {P = P} {xs = xs} {ys} i pxys =
- subst (uncurry P) (lookup-zip i xs ys) (lookup i pxys)
-
--- A variant of map tailored to All₂.
-
-map₂ : ∀ {a p q} {A B : Set a} {P : A → B → Set p} {Q : A → B → Set q} →
- (∀ {x y} → P x y → Q x y) →
- ∀ {k xs ys} → All₂ P {k} xs ys → All₂ Q {k} xs ys
-map₂ g = map g
+ ∀ i → All₂ P xs ys → P (Vec.lookup i xs) (Vec.lookup i ys)
+lookup₂ zero (pxy ∷ _) = pxy
+lookup₂ (suc i) (_ ∷ pxys) = lookup₂ i pxys
+
+map₂ : ∀ {a b p q} {A : Set a} {B : Set b}
+ {P : A → B → Set p} {Q : A → B → Set q} →
+ (∀ {x y} → P x y → Q x y) →
+ ∀ {k xs ys} → All₂ P {k} xs ys → All₂ Q {k} xs ys
+map₂ g [] = []
+map₂ g (pxy ∷ pxys) = g pxy ∷ map₂ g pxys
diff --git a/src/Data/Vec/All/Properties.agda b/src/Data/Vec/All/Properties.agda
index af013d5..3af19c7 100644
--- a/src/Data/Vec/All/Properties.agda
+++ b/src/Data/Vec/All/Properties.agda
@@ -6,14 +6,14 @@
module Data.Vec.All.Properties where
-open import Data.Vec as Vec using (Vec; []; _∷_; zip)
+open import Data.Vec as Vec using (Vec; []; _∷_; zip; concat; _++_)
open import Data.Vec.Properties using (map-id; lookup-zip)
open import Data.Product as Prod using (_×_; _,_; uncurry; uncurry′)
-open import Data.Vec using (Vec; _++_)
open import Data.Vec.All as All using (All; All₂; []; _∷_)
open import Function
open import Function.Inverse using (_↔_)
open import Relation.Unary using () renaming (_⊆_ to _⋐_)
+open import Relation.Binary using (REL; Trans)
open import Relation.Binary.PropositionalEquality as P using (_≡_)
-- Functions can be shifted between the predicate and the vector.
@@ -38,88 +38,162 @@ gmap : ∀ {a b p q}
P ⋐ Q ∘ f → All P {k} ⋐ All Q {k} ∘ Vec.map f
gmap g = All-map ∘ All.map g
--- A variant of All-map tailored to All₂.
+-- A variant of All-map for All₂.
-All₂-map : ∀ {a b p} {A₁ A₂ : Set a} {B₁ B₂ : Set b} {P : B₁ → B₂ → Set p}
- {f₁ : A₁ → B₁} {f₂ : A₂ → B₂} {k} {xs : Vec A₁ k} {ys : Vec A₂ k} →
+All₂-map : ∀ {a b c d p} {A : Set a} {B : Set b} {C : Set c} {D : Set d}
+ {P : C → D → Set p}
+ {f₁ : A → C} {f₂ : B → D} {k} {xs : Vec A k} {ys : Vec B k} →
All₂ (λ x y → P (f₁ x) (f₂ y)) xs ys →
All₂ P (Vec.map f₁ xs) (Vec.map f₂ ys)
-All₂-map {xs = []} {ys = []} [] = []
-All₂-map {xs = _ ∷ _} {ys = _ ∷ _} (px ∷ pxs) = px ∷ All₂-map pxs
+All₂-map [] = []
+All₂-map (px ∷ pxs) = px ∷ All₂-map pxs
--- A variant of gmap tailored to All₂.
+-- Abstract composition of binary relations lifted to All₂.
+
+comp : ∀ {a b c p q r} {A : Set a} {B : Set b} {C : Set c}
+ {P : A → B → Set p} {Q : B → C → Set q} {R : A → C → Set r} →
+ Trans P Q R → ∀ {k} → Trans (All₂ P {k}) (All₂ Q) (All₂ R)
+comp _⊙_ [] [] = []
+comp _⊙_ (pxy ∷ pxys) (qzx ∷ qzxs) = pxy ⊙ qzx ∷ comp _⊙_ pxys qzxs
-gmap₂ : ∀ {a b p q} {A₁ A₂ : Set a} {B₁ B₂ : Set b}
- {P : A₁ → A₂ → Set p} {Q : B₁ → B₂ → Set q}
- {f₁ : A₁ → B₁} {f₂ : A₂ → B₂} →
- (∀ {x y} → P x y → Q (f₁ x) (f₂ y)) → ∀ {k xs ys} →
- All₂ P {k} xs ys → All₂ Q {k} (Vec.map f₁ xs) (Vec.map f₂ ys)
-gmap₂ g = All₂-map ∘ All.map g
+------------------------------------------------------------------------
+-- Variants of gmap for All₂.
--- A variant of gmap₂ shifting only the first function from the binary
--- relation to the vector.
+module _ {a b c p q} {A : Set a} {B : Set b} {C : Set c} where
-gmap₂₁ : ∀ {a p q} {A₁ A₂ : Set a} {B : Set a}
- {P : A₁ → A₂ → Set p} {Q : B → A₂ → Set q} {f : A₁ → B} →
- (∀ {x y} → P x y → Q (f x) y) → ∀ {k xs ys} →
- All₂ P {k} xs ys → All₂ Q {k} (Vec.map f xs) ys
-gmap₂₁ g = P.subst (All₂ _ _) (map-id _) ∘ All₂-map {f₂ = id} ∘ All.map g
+ -- A variant of gmap₂ shifting two functions from the binary
+ -- relation to the vector.
--- A variant of gmap₂ shifting only the second function from the
--- binary relation to the vector.
+ gmap₂ : ∀ {d} {D : Set d} {P : A → B → Set p} {Q : C → D → Set q}
+ {f₁ : A → C} {f₂ : B → D} →
+ (∀ {x y} → P x y → Q (f₁ x) (f₂ y)) → ∀ {k xs ys} →
+ All₂ P {k} xs ys → All₂ Q {k} (Vec.map f₁ xs) (Vec.map f₂ ys)
+ gmap₂ g [] = []
+ gmap₂ g (pxy ∷ pxys) = g pxy ∷ gmap₂ g pxys
-gmap₂₂ : ∀ {a p q} {A₁ A₂ : Set a} {B : Set a}
- {P : A₁ → A₂ → Set p} {Q : A₁ → B → Set q} {f : A₂ → B} →
- (∀ {x y} → P x y → Q x (f y)) → ∀ {k xs ys} →
- All₂ P {k} xs ys → All₂ Q {k} xs (Vec.map f ys)
-gmap₂₂ g =
- P.subst (flip (All₂ _) _) (map-id _) ∘ All₂-map {f₁ = id} ∘ All.map g
+ -- A variant of gmap₂ shifting only the first function from the binary
+ -- relation to the vector.
+
+ gmap₂₁ : ∀ {P : A → B → Set p} {Q : C → B → Set q} {f : A → C} →
+ (∀ {x y} → P x y → Q (f x) y) → ∀ {k xs ys} →
+ All₂ P {k} xs ys → All₂ Q {k} (Vec.map f xs) ys
+ gmap₂₁ g [] = []
+ gmap₂₁ g (pxy ∷ pxys) = g pxy ∷ gmap₂₁ g pxys
+
+ -- A variant of gmap₂ shifting only the second function from the
+ -- binary relation to the vector.
+
+ gmap₂₂ : ∀ {P : A → B → Set p} {Q : A → C → Set q} {f : B → C} →
+ (∀ {x y} → P x y → Q x (f y)) → ∀ {k xs ys} →
+ All₂ P {k} xs ys → All₂ Q {k} xs (Vec.map f ys)
+ gmap₂₂ g [] = []
+ gmap₂₂ g (pxy ∷ pxys) = g pxy ∷ gmap₂₂ g pxys
--- Abstract composition of binary relations lifted to All₂.
-comp : ∀ {a p} {A : Set a} {B : Set a} {C : Set a}
- {P : A → B → Set p} {Q : B → C → Set p} {R : A → C → Set p} →
- (∀ {x y z} → P x y → Q y z → R x z) →
- ∀ {k xs ys zs} → All₂ P {k} xs ys → All₂ Q {k} ys zs → All₂ R {k} xs zs
-comp _⊙_ {xs = []} {[]} {[]} [] [] = []
-comp _⊙_ {xs = x ∷ xs} {y ∷ ys} {z ∷ zs} (pxy ∷ pxys) (qzx ∷ qzxs) =
- pxy ⊙ qzx ∷ comp _⊙_ pxys qzxs
+------------------------------------------------------------------------
+-- All and _++_
+
+module _ {a n p} {A : Set a} {P : A → Set p} where
--- All P (xs ++ ys) is isomorphic to All P xs and All P ys.
+ All-++⁺ : ∀ {m} {xs : Vec A m} {ys : Vec A n} →
+ All P xs → All P ys → All P (xs ++ ys)
+ All-++⁺ [] pys = pys
+ All-++⁺ (px ∷ pxs) pys = px ∷ All-++⁺ pxs pys
-private
+ All-++ˡ⁻ : ∀ {m} (xs : Vec A m) {ys : Vec A n} →
+ All P (xs ++ ys) → All P xs
+ All-++ˡ⁻ [] _ = []
+ All-++ˡ⁻ (x ∷ xs) (px ∷ pxs) = px ∷ All-++ˡ⁻ xs pxs
- ++⁺ : ∀ {a p} {A : Set a} {P : A → Set p} {k i}
- {xs : Vec A k} {ys : Vec A i} →
- All P xs → All P ys → All P (xs ++ ys)
- ++⁺ [] pys = pys
- ++⁺ (px ∷ pxs) pys = px ∷ ++⁺ pxs pys
+ All-++ʳ⁻ : ∀ {m} (xs : Vec A m) {ys : Vec A n} →
+ All P (xs ++ ys) → All P ys
+ All-++ʳ⁻ [] pys = pys
+ All-++ʳ⁻ (x ∷ xs) (px ∷ pxs) = All-++ʳ⁻ xs pxs
- ++⁻ : ∀ {a p} {A : Set a} {P : A → Set p} {k i}
- (xs : Vec A k) {ys : Vec A i} →
+ All-++⁻ : ∀ {m} (xs : Vec A m) {ys : Vec A n} →
All P (xs ++ ys) → All P xs × All P ys
- ++⁻ [] p = [] , p
- ++⁻ (x ∷ xs) (px ∷ pxs) = Prod.map (_∷_ px) id $ ++⁻ _ pxs
-
- ++⁺∘++⁻ : ∀ {a p} {A : Set a} {P : A → Set p} {k i}
- (xs : Vec A k) {ys : Vec A i} (p : All P (xs ++ ys)) →
- uncurry′ ++⁺ (++⁻ xs p) ≡ p
- ++⁺∘++⁻ [] p = P.refl
- ++⁺∘++⁻ (x ∷ xs) (px ∷ pxs) = P.cong (_∷_ px) $ ++⁺∘++⁻ xs pxs
-
- ++⁻∘++⁺ : ∀ {a p} {A : Set a} {P : A → Set p} {k i}
- {xs : Vec A k} {ys : Vec A i} (p : All P xs × All P ys) →
- ++⁻ xs (uncurry ++⁺ p) ≡ p
- ++⁻∘++⁺ ([] , pys) = P.refl
- ++⁻∘++⁺ (px ∷ pxs , pys) rewrite ++⁻∘++⁺ (pxs , pys) = P.refl
-
-++↔ : ∀ {a p} {A : Set a} {P : A → Set p} {k i} {xs : Vec A k} {ys : Vec A i} →
- (All P xs × All P ys) ↔ All P (xs ++ ys)
-++↔ {P = P} {xs = xs} = record
- { to = P.→-to-⟶ $ uncurry ++⁺
- ; from = P.→-to-⟶ $ ++⁻ xs
- ; inverse-of = record
- { left-inverse-of = ++⁻∘++⁺
- ; right-inverse-of = ++⁺∘++⁻ xs
+ All-++⁻ [] p = [] , p
+ All-++⁻ (x ∷ xs) (px ∷ pxs) = Prod.map (px ∷_) id (All-++⁻ _ pxs)
+
+ All-++⁺∘++⁻ : ∀ {m} (xs : Vec A m) {ys : Vec A n} →
+ (p : All P (xs ++ ys)) →
+ uncurry′ All-++⁺ (All-++⁻ xs p) ≡ p
+ All-++⁺∘++⁻ [] p = P.refl
+ All-++⁺∘++⁻ (x ∷ xs) (px ∷ pxs) = P.cong (px ∷_) (All-++⁺∘++⁻ xs pxs)
+
+ All-++⁻∘++⁺ : ∀ {m} {xs : Vec A m} {ys : Vec A n} →
+ (p : All P xs × All P ys) →
+ All-++⁻ xs (uncurry All-++⁺ p) ≡ p
+ All-++⁻∘++⁺ ([] , pys) = P.refl
+ All-++⁻∘++⁺ (px ∷ pxs , pys) rewrite All-++⁻∘++⁺ (pxs , pys) = P.refl
+
+ ++↔ : ∀ {m} {xs : Vec A m} {ys : Vec A n} →
+ (All P xs × All P ys) ↔ All P (xs ++ ys)
+ ++↔ {xs = xs} = record
+ { to = P.→-to-⟶ $ uncurry All-++⁺
+ ; from = P.→-to-⟶ $ All-++⁻ xs
+ ; inverse-of = record
+ { left-inverse-of = All-++⁻∘++⁺
+ ; right-inverse-of = All-++⁺∘++⁻ xs
+ }
}
- }
+
+------------------------------------------------------------------------
+-- All₂ and _++_
+
+module _ {a b n p} {A : Set a} {B : Set b} {_~_ : REL A B p} where
+
+ All₂-++⁺ : ∀ {m} {ws : Vec A m} {xs} {ys : Vec A n} {zs} →
+ All₂ _~_ ws xs → All₂ _~_ ys zs →
+ All₂ _~_ (ws ++ ys) (xs ++ zs)
+ All₂-++⁺ [] ys~zs = ys~zs
+ All₂-++⁺ (w~x ∷ ws~xs) ys~zs = w~x ∷ (All₂-++⁺ ws~xs ys~zs)
+
+ All₂-++ˡ⁻ : ∀ {m} (ws : Vec A m) xs {ys : Vec A n} {zs} →
+ All₂ _~_ (ws ++ ys) (xs ++ zs) → All₂ _~_ ws xs
+ All₂-++ˡ⁻ [] [] _ = []
+ All₂-++ˡ⁻ (w ∷ ws) (x ∷ xs) (w~x ∷ ps) = w~x ∷ All₂-++ˡ⁻ ws xs ps
+
+ All₂-++ʳ⁻ : ∀ {m} (ws : Vec A m) (xs : Vec B m) {ys : Vec A n} {zs} →
+ All₂ _~_ (ws ++ ys) (xs ++ zs) → All₂ _~_ ys zs
+ All₂-++ʳ⁻ [] [] ys~zs = ys~zs
+ All₂-++ʳ⁻ (w ∷ ws) (x ∷ xs) (_ ∷ ps) = All₂-++ʳ⁻ ws xs ps
+
+ All₂-++⁻ : ∀ {m} (ws : Vec A m) (xs : Vec B m) {ys : Vec A n} {zs} →
+ All₂ _~_ (ws ++ ys) (xs ++ zs) →
+ All₂ _~_ ws xs × All₂ _~_ ys zs
+ All₂-++⁻ ws xs ps = All₂-++ˡ⁻ ws xs ps , All₂-++ʳ⁻ ws xs ps
+
+------------------------------------------------------------------------
+-- All and concat
+
+module _ {a m p} {A : Set a} {P : A → Set p} where
+
+ All-concat⁺ : ∀ {n} {xss : Vec (Vec A m) n} →
+ All (All P) xss → All P (concat xss)
+ All-concat⁺ [] = []
+ All-concat⁺ (pxs ∷ pxss) = All-++⁺ pxs (All-concat⁺ pxss)
+
+ All-concat⁻ : ∀ {n} (xss : Vec (Vec A m) n) →
+ All P (concat xss) → All (All P) xss
+ All-concat⁻ [] [] = []
+ All-concat⁻ (xs ∷ xss) pxss = All-++ˡ⁻ xs pxss ∷ All-concat⁻ xss (All-++ʳ⁻ xs pxss)
+
+------------------------------------------------------------------------
+-- All₂ and concat
+
+module _ {a b m p} {A : Set a} {B : Set b} {_~_ : REL A B p} where
+
+ All₂-concat⁺ : ∀ {n} {xss : Vec (Vec A m) n} {yss} →
+ All₂ (All₂ _~_) xss yss →
+ All₂ _~_ (concat xss) (concat yss)
+ All₂-concat⁺ [] = []
+ All₂-concat⁺ (xs~ys ∷ ps) = All₂-++⁺ xs~ys (All₂-concat⁺ ps)
+
+ All₂-concat⁻ : ∀ {n} (xss : Vec (Vec A m) n) yss →
+ All₂ _~_ (concat xss) (concat yss) →
+ All₂ (All₂ _~_) xss yss
+ All₂-concat⁻ [] [] [] = []
+ All₂-concat⁻ (xs ∷ xss) (ys ∷ yss) ps =
+ All₂-++ˡ⁻ xs ys ps ∷ All₂-concat⁻ xss yss (All₂-++ʳ⁻ xs ys ps)
+
diff --git a/src/Data/Vec/Equality.agda b/src/Data/Vec/Equality.agda
index fb12aab..8933a6d 100644
--- a/src/Data/Vec/Equality.agda
+++ b/src/Data/Vec/Equality.agda
@@ -12,7 +12,7 @@ open import Function
open import Level using (_⊔_)
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as P using (_≡_)
-
+open import Relation.Binary.HeterogeneousEquality as H using (_≅_)
module Equality {s₁ s₂} (S : Setoid s₁ s₂) where
private
@@ -50,6 +50,10 @@ module Equality {s₁ s₂} (S : Setoid s₁ s₂) where
trans (x≈y ∷-cong xs≈ys) (y≈z ∷-cong ys≈zs) =
SS.trans x≈y y≈z ∷-cong trans xs≈ys ys≈zs
+ xs++[]≈xs : ∀ {n} (xs : Vec A n) → xs ++ [] ≈ xs
+ xs++[]≈xs [] = []-cong
+ xs++[]≈xs (x ∷ xs) = SS.refl ∷-cong (xs++[]≈xs xs)
+
_++-cong_ : ∀ {n₁¹ n₂¹} {xs₁¹ : Vec A n₁¹} {xs₂¹ : Vec A n₂¹}
{n₁² n₂²} {xs₁² : Vec A n₁²} {xs₂² : Vec A n₂²} →
xs₁¹ ≈ xs₁² → xs₂¹ ≈ xs₂² →
@@ -91,3 +95,11 @@ module PropositionalEquality {a} {A : Set a} where
from-≡ : ∀ {n} {xs ys : Vec A n} → xs ≡ ys → xs ≈ ys
from-≡ P.refl = refl _
+
+ to-≅ : ∀ {m n} {xs : Vec A m} {ys : Vec A n} →
+ xs ≈ ys → xs ≅ ys
+ to-≅ p with length-equal p
+ to-≅ p | P.refl = H.≡-to-≅ (to-≡ p)
+
+ xs++[]≅xs : ∀ {n} → (xs : Vec A n) → (xs ++ []) ≅ xs
+ xs++[]≅xs xs = to-≅ (xs++[]≈xs xs)
diff --git a/src/Data/Vec/Properties.agda b/src/Data/Vec/Properties.agda
index bc25e92..8a3d370 100644
--- a/src/Data/Vec/Properties.agda
+++ b/src/Data/Vec/Properties.agda
@@ -8,17 +8,18 @@ module Data.Vec.Properties where
open import Algebra
open import Category.Applicative.Indexed
+import Category.Functor as Fun
+open import Category.Functor.Identity using (IdentityFunctor)
open import Category.Monad
open import Category.Monad.Identity
open import Data.Vec
-open import Data.List.Any
- using (here; there) renaming (module Membership-≡ to List)
+open import Data.List.Any using (here; there)
+import Data.List.Any.Membership.Propositional as List
open import Data.Nat
+open import Data.Nat.Properties using (+-assoc)
open import Data.Empty using (⊥-elim)
-import Data.Nat.Properties as Nat
open import Data.Fin as Fin using (Fin; zero; suc; toℕ; fromℕ)
open import Data.Fin.Properties using (_+′_)
-open import Data.Empty using (⊥-elim)
open import Data.Product as Prod using (_×_; _,_; proj₁; proj₂; <_,_>)
open import Function
open import Function.Inverse using (_↔_)
@@ -55,7 +56,18 @@ open import Relation.Binary.HeterogeneousEquality using (_≅_; refl)
(x ∷ xs) ≡ (y ∷ ys) → x ≡ y × xs ≡ ys
∷-injective refl = refl , refl
--- lookup is an applicative functor morphism.
+-- lookup is a functor morphism from Vec to Identity.
+
+lookup-map : ∀ {a b n} {A : Set a} {B : Set b} (i : Fin n) (f : A → B) (xs : Vec A n) →
+ lookup i (map f xs) ≡ f (lookup i xs)
+lookup-map zero f (x ∷ xs) = refl
+lookup-map (suc i) f (x ∷ xs) = lookup-map i f xs
+
+lookup-functor-morphism : ∀ {a n} (i : Fin n) →
+ Fun.Morphism (functor {a = a} {n = n}) IdentityFunctor
+lookup-functor-morphism i = record { op = lookup i ; op-<$> = lookup-map i }
+
+-- lookup is even an applicative functor morphism.
lookup-morphism :
∀ {a n} (i : Fin n) →
@@ -67,8 +79,8 @@ lookup-morphism i = record
; op-⊛ = lookup-⊛ i
}
where
- lookup-replicate : ∀ {a n} {A : Set a} (i : Fin n) →
- lookup i ∘ replicate {A = A} ≗ id {A = A}
+ lookup-replicate : ∀ {a n} {A : Set a} (i : Fin n) (x : A) →
+ lookup i (replicate x) ≡ x
lookup-replicate zero = λ _ → refl
lookup-replicate (suc i) = lookup-replicate i
@@ -162,14 +174,80 @@ map-∘ : ∀ {a b c n} {A : Set a} {B : Set b} {C : Set c}
map-∘ f g [] = refl
map-∘ f g (x ∷ xs) = P.cong (_∷_ (f (g x))) (map-∘ f g xs)
--- tabulate can be expressed using map and allFin.
+-- Laws of the applicative.
+
+-- Idiomatic application is a special case of zipWith:
+-- _⊛_ = zipWith id
+
+⊛-is-zipWith : ∀ {a b n} {A : Set a} {B : Set b} →
+ (fs : Vec (A → B) n) (xs : Vec A n) →
+ (fs ⊛ xs) ≡ zipWith _$_ fs xs
+⊛-is-zipWith [] [] = refl
+⊛-is-zipWith (f ∷ fs) (x ∷ xs) = P.cong (f x ∷_) (⊛-is-zipWith fs xs)
+
+-- map is expressible via idiomatic application
+
+map-is-⊛ : ∀ {a b n} {A : Set a} {B : Set b} (f : A → B) (xs : Vec A n) →
+ map f xs ≡ (replicate f ⊛ xs)
+map-is-⊛ f [] = refl
+map-is-⊛ f (x ∷ xs) = P.cong (_ ∷_) (map-is-⊛ f xs)
+
+-- zipWith is expressible via idiomatic application
+
+zipWith-is-⊛ : ∀ {a b c n} {A : Set a} {B : Set b} {C : Set c} →
+ (f : A → B → C) (xs : Vec A n) (ys : Vec B n) →
+ zipWith f xs ys ≡ (replicate f ⊛ xs ⊛ ys)
+zipWith-is-⊛ f [] [] = refl
+zipWith-is-⊛ f (x ∷ xs) (y ∷ ys) = P.cong (_ ∷_) (zipWith-is-⊛ f xs ys)
+
+-- Applicative fusion laws for map and zipWith.
+
+-- pulling a replicate into a map
+
+map-replicate : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) (x : A) (n : ℕ) →
+ map f (replicate {n = n} x) ≡ replicate {n = n} (f x)
+map-replicate f x zero = refl
+map-replicate f x (suc n) = P.cong (f x ∷_) (map-replicate f x n)
+
+-- pulling a replicate into a zipWith
+
+zipWith-replicate₁ : ∀ {a b c n} {A : Set a} {B : Set b} {C : Set c} →
+ (_⊕_ : A → B → C) (x : A) (ys : Vec B n) →
+ zipWith _⊕_ (replicate x) ys ≡ map (x ⊕_) ys
+zipWith-replicate₁ _⊕_ x [] = refl
+zipWith-replicate₁ _⊕_ x (y ∷ ys) = P.cong ((x ⊕ y) ∷_) (zipWith-replicate₁ _⊕_ x ys)
+
+zipWith-replicate₂ : ∀ {a b c n} {A : Set a} {B : Set b} {C : Set c} →
+ (_⊕_ : A → B → C) (xs : Vec A n) (y : B) →
+ zipWith _⊕_ xs (replicate y) ≡ map (_⊕ y) xs
+zipWith-replicate₂ _⊕_ [] y = refl
+zipWith-replicate₂ _⊕_ (x ∷ xs) y = P.cong ((x ⊕ y) ∷_) (zipWith-replicate₂ _⊕_ xs y)
+
+-- pulling a map into a zipWith
+
+zipWith-map₁ : ∀ {a b c d n} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
+ (_⊕_ : B → C → D) (f : A → B) (xs : Vec A n) (ys : Vec C n) →
+ zipWith _⊕_ (map f xs) ys ≡ zipWith (λ x y → f x ⊕ y) xs ys
+zipWith-map₁ _⊕_ f [] [] = refl
+zipWith-map₁ _⊕_ f (x ∷ xs) (y ∷ ys) = P.cong ((f x ⊕ y) ∷_) (zipWith-map₁ _⊕_ f xs ys)
+
+zipWith-map₂ : ∀ {a b c d n} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
+ (_⊕_ : A → C → D) (f : B → C) (xs : Vec A n) (ys : Vec B n) →
+ zipWith _⊕_ xs (map f ys) ≡ zipWith (λ x y → x ⊕ f y) xs ys
+zipWith-map₂ _⊕_ f [] [] = refl
+zipWith-map₂ _⊕_ f (x ∷ xs) (y ∷ ys) = P.cong ((x ⊕ f y) ∷_) (zipWith-map₂ _⊕_ f xs ys)
+
+-- Tabulation.
+
+-- mapping over a tabulation is the tabulation of the composition
tabulate-∘ : ∀ {n a b} {A : Set a} {B : Set b}
(f : A → B) (g : Fin n → A) →
tabulate (f ∘ g) ≡ map f (tabulate g)
tabulate-∘ {zero} f g = refl
-tabulate-∘ {suc n} f g =
- P.cong (_∷_ (f (g zero))) (tabulate-∘ f (g ∘ suc))
+tabulate-∘ {suc n} f g = P.cong (f (g zero) ∷_) (tabulate-∘ f (g ∘ suc))
+
+-- tabulate can be expressed using map and allFin.
tabulate-allFin : ∀ {n a} {A : Set a} (f : Fin n → A) →
tabulate f ≡ map f (allFin n)
@@ -242,7 +320,6 @@ sum-++-commute (x ∷ xs) {ys} = begin
where
open P.≡-Reasoning
- open CommutativeSemiring Nat.commutativeSemiring hiding (_+_; sym)
-- foldr is a congruence.
diff --git a/src/Foreign/Haskell.agda b/src/Foreign/Haskell.agda
index 1a88fb4..55572be 100644
--- a/src/Foreign/Haskell.agda
+++ b/src/Foreign/Haskell.agda
@@ -11,5 +11,5 @@ module Foreign.Haskell where
data Unit : Set where
unit : Unit
-{-# COMPILED_DATA Unit () () #-}
-{-# COMPILED_DATA_UHC Unit __UNIT__ __UNIT__ #-}
+{-# COMPILE GHC Unit = data () (()) #-}
+{-# COMPILE UHC Unit = data __UNIT__ (__UNIT__) #-}
diff --git a/src/Function.agda b/src/Function.agda
index 52bcc32..e04bcc7 100644
--- a/src/Function.agda
+++ b/src/Function.agda
@@ -7,12 +7,13 @@
module Function where
open import Level
+open import Strict
infixr 9 _∘_ _∘′_
infixl 8 _ˢ_
infixl 1 _on_
infixl 1 _⟨_⟩_
-infixr 0 _-[_]-_ _$_
+infixr 0 _-[_]-_ _$_ _$′_ _$!_ _$!′_
infixl 0 _∋_
------------------------------------------------------------------------
@@ -71,6 +72,21 @@ _$_ : ∀ {a b} {A : Set a} {B : A → Set b} →
((x : A) → B x) → ((x : A) → B x)
f $ x = f x
+_$′_ : ∀ {a b} {A : Set a} {B : Set b} →
+ (A → B) → (A → B)
+_$′_ = _$_
+
+-- Strict (call-by-value) application
+
+_$!_ : ∀ {a b} {A : Set a} {B : A → Set b} →
+ ((x : A) → B x) → ((x : A) → B x)
+_$!_ = flip force
+
+_$!′_ : ∀ {a b} {A : Set a} {B : Set b} →
+ (A → B) → (A → B)
+_$!′_ = _$!_
+
+
_⟨_⟩_ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} →
A → (A → B → C) → B → C
x ⟨ f ⟩ y = f x y
diff --git a/src/IO/Primitive.agda b/src/IO/Primitive.agda
index b6a33fc..8793ea0 100644
--- a/src/IO/Primitive.agda
+++ b/src/IO/Primitive.agda
@@ -21,10 +21,10 @@ 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 _>>=_ (\_ _ _ _ -> (>>=)) #-}
-{-# COMPILED_UHC return (\_ _ x -> UHC.Agda.Builtins.primReturn x) #-}
-{-# COMPILED_UHC _>>=_ (\_ _ _ _ x y -> UHC.Agda.Builtins.primBind x y) #-}
+{-# COMPILE GHC return = \_ _ -> return #-}
+{-# COMPILE GHC _>>=_ = \_ _ _ _ -> (>>=) #-}
+{-# COMPILE UHC return = \_ _ x -> UHC.Agda.Builtins.primReturn x #-}
+{-# COMPILE UHC _>>=_ = \_ _ _ _ x y -> UHC.Agda.Builtins.primBind x y #-}
------------------------------------------------------------------------
-- Simple lazy IO
@@ -38,10 +38,10 @@ postulate
-- the locale. For older versions of the library (going back to at
-- least version 3) the functions use ISO-8859-1.
-{-# IMPORT Data.Text #-}
-{-# IMPORT Data.Text.IO #-}
-{-# IMPORT System.IO #-}
-{-# IMPORT Control.Exception #-}
+{-# FOREIGN GHC import qualified Data.Text #-}
+{-# FOREIGN GHC import qualified Data.Text.IO #-}
+{-# FOREIGN GHC import qualified System.IO #-}
+{-# FOREIGN GHC import qualified Control.Exception #-}
postulate
getContents : IO Costring
@@ -56,7 +56,7 @@ postulate
readFiniteFile : String → IO String
-{-# HASKELL
+{-# FOREIGN GHC
readFiniteFile :: Data.Text.Text -> IO Data.Text.Text
readFiniteFile f = do
h <- System.IO.openFile (Data.Text.unpack f) System.IO.ReadMode
@@ -65,17 +65,17 @@ postulate
Data.Text.IO.hGetContents h
#-}
-{-# COMPILED getContents getContents #-}
-{-# COMPILED readFile (readFile . Data.Text.unpack) #-}
-{-# COMPILED writeFile (\x -> writeFile (Data.Text.unpack x)) #-}
-{-# COMPILED appendFile (\x -> appendFile (Data.Text.unpack x)) #-}
-{-# COMPILED putStr putStr #-}
-{-# COMPILED putStrLn putStrLn #-}
-{-# COMPILED readFiniteFile readFiniteFile #-}
-{-# COMPILED_UHC getContents (UHC.Agda.Builtins.primGetContents) #-}
-{-# COMPILED_UHC readFile (UHC.Agda.Builtins.primReadFile) #-}
-{-# COMPILED_UHC writeFile (UHC.Agda.Builtins.primWriteFile) #-}
-{-# COMPILED_UHC appendFile (UHC.Agda.Builtins.primAppendFile) #-}
-{-# COMPILED_UHC putStr (UHC.Agda.Builtins.primPutStr) #-}
-{-# COMPILED_UHC putStrLn (UHC.Agda.Builtins.primPutStrLn) #-}
-{-# COMPILED_UHC readFiniteFile (UHC.Agda.Builtins.primReadFiniteFile) #-}
+{-# COMPILE GHC getContents = getContents #-}
+{-# COMPILE GHC readFile = readFile . Data.Text.unpack #-}
+{-# COMPILE GHC writeFile = \x -> writeFile (Data.Text.unpack x) #-}
+{-# COMPILE GHC appendFile = \x -> appendFile (Data.Text.unpack x) #-}
+{-# COMPILE GHC putStr = putStr #-}
+{-# COMPILE GHC putStrLn = putStrLn #-}
+{-# COMPILE GHC readFiniteFile = readFiniteFile #-}
+{-# COMPILE UHC getContents = UHC.Agda.Builtins.primGetContents #-}
+{-# COMPILE UHC readFile = UHC.Agda.Builtins.primReadFile #-}
+{-# COMPILE UHC writeFile = UHC.Agda.Builtins.primWriteFile #-}
+{-# COMPILE UHC appendFile = UHC.Agda.Builtins.primAppendFile #-}
+{-# COMPILE UHC putStr = UHC.Agda.Builtins.primPutStr #-}
+{-# COMPILE UHC putStrLn = UHC.Agda.Builtins.primPutStrLn #-}
+{-# COMPILE UHC readFiniteFile = UHC.Agda.Builtins.primReadFiniteFile #-}
diff --git a/src/Induction/Nat.agda b/src/Induction/Nat.agda
index 23a61b4..6a20675 100644
--- a/src/Induction/Nat.agda
+++ b/src/Induction/Nat.agda
@@ -8,6 +8,7 @@ module Induction.Nat where
open import Function
open import Data.Nat
+open import Data.Nat.Properties using (≤⇒≤′)
open import Data.Fin using (_≺_)
open import Data.Fin.Properties
open import Data.Product
@@ -50,18 +51,33 @@ cRec = build cRec-builder
------------------------------------------------------------------------
-- Complete induction based on _<′_
-<-Rec : ∀ {ℓ} → RecStruct ℕ ℓ ℓ
-<-Rec = WfRec _<′_
+<′-Rec : ∀ {ℓ} → RecStruct ℕ ℓ ℓ
+<′-Rec = WfRec _<′_
mutual
- <-well-founded : Well-founded _<′_
- <-well-founded n = acc (<-well-founded′ n)
+ <′-well-founded : Well-founded _<′_
+ <′-well-founded n = acc (<′-well-founded′ n)
+
+ <′-well-founded′ : ∀ n → <′-Rec (Acc _<′_) n
+ <′-well-founded′ zero _ ()
+ <′-well-founded′ (suc n) .n ≤′-refl = <′-well-founded n
+ <′-well-founded′ (suc n) m (≤′-step m<n) = <′-well-founded′ n m m<n
+
+module _ {ℓ} where
+ open WF.All <′-well-founded ℓ public
+ renaming ( wfRec-builder to <′-rec-builder
+ ; wfRec to <′-rec
+ )
+
+------------------------------------------------------------------------
+-- Complete induction based on _<_
+
+<-Rec : ∀ {ℓ} → RecStruct ℕ ℓ ℓ
+<-Rec = WfRec _<_
- <-well-founded′ : ∀ n → <-Rec (Acc _<′_) n
- <-well-founded′ zero _ ()
- <-well-founded′ (suc n) .n ≤′-refl = <-well-founded n
- <-well-founded′ (suc n) m (≤′-step m<n) = <-well-founded′ n m m<n
+<-well-founded : Well-founded _<_
+<-well-founded = Subrelation.well-founded ≤⇒≤′ <′-well-founded
module _ {ℓ} where
open WF.All <-well-founded ℓ public
@@ -76,7 +92,7 @@ module _ {ℓ} where
≺-Rec = WfRec _≺_
≺-well-founded : Well-founded _≺_
-≺-well-founded = Subrelation.well-founded ≺⇒<′ <-well-founded
+≺-well-founded = Subrelation.well-founded ≺⇒<′ <′-well-founded
module _ {ℓ} where
open WF.All ≺-well-founded ℓ public
@@ -127,7 +143,7 @@ private
}
half₂ : ℕ → ℕ
- half₂ = <-rec _ half₂-step
+ half₂ = <′-rec _ half₂-step
-- The application half₁ (2 + n) is definitionally equal to
-- 1 + half₁ n. Perhaps it is instructive to see why.
@@ -165,32 +181,32 @@ private
half₂ (2 + n) ≡⟨⟩
- <-rec (λ _ → ℕ) half₂-step (2 + n) ≡⟨⟩
+ <′-rec (λ _ → ℕ) half₂-step (2 + n) ≡⟨⟩
- half₂-step (2 + n) (<-rec-builder (λ _ → ℕ) half₂-step (2 + n)) ≡⟨⟩
+ half₂-step (2 + n) (<′-rec-builder (λ _ → ℕ) half₂-step (2 + n)) ≡⟨⟩
- 1 + <-rec-builder (λ _ → ℕ) half₂-step (2 + n) n (≤′-step ≤′-refl) ≡⟨⟩
+ 1 + <′-rec-builder (λ _ → ℕ) half₂-step (2 + n) n (≤′-step ≤′-refl) ≡⟨⟩
1 + Some.wfRec-builder (λ _ → ℕ) half₂-step (2 + n)
- (<-well-founded (2 + n)) n (≤′-step ≤′-refl) ≡⟨⟩
+ (<′-well-founded (2 + n)) n (≤′-step ≤′-refl) ≡⟨⟩
1 + Some.wfRec-builder (λ _ → ℕ) half₂-step (2 + n)
- (acc (<-well-founded′ (2 + n))) n (≤′-step ≤′-refl) ≡⟨⟩
+ (acc (<′-well-founded′ (2 + n))) n (≤′-step ≤′-refl) ≡⟨⟩
1 + half₂-step n
(Some.wfRec-builder (λ _ → ℕ) half₂-step n
- (<-well-founded′ (2 + n) n (≤′-step ≤′-refl))) ≡⟨⟩
+ (<′-well-founded′ (2 + n) n (≤′-step ≤′-refl))) ≡⟨⟩
1 + half₂-step n
(Some.wfRec-builder (λ _ → ℕ) half₂-step n
- (<-well-founded′ (1 + n) n ≤′-refl)) ≡⟨⟩
+ (<′-well-founded′ (1 + n) n ≤′-refl)) ≡⟨⟩
1 + half₂-step n
- (Some.wfRec-builder (λ _ → ℕ) half₂-step n (<-well-founded n)) ≡⟨⟩
+ (Some.wfRec-builder (λ _ → ℕ) half₂-step n (<′-well-founded n)) ≡⟨⟩
- 1 + half₂-step n (<-rec-builder (λ _ → ℕ) half₂-step n) ≡⟨⟩
+ 1 + half₂-step n (<′-rec-builder (λ _ → ℕ) half₂-step n) ≡⟨⟩
- 1 + <-rec (λ _ → ℕ) half₂-step n ≡⟨⟩
+ 1 + <′-rec (λ _ → ℕ) half₂-step n ≡⟨⟩
1 + half₂ n ∎
@@ -216,10 +232,10 @@ private
}
-- Some properties that the functions above satisfy, proved using
- -- <-rec.
+ -- <′-rec.
half₁-+₂ : ∀ n → half₁ (twice n) ≡ n
- half₁-+₂ = <-rec _ λ
+ half₁-+₂ = <′-rec _ λ
{ zero _ → refl
; (suc zero) _ → refl
; (suc (suc n)) rec →
@@ -227,7 +243,7 @@ private
}
half₂-+₂ : ∀ n → half₂ (twice n) ≡ n
- half₂-+₂ = <-rec _ λ
+ half₂-+₂ = <′-rec _ λ
{ zero _ → refl
; (suc zero) _ → refl
; (suc (suc n)) rec →
diff --git a/src/Irrelevance.agda b/src/Irrelevance.agda
deleted file mode 100644
index 31fb54a..0000000
--- a/src/Irrelevance.agda
+++ /dev/null
@@ -1,9 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- The irrelevance axiom
-------------------------------------------------------------------------
-
-module Irrelevance where
-
-import Level
diff --git a/src/Relation/Binary/Consequences.agda b/src/Relation/Binary/Consequences.agda
index 4b0f071..dafdc88 100644
--- a/src/Relation/Binary/Consequences.agda
+++ b/src/Relation/Binary/Consequences.agda
@@ -114,6 +114,11 @@ trans∧tri⟶resp≈ {_≈_ = _≈_} {_<_} sym trans <-trans tri =
... | tri≈ _ y≈z _ = ⊥-elim (tri⟶irr tri (trans x≈y y≈z) x<z)
... | tri> _ _ z<y = ⊥-elim (tri⟶irr tri x≈y (<-trans x<z z<y))
+P-resp⟶¬P-resp :
+ ∀ {a p ℓ} {A : Set a} {_≈_ : Rel A ℓ} {P : A → Set p} →
+ Symmetric _≈_ → P Respects _≈_ → (¬_ ∘ P) Respects _≈_
+P-resp⟶¬P-resp sym resp x≈y ¬Px Py = ¬Px (resp (sym x≈y) Py)
+
tri⟶dec≈ :
∀ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} →
Trichotomous _≈_ _<_ → Decidable _≈_
diff --git a/src/Relation/Binary/HeterogeneousEquality.agda b/src/Relation/Binary/HeterogeneousEquality.agda
index a15d71d..12ab465 100644
--- a/src/Relation/Binary/HeterogeneousEquality.agda
+++ b/src/Relation/Binary/HeterogeneousEquality.agda
@@ -34,10 +34,18 @@ x ≇ y = ¬ x ≅ y
------------------------------------------------------------------------
-- Conversion
+open Core public using (≅-to-≡)
+
≡-to-≅ : ∀ {a} {A : Set a} {x y : A} → x ≡ y → x ≅ y
≡-to-≅ refl = refl
-open Core public using (≅-to-≡)
+≅-to-type-≡ : ∀ {a} {A B : Set a} {x : A} {y : B} →
+ x ≅ y → A ≡ B
+≅-to-type-≡ refl = refl
+
+≅-to-subst-≡ : ∀ {a} {A B : Set a} {x : A} {y : B} → (p : x ≅ y) →
+ P.subst (λ x → x) (≅-to-type-≡ p) x ≡ y
+≅-to-subst-≡ refl = refl
------------------------------------------------------------------------
-- Some properties
diff --git a/src/Relation/Binary/StrictPartialOrderReasoning.agda b/src/Relation/Binary/StrictPartialOrderReasoning.agda
index 62a13eb..bf774f8 100644
--- a/src/Relation/Binary/StrictPartialOrderReasoning.agda
+++ b/src/Relation/Binary/StrictPartialOrderReasoning.agda
@@ -12,4 +12,10 @@ module Relation.Binary.StrictPartialOrderReasoning
import Relation.Binary.PreorderReasoning as PreR
import Relation.Binary.Properties.StrictPartialOrder as SPO
-open PreR (SPO.preorder S) public renaming (_∼⟨_⟩_ to _<⟨_⟩_)
+open PreR (SPO.preorder S) public
+open import Data.Sum
+
+_<⟨_⟩_ : ∀ x {y z} → _ → y IsRelatedTo z → x IsRelatedTo z
+x <⟨ x∼y ⟩ y∼z = x ∼⟨ inj₁ x∼y ⟩ y∼z
+
+infixr 2 _<⟨_⟩_
diff --git a/src/Relation/Binary/Vec/Pointwise.agda b/src/Relation/Binary/Vec/Pointwise.agda
index 722645f..f944ea2 100644
--- a/src/Relation/Binary/Vec/Pointwise.agda
+++ b/src/Relation/Binary/Vec/Pointwise.agda
@@ -6,7 +6,7 @@
module Relation.Binary.Vec.Pointwise where
-open import Category.Applicative.Indexed
+open import Category.Functor
open import Data.Fin
open import Data.Nat
open import Data.Plus as Plus hiding (equivalent; map)
@@ -227,7 +227,7 @@ gmap : ∀ {ℓ} {A A′ : Set ℓ}
_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
+ let module M = Morphism (VecProp.lookup-functor-morphism i) in
P.subst₂ _R′_ (P.sym $ M.op-<$> f xs)
(P.sym $ M.op-<$> f ys)
(R⇒R′ (Pointwise.app xsRys i))
diff --git a/src/Relation/Nullary/Negation.agda b/src/Relation/Nullary/Negation.agda
index cd3145e..1376f2b 100644
--- a/src/Relation/Nullary/Negation.agda
+++ b/src/Relation/Nullary/Negation.agda
@@ -12,8 +12,6 @@ open import Data.Bool.Base using (Bool; false; true; if_then_else_)
open import Data.Empty
open import Function
open import Data.Product as Prod
-open import Data.Fin using (Fin)
-open import Data.Fin.Dec
open import Data.Sum as Sum
open import Category.Monad
open import Level
@@ -62,13 +60,6 @@ private
∃ (λ x → ¬ P x) → ¬ (∀ x → P x)
∃¬⟶¬∀ = flip ∀⟶¬∃¬
--- When P is a decidable predicate over a finite set the following
--- lemma can be proved.
-
-¬∀⟶∃¬ : ∀ n {p} (P : Fin n → Set p) → Decidable P →
- ¬ (∀ i → P i) → ∃ λ i → ¬ P i
-¬∀⟶∃¬ n P dec ¬P = Prod.map id proj₁ $ ¬∀⟶∃¬-smallest n P dec ¬P
-
------------------------------------------------------------------------
-- Double-negation
diff --git a/src/Strict.agda b/src/Strict.agda
new file mode 100644
index 0000000..78e49b6
--- /dev/null
+++ b/src/Strict.agda
@@ -0,0 +1,29 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Strictness combinators
+------------------------------------------------------------------------
+
+module Strict where
+
+open import Level
+open import Agda.Builtin.Equality
+
+open import Agda.Builtin.Strict
+ renaming ( primForce to force
+ ; primForceLemma to force-≡) public
+
+-- Derived combinators
+module _ {ℓ ℓ′ : Level} {A : Set ℓ} {B : Set ℓ′} where
+
+ force′ : A → (A → B) → B
+ force′ = force
+
+ force′-≡ : (a : A) (f : A → B) → force′ a f ≡ f a
+ force′-≡ = force-≡
+
+ seq : A → B → B
+ seq a b = force a (λ _ → b)
+
+ seq-≡ : (a : A) (b : B) → seq a b ≡ b
+ seq-≡ a b = force-≡ a (λ _ → b)